ImpCEvalFunAn Evaluation Function for Imp
Here was our first try at an evaluation function for commands,
omitting WHILE.
Open Scope imp_scope.
Fixpoint ceval_step1 (st : state) (c : com) : state :=
match c with
| SKIP ⇒
st
| l ::= a1 ⇒
(l !-> aeval st a1 ; st)
| c1 ;; c2 ⇒
let st' := ceval_step1 st c1 in
ceval_step1 st' c2
| TEST b THEN c1 ELSE c2 FI ⇒
if (beval st b)
then ceval_step1 st c1
else ceval_step1 st c2
| WHILE b1 DO c1 END ⇒
st (* bogus *)
end.
Close Scope imp_scope.
As we remarked in chapter Imp, in a traditional functional programming language like ML or Haskell we could write the WHILE case as follows:
| WHILE b1 DO c1 END => if (beval st b1) then ceval_step1 st (c1;; WHILE b1 DO c1 END) else stCoq doesn't accept such a definition (Error: Cannot guess decreasing argument of fix) because the function we want to define is not guaranteed to terminate. Indeed, the changed ceval_step1 function applied to the loop program from Imp.v would never terminate. Since Coq is not just a functional programming language, but also a consistent logic, any potentially non-terminating function needs to be rejected. Here is an invalid(!) Coq program showing what would go wrong if Coq allowed non-terminating recursive functions:
Fixpoint loop_false (n : nat) : False := loop_false n.That is, propositions like False would become provable (e.g., loop_false 0 would be a proof of False), which would be a disaster for Coq's logical consistency.
A Step-Indexed Evaluator
Open Scope imp_scope.
Fixpoint ceval_step2 (st : state) (c : com) (i : nat) : state :=
match i with
| O ⇒ empty_st
| S i' ⇒
match c with
| SKIP ⇒
st
| l ::= a1 ⇒
(l !-> aeval st a1 ; st)
| c1 ;; c2 ⇒
let st' := ceval_step2 st c1 i' in
ceval_step2 st' c2 i'
| TEST b THEN c1 ELSE c2 FI ⇒
if (beval st b)
then ceval_step2 st c1 i'
else ceval_step2 st c2 i'
| WHILE b1 DO c1 END ⇒
if (beval st b1)
then let st' := ceval_step2 st c1 i' in
ceval_step2 st' c i'
else st
end
end.
Close Scope imp_scope.
Open Scope imp_scope.
Fixpoint ceval_step3 (st : state) (c : com) (i : nat)
: option state :=
match i with
| O ⇒ None
| S i' ⇒
match c with
| SKIP ⇒
Some st
| l ::= a1 ⇒
Some (l !-> aeval st a1 ; st)
| c1 ;; c2 ⇒
match (ceval_step3 st c1 i') with
| Some st' ⇒ ceval_step3 st' c2 i'
| None ⇒ None
end
| TEST b THEN c1 ELSE c2 FI ⇒
if (beval st b)
then ceval_step3 st c1 i'
else ceval_step3 st c2 i'
| WHILE b1 DO c1 END ⇒
if (beval st b1)
then match (ceval_step3 st c1 i') with
| Some st' ⇒ ceval_step3 st' c i'
| None ⇒ None
end
else Some st
end
end.
Close Scope imp_scope.
We can improve the readability of this version by introducing a
bit of auxiliary notation to hide the plumbing involved in
repeatedly matching against optional states.
Notation "'LETOPT' x <== e1 'IN' e2"
:= (match e1 with
| Some x ⇒ e2
| None ⇒ None
end)
(right associativity, at level 60).
Open Scope imp_scope.
Fixpoint ceval_step (st : state) (c : com) (i : nat)
: option state :=
match i with
| O ⇒ None
| S i' ⇒
match c with
| SKIP ⇒
Some st
| l ::= a1 ⇒
Some (l !-> aeval st a1 ; st)
| c1 ;; c2 ⇒
LETOPT st' <== ceval_step st c1 i' IN
ceval_step st' c2 i'
| TEST b THEN c1 ELSE c2 FI ⇒
if (beval st b)
then ceval_step st c1 i'
else ceval_step st c2 i'
| WHILE b1 DO c1 END ⇒
if (beval st b1)
then LETOPT st' <== ceval_step st c1 i' IN
ceval_step st' c i'
else Some st
end
end.
Close Scope imp_scope.
Definition test_ceval (st:state) (c:com) :=
match ceval_step st c 500 with
| None ⇒ None
| Some st ⇒ Some (st X, st Y, st Z)
end.
(* Compute
(test_ceval empty_st
(X ::= 2;;
TEST (X <= 1)
THEN Y ::= 3
ELSE Z ::= 4
FI)).
====>
Some (2, 0, 4) *)
Theorem ceval_step__ceval: ∀ c st st',
(∃ i, ceval_step st c i = Some st') →
st =[ c ]=> st'.
Theorem ceval_step_more: ∀ i1 i2 st st' c,
i1 ≤ i2 →
ceval_step st c i1 = Some st' →
ceval_step st c i2 = Some st'.
Proof.
induction i1 as [|i1']; intros i2 st st' c Hle Hceval.
- (* i1 = 0 *)
simpl in Hceval. discriminate Hceval.
- (* i1 = S i1' *)
destruct i2 as [|i2']. inversion Hle.
assert (Hle': i1' ≤ i2') by omega.
destruct c.
+ (* SKIP *)
simpl in Hceval. inversion Hceval.
reflexivity.
+ (* ::= *)
simpl in Hceval. inversion Hceval.
reflexivity.
+ (* ;; *)
simpl in Hceval. simpl.
destruct (ceval_step st c1 i1') eqn:Heqst1'o.
× (* st1'o = Some *)
apply (IHi1' i2') in Heqst1'o; try assumption.
rewrite Heqst1'o. simpl. simpl in Hceval.
apply (IHi1' i2') in Hceval; try assumption.
× (* st1'o = None *)
discriminate Hceval.
+ (* TEST *)
simpl in Hceval. simpl.
destruct (beval st b); apply (IHi1' i2') in Hceval;
assumption.
+ (* WHILE *)
simpl in Hceval. simpl.
destruct (beval st b); try assumption.
destruct (ceval_step st c i1') eqn: Heqst1'o.
× (* st1'o = Some *)
apply (IHi1' i2') in Heqst1'o; try assumption.
rewrite → Heqst1'o. simpl. simpl in Hceval.
apply (IHi1' i2') in Hceval; try assumption.
× (* i1'o = None *)
simpl in Hceval. discriminate Hceval. Qed.
Theorem ceval__ceval_step: ∀ c st st',
st =[ c ]=> st' →
∃ i, ceval_step st c i = Some st'.
Proof.
intros c st st' Hce.
induction Hce.
(* FILL IN HERE *) Admitted.
Theorem ceval_and_ceval_step_coincide: ∀ c st st',
st =[ c ]=> st'
↔ ∃ i, ceval_step st c i = Some st'.
Proof.
intros c st st'.
split. apply ceval__ceval_step. apply ceval_step__ceval.
Qed.
Determinism of Evaluation Again
Theorem ceval_deterministic' : ∀ c st st1 st2,
st =[ c ]=> st1 →
st =[ c ]=> st2 →
st1 = st2.