EquivProgram Equivalence
Require Export Imp.
Behavioral Equivalence
Definitions
Definition aequiv (a1 a2 : aexp) : Prop :=
∀(st:state),
aeval st a1 = aeval st a2.
Definition bequiv (b1 b2 : bexp) : Prop :=
∀(st:state),
beval st b1 = beval st b2.
For commands, the situation is a little more subtle.  We can't
    simply say "two commands are behaviorally equivalent if they
    evaluate to the same ending state whenever they are started in the
    same initial state," because some commands (in some starting
    states) don't terminate in any final state at all!  What we need
    instead is this: two commands are behaviorally equivalent if, for
    any given starting state, they either both diverge or both
    terminate in the same final state.  A compact way to express this
    is "if the first one terminates in a particular state then so does
    the second, and vice versa." 
Definition cequiv (c1 c2 : com) : Prop :=
∀(st st' : state),
(c1 / st ⇓ st') ↔ (c2 / st ⇓ st').
Are these two programs equivalent?  
 
 
 
    (1) Yes
 
    (2) No
 
    (3) Not sure
    X ::= 1;;
Y ::= 2 
    and
Y ::= 2
    Y ::= 2;;
X ::= 1 
X ::= 1
What about these?
 
 
 
 
    (1) Yes
 
    (2) No
 
    (3) Not sure
    X ::= 1;;
Y ::= 2 
    and
Y ::= 2
    X ::= 2;;
Y ::= 1 
Y ::= 1
What about these?
 
 
 
 
    WHILE 1 ≤ X DO
X ::= X + 1
END 
    and
X ::= X + 1
END
    WHILE 2 ≤ X DO
X ::= X + 1
END 
X ::= X + 1
END
These?
 
 
    WHILE TRUE DO
WHILE FALSE DO X ::= X + 1 END
END 
    and
WHILE FALSE DO X ::= X + 1 END
END
    WHILE FALSE DO
WHILE TRUE DO X ::= X + 1 END
END 
WHILE TRUE DO X ::= X + 1 END
END
Theorem aequiv_example:
aequiv (AMinus (AId X) (AId X)) (ANum 0).
Proof.
intros st. simpl. omega.
Qed.
intros st. simpl. omega.
Qed.
Theorem bequiv_example:
bequiv (BEq (AMinus (AId X) (AId X)) (ANum 0)) BTrue.
Proof.
intros st. unfold beval.
rewrite aequiv_example. reflexivity.
Qed.
intros st. unfold beval.
rewrite aequiv_example. reflexivity.
Qed.
For examples of command equivalence, let's start by looking at
    some trivial program transformations involving SKIP: 
Theorem skip_left: ∀c,
cequiv
(SKIP;; c)
c.
Proof.
(* WORK IN CLASS *) Admitted.
Similarly, here is a simple transformations that simplifies IFB
    commands: 
Theorem IFB_true_simple: ∀c1 c2,
cequiv
(IFB BTrue THEN c1 ELSE c2 FI)
c1.
Proof.
intros c1 c2.
split; intros H.
Case "→".
inversion H; subst. assumption. inversion H5.
Case "←".
apply E_IfTrue. reflexivity. assumption. Qed.
intros c1 c2.
split; intros H.
Case "→".
inversion H; subst. assumption. inversion H5.
Case "←".
apply E_IfTrue. reflexivity. assumption. Qed.
Of course, few programmers would be tempted to write a conditional
    whose guard is literally BTrue.  A more interesting case is when
    the guard is equivalent to true:
 
   Theorem: If b is equivalent to BTrue, then IFB b THEN c1
   ELSE c2 FI is equivalent to c1.
 
   Proof: 
 
 
   Here is the formal version of this proof: 
-  (→) We must show, for all st and st', that if IFB b
       THEN c1 ELSE c2 FI / st ⇓ st' then c1 / st ⇓ st'.
Proceed by cases on the rules that could possibly have been used to show IFB b THEN c1 ELSE c2 FI / st ⇓ st', namely E_IfTrue and E_IfFalse.-  Suppose the final rule rule in the derivation of IFB b THEN
         c1 ELSE c2 FI / st ⇓ st' was E_IfTrue.  We then have, by
         the premises of E_IfTrue, that c1 / st ⇓ st'.  This is
         exactly what we set out to prove.
-  On the other hand, suppose the final rule in the derivation
         of IFB b THEN c1 ELSE c2 FI / st ⇓ st' was E_IfFalse.
         We then know that beval st b = false and c2 / st ⇓ st'.
Recall that b is equivalent to BTrue, i.e. forall st, beval st b = beval st BTrue. In particular, this means that beval st b = true, since beval st BTrue = true. But this is a contradiction, since E_IfFalse requires that beval st b = false. Thus, the final rule could not have been E_IfFalse.
 
-  Suppose the final rule rule in the derivation of IFB b THEN
         c1 ELSE c2 FI / st ⇓ st' was E_IfTrue.  We then have, by
         the premises of E_IfTrue, that c1 / st ⇓ st'.  This is
         exactly what we set out to prove.
-  (←) We must show, for all st and st', that if c1 / st
       ⇓ st' then IFB b THEN c1 ELSE c2 FI / st ⇓ st'.
Since b is equivalent to BTrue, we know that beval st b = beval st BTrue = true. Together with the assumption that c1 / st ⇓ st', we can apply E_IfTrue to derive IFB b THEN c1 ELSE c2 FI / st ⇓ st'. ☐
Theorem IFB_true: ∀b c1 c2,
bequiv b BTrue →
cequiv
(IFB b THEN c1 ELSE c2 FI)
c1.
Proof.
intros b c1 c2 Hb.
split; intros H.
Case "→".
inversion H; subst.
SCase "b evaluates to true".
assumption.
SCase "b evaluates to false (contradiction)".
rewrite Hb in H5.
inversion H5.
Case "←".
apply E_IfTrue; try assumption.
rewrite Hb. reflexivity. Qed.
Similarly: 
Theorem IFB_false: ∀b c1 c2,
bequiv b BFalse →
cequiv
(IFB b THEN c1 ELSE c2 FI)
c2.
Proof.
(* FILL IN HERE *) Admitted.
(* FILL IN HERE *) Admitted.
For WHILE loops, we can give a similar pair of theorems.  A loop
    whose guard is equivalent to BFalse is equivalent to SKIP,
    while a loop whose guard is equivalent to BTrue is equivalent to
    WHILE BTrue DO SKIP END (or any other non-terminating program).
    The first of these facts is easy. 
Theorem WHILE_false : ∀b c,
bequiv b BFalse →
cequiv
(WHILE b DO c END)
SKIP.
Proof.
intros b c Hb. split; intros H.
Case "→".
inversion H; subst.
SCase "E_WhileEnd".
apply E_Skip.
SCase "E_WhileLoop".
rewrite Hb in H2. inversion H2.
Case "←".
inversion H; subst.
apply E_WhileEnd.
rewrite Hb.
reflexivity. Qed.
intros b c Hb. split; intros H.
Case "→".
inversion H; subst.
SCase "E_WhileEnd".
apply E_Skip.
SCase "E_WhileLoop".
rewrite Hb in H2. inversion H2.
Case "←".
inversion H; subst.
apply E_WhileEnd.
rewrite Hb.
reflexivity. Qed.
To prove the second fact, we need an auxiliary lemma stating that
    WHILE loops whose guards are equivalent to BTrue never
    terminate:
 
    Lemma: If b is equivalent to BTrue, then it cannot be the
    case that (WHILE b DO c END) / st ⇓ st'.
 
    Proof: Suppose that (WHILE b DO c END) / st ⇓ st'.  We show,
    by induction on a derivation of (WHILE b DO c END) / st ⇓ st',
    that this assumption leads to a contradiction.  
 
-  Suppose (WHILE b DO c END) / st ⇓ st' is proved using rule
        E_WhileEnd.  Then by assumption beval st b = false.  But
        this contradicts the assumption that b is equivalent to
        BTrue.
-  Suppose (WHILE b DO c END) / st ⇓ st' is proved using rule
        E_WhileLoop.  Then we are given the induction hypothesis
        that (WHILE b DO c END) / st ⇓ st' is contradictory, which
        is exactly what we are trying to prove!
- Since these are the only rules that could have been used to prove (WHILE b DO c END) / st ⇓ st', the other cases of the induction are immediately contradictory. ☐
Lemma WHILE_true_nonterm : ∀b c st st',
bequiv b BTrue →
~( (WHILE b DO c END) / st ⇓ st' ).
Proof.
(* WORK IN CLASS *) Admitted.
Theorem WHILE_true: ∀b c,
bequiv b BTrue →
cequiv
(WHILE b DO c END)
(WHILE BTrue DO SKIP END).
Proof.
(* FILL IN HERE *) Admitted.
(* FILL IN HERE *) Admitted.
Theorem loop_unrolling: ∀b c,
cequiv
(WHILE b DO c END)
(IFB b THEN (c;; WHILE b DO c END) ELSE SKIP FI).
Proof.
(* WORK IN CLASS *) Admitted.
The Functional Equivalence Axiom
Theorem identity_assignment_first_try : ∀(X:id),
cequiv (X ::= AId X) SKIP.
Proof.
intros. split; intro H.
Case "→".
inversion H; subst. simpl.
replace (update st X (st X)) with st.
constructor.
(* Stuck... *) Abort.
Here we're stuck. The goal looks reasonable, but in fact it is not
    provable!  If we look back at the set of lemmas we proved about
    update in the last chapter, we can see that lemma update_same
    almost does the job, but not quite: it says that the original and
    updated states agree at all values, but this is not the same thing
    as saying that they are = in Coq's sense! 
 
 What's going on?
 
 
 This problem is actually quite general. If we try to prove other
    simple facts, such as
 
 
 
    The reasoning principle we would like to use in these situations
    is called functional extensionality:
- States are function from ids to values
- Functions are equal when their definitions are syntactically the same (modulo simplification)
- Here we have two functions that behave the same on all inputs but are not syntactically the same.
    cequiv (X ::= X + 1;;
X ::= X + 1)
(X ::= X + 2) 
    or
X ::= X + 1)
(X ::= X + 2)
    cequiv (X ::= 1;; Y ::= 2)
(y ::= 2;; X ::= 1)
  
 
    we'll get stuck in the same way: we'll have two functions that
    behave the same way on all inputs, but cannot be proven to be eq
    to each other.
(y ::= 2;; X ::= 1)
| ∀x, f x = g x | |
| f = g | 
Axiom functional_extensionality : ∀{X Y: Type} {f g : X → Y},
(∀(x: X), f x = g x) → f = g.
With the benefit of this axiom we can prove our theorem.  
Theorem identity_assignment : ∀(X:id),
cequiv
(X ::= AId X)
SKIP.
Proof.
intros. split; intro H.
Case "→".
inversion H; subst. simpl.
replace (update st X (st X)) with st.
constructor.
apply functional_extensionality. intro.
rewrite update_same; reflexivity.
Case "←".
inversion H; subst.
assert (st' = (update st' X (st' X))).
apply functional_extensionality. intro.
rewrite update_same; reflexivity.
rewrite H0 at 2.
constructor. reflexivity.
Qed.
Properties of Behavioral Equivalence
Behavioral Equivalence is an Equivalence
Lemma refl_aequiv : ∀(a : aexp), aequiv a a.
Proof.
intros a st. reflexivity. Qed.
intros a st. reflexivity. Qed.
Lemma sym_aequiv : ∀(a1 a2 : aexp),
aequiv a1 a2 → aequiv a2 a1.
Proof.
intros a1 a2 H. intros st. symmetry. apply H. Qed.
intros a1 a2 H. intros st. symmetry. apply H. Qed.
Lemma trans_aequiv : ∀(a1 a2 a3 : aexp),
aequiv a1 a2 → aequiv a2 a3 → aequiv a1 a3.
Proof.
unfold aequiv. intros a1 a2 a3 H12 H23 st.
rewrite (H12 st). rewrite (H23 st). reflexivity. Qed.
unfold aequiv. intros a1 a2 a3 H12 H23 st.
rewrite (H12 st). rewrite (H23 st). reflexivity. Qed.
Lemma refl_bequiv : ∀(b : bexp), bequiv b b.
Proof.
unfold bequiv. intros b st. reflexivity. Qed.
unfold bequiv. intros b st. reflexivity. Qed.
Lemma sym_bequiv : ∀(b1 b2 : bexp),
bequiv b1 b2 → bequiv b2 b1.
Proof.
unfold bequiv. intros b1 b2 H. intros st. symmetry. apply H. Qed.
unfold bequiv. intros b1 b2 H. intros st. symmetry. apply H. Qed.
Lemma trans_bequiv : ∀(b1 b2 b3 : bexp),
bequiv b1 b2 → bequiv b2 b3 → bequiv b1 b3.
Proof.
unfold bequiv. intros b1 b2 b3 H12 H23 st.
rewrite (H12 st). rewrite (H23 st). reflexivity. Qed.
unfold bequiv. intros b1 b2 b3 H12 H23 st.
rewrite (H12 st). rewrite (H23 st). reflexivity. Qed.
Lemma refl_cequiv : ∀(c : com), cequiv c c.
Proof.
unfold cequiv. intros c st st'. apply iff_refl. Qed.
unfold cequiv. intros c st st'. apply iff_refl. Qed.
Lemma sym_cequiv : ∀(c1 c2 : com),
cequiv c1 c2 → cequiv c2 c1.
Proof.
unfold cequiv. intros c1 c2 H st st'.
assert (c1 / st ⇓ st' ↔ c2 / st ⇓ st') as H'.
SCase "Proof of assertion". apply H.
apply iff_sym. assumption.
Qed.
unfold cequiv. intros c1 c2 H st st'.
assert (c1 / st ⇓ st' ↔ c2 / st ⇓ st') as H'.
SCase "Proof of assertion". apply H.
apply iff_sym. assumption.
Qed.
Lemma iff_trans : ∀(P1 P2 P3 : Prop),
(P1 ↔ P2) → (P2 ↔ P3) → (P1 ↔ P3).
Proof.
intros P1 P2 P3 H12 H23.
inversion H12. inversion H23.
split; intros A.
apply H1. apply H. apply A.
apply H0. apply H2. apply A. Qed.
intros P1 P2 P3 H12 H23.
inversion H12. inversion H23.
split; intros A.
apply H1. apply H. apply A.
apply H0. apply H2. apply A. Qed.
Lemma trans_cequiv : ∀(c1 c2 c3 : com),
cequiv c1 c2 → cequiv c2 c3 → cequiv c1 c3.
Proof.
unfold cequiv. intros c1 c2 c3 H12 H23 st st'.
apply iff_trans with (c2 / st ⇓ st'). apply H12. apply H23. Qed.
unfold cequiv. intros c1 c2 c3 H12 H23 st st'.
apply iff_trans with (c2 / st ⇓ st'). apply H12. apply H23. Qed.
Behavioral Equivalence is a Congruence
| aequiv a1 a1' | |
| cequiv (i ::= a1) (i ::= a1') | 
| cequiv c1 c1' | |
| cequiv c2 c2' | |
| cequiv (c1;;c2) (c1';;c2') | 
Theorem CAss_congruence : ∀i a1 a1',
aequiv a1 a1' →
cequiv (CAss i a1) (CAss i a1').
Proof.
intros i a1 a2 Heqv st st'.
split; intros Hceval.
Case "→".
inversion Hceval. subst. apply E_Ass.
rewrite Heqv. reflexivity.
Case "←".
inversion Hceval. subst. apply E_Ass.
rewrite Heqv. reflexivity. Qed.
intros i a1 a2 Heqv st st'.
split; intros Hceval.
Case "→".
inversion Hceval. subst. apply E_Ass.
rewrite Heqv. reflexivity.
Case "←".
inversion Hceval. subst. apply E_Ass.
rewrite Heqv. reflexivity. Qed.
Theorem CWhile_congruence : ∀b1 b1' c1 c1',
bequiv b1 b1' → cequiv c1 c1' →
cequiv (WHILE b1 DO c1 END) (WHILE b1' DO c1' END).
Proof.
(* WORK IN CLASS *) Admitted.
Theorem CSeq_congruence : ∀c1 c1' c2 c2',
cequiv c1 c1' → cequiv c2 c2' →
cequiv (c1;;c2) (c1';;c2').
Proof.
(* FILL IN HERE *) Admitted.
(* FILL IN HERE *) Admitted.
Theorem CIf_congruence : ∀b b' c1 c1' c2 c2',
bequiv b b' → cequiv c1 c1' → cequiv c2 c2' →
cequiv (IFB b THEN c1 ELSE c2 FI) (IFB b' THEN c1' ELSE c2' FI).
Proof.
(* FILL IN HERE *) Admitted.
(* FILL IN HERE *) Admitted.
For example, here are two equivalent programs and a proof of their
    equivalence... 
Example congruence_example:
cequiv
(* Program 1: *)
(X ::= ANum 0;;
IFB (BEq (AId X) (ANum 0))
THEN
Y ::= ANum 0
ELSE
Y ::= ANum 42
FI)
(* Program 2: *)
(X ::= ANum 0;;
IFB (BEq (AId X) (ANum 0))
THEN
Y ::= AMinus (AId X) (AId X) (* <--- changed here *)
ELSE
Y ::= ANum 42
FI).
Proof.
apply CSeq_congruence.
apply refl_cequiv.
apply CIf_congruence.
apply refl_bequiv.
apply CAss_congruence. unfold aequiv. simpl.
symmetry. apply minus_diag.
apply refl_cequiv.
Qed.
Case Study: Constant Folding
Soundness of Program Transformations
Definition atrans_sound (atrans : aexp → aexp) : Prop :=
∀(a : aexp),
aequiv a (atrans a).
Definition btrans_sound (btrans : bexp → bexp) : Prop :=
∀(b : bexp),
bequiv b (btrans b).
Definition ctrans_sound (ctrans : com → com) : Prop :=
∀(c : com),
cequiv c (ctrans c).
The Constant-Folding Transformation
Fixpoint fold_constants_aexp (a : aexp) : aexp :=
match a with
| ANum n ⇒ ANum n
| AId i ⇒ AId i
| APlus a1 a2 ⇒
match (fold_constants_aexp a1, fold_constants_aexp a2) with
| (ANum n1, ANum n2) ⇒ ANum (n1 + n2)
| (a1', a2') ⇒ APlus a1' a2'
end
| AMinus a1 a2 ⇒
match (fold_constants_aexp a1, fold_constants_aexp a2) with
| (ANum n1, ANum n2) ⇒ ANum (n1 - n2)
| (a1', a2') ⇒ AMinus a1' a2'
end
| AMult a1 a2 ⇒
match (fold_constants_aexp a1, fold_constants_aexp a2) with
| (ANum n1, ANum n2) ⇒ ANum (n1 × n2)
| (a1', a2') ⇒ AMult a1' a2'
end
end.
Example fold_aexp_ex1 :
fold_constants_aexp
(AMult (APlus (ANum 1) (ANum 2)) (AId X))
= AMult (ANum 3) (AId X).
Proof. reflexivity. Qed.
Note that this version of constant folding doesn't eliminate
    trivial additions, etc. — we are focusing attention on a single
    optimization for the sake of simplicity.  It is not hard to
    incorporate other ways of simplifying expressions; the definitions
    and proofs just get longer. 
Example fold_aexp_ex2 :
fold_constants_aexp
(AMinus (AId X) (APlus (AMult (ANum 0) (ANum 6)) (AId Y)))
= AMinus (AId X) (APlus (ANum 0) (AId Y)).
Proof. reflexivity. Qed.
Not only can we lift fold_constants_aexp to bexps (in the
    BEq and BLe cases), we can also find constant boolean
    expressions and reduce them in-place. 
Fixpoint fold_constants_bexp (b : bexp) : bexp :=
match b with
| BTrue ⇒ BTrue
| BFalse ⇒ BFalse
| BEq a1 a2 ⇒
match (fold_constants_aexp a1, fold_constants_aexp a2) with
| (ANum n1, ANum n2) ⇒ if beq_nat n1 n2 then BTrue else BFalse
| (a1', a2') ⇒ BEq a1' a2'
end
| BLe a1 a2 ⇒
match (fold_constants_aexp a1, fold_constants_aexp a2) with
| (ANum n1, ANum n2) ⇒ if ble_nat n1 n2 then BTrue else BFalse
| (a1', a2') ⇒ BLe a1' a2'
end
| BNot b1 ⇒
match (fold_constants_bexp b1) with
| BTrue ⇒ BFalse
| BFalse ⇒ BTrue
| b1' ⇒ BNot b1'
end
| BAnd b1 b2 ⇒
match (fold_constants_bexp b1, fold_constants_bexp b2) with
| (BTrue, BTrue) ⇒ BTrue
| (BTrue, BFalse) ⇒ BFalse
| (BFalse, BTrue) ⇒ BFalse
| (BFalse, BFalse) ⇒ BFalse
| (b1', b2') ⇒ BAnd b1' b2'
end
end.
Example fold_bexp_ex1 :
fold_constants_bexp (BAnd BTrue (BNot (BAnd BFalse BTrue)))
= BTrue.
Proof. reflexivity. Qed.
Example fold_bexp_ex2 :
fold_constants_bexp
(BAnd (BEq (AId X) (AId Y))
(BEq (ANum 0)
(AMinus (ANum 2) (APlus (ANum 1) (ANum 1)))))
= BAnd (BEq (AId X) (AId Y)) BTrue.
Proof. reflexivity. Qed.
To fold constants in a command, we apply the appropriate folding
    functions on all embedded expressions. 
Fixpoint fold_constants_com (c : com) : com :=
match c with
| SKIP ⇒
SKIP
| i ::= a ⇒
CAss i (fold_constants_aexp a)
| c1 ;; c2 ⇒
(fold_constants_com c1) ;; (fold_constants_com c2)
| IFB b THEN c1 ELSE c2 FI ⇒
match fold_constants_bexp b with
| BTrue ⇒ fold_constants_com c1
| BFalse ⇒ fold_constants_com c2
| b' ⇒ IFB b' THEN fold_constants_com c1
ELSE fold_constants_com c2 FI
end
| WHILE b DO c END ⇒
match fold_constants_bexp b with
| BTrue ⇒ WHILE BTrue DO SKIP END
| BFalse ⇒ SKIP
| b' ⇒ WHILE b' DO (fold_constants_com c) END
end
end.
Example fold_com_ex1 :
fold_constants_com
(* Original program: *)
(X ::= APlus (ANum 4) (ANum 5);;
Y ::= AMinus (AId X) (ANum 3);;
IFB BEq (AMinus (AId X) (AId Y)) (APlus (ANum 2) (ANum 4)) THEN
SKIP
ELSE
Y ::= ANum 0
FI;;
IFB BLe (ANum 0) (AMinus (ANum 4) (APlus (ANum 2) (ANum 1))) THEN
Y ::= ANum 0
ELSE
SKIP
FI;;
WHILE BEq (AId Y) (ANum 0) DO
X ::= APlus (AId X) (ANum 1)
END)
= (* After constant folding: *)
(X ::= ANum 9;;
Y ::= AMinus (AId X) (ANum 3);;
IFB BEq (AMinus (AId X) (AId Y)) (ANum 6) THEN
SKIP
ELSE
(Y ::= ANum 0)
FI;;
Y ::= ANum 0;;
WHILE BEq (AId Y) (ANum 0) DO
X ::= APlus (AId X) (ANum 1)
END).
Proof. reflexivity. Qed.
Theorem fold_constants_aexp_sound :
atrans_sound fold_constants_aexp.
Proof.
unfold atrans_sound. intros a. unfold aequiv. intros st.
aexp_cases (induction a) Case; simpl;
(* ANum and AId follow immediately *)
try reflexivity;
(* APlus, AMinus, and AMult follow from the IH
and the observation that
aeval st (APlus a1 a2)
= ANum ((aeval st a1) + (aeval st a2))
= aeval st (ANum ((aeval st a1) + (aeval st a2)))
(and similarly for AMinus/minus and AMult/mult) *)
try (destruct (fold_constants_aexp a1);
destruct (fold_constants_aexp a2);
rewrite IHa1; rewrite IHa2; reflexivity). Qed.
unfold atrans_sound. intros a. unfold aequiv. intros st.
aexp_cases (induction a) Case; simpl;
(* ANum and AId follow immediately *)
try reflexivity;
(* APlus, AMinus, and AMult follow from the IH
and the observation that
aeval st (APlus a1 a2)
= ANum ((aeval st a1) + (aeval st a2))
= aeval st (ANum ((aeval st a1) + (aeval st a2)))
(and similarly for AMinus/minus and AMult/mult) *)
try (destruct (fold_constants_aexp a1);
destruct (fold_constants_aexp a2);
rewrite IHa1; rewrite IHa2; reflexivity). Qed.
Theorem fold_constants_bexp_sound:
btrans_sound fold_constants_bexp.
Proof.
unfold btrans_sound. intros b. unfold bequiv. intros st.
bexp_cases (induction b) Case;
(* BTrue and BFalse are immediate *)
try reflexivity.
Case "BEq".
(* Doing induction when there are a lot of constructors makes
specifying variable names a chore, but Coq doesn't always
choose nice variable names. We can rename entries in the
context with the rename tactic: rename a into a1 will
change a to a1 in the current goal and context. *)
rename a into a1. rename a0 into a2. simpl.
remember (fold_constants_aexp a1) as a1' eqn:Heqa1'.
remember (fold_constants_aexp a2) as a2' eqn:Heqa2'.
replace (aeval st a1) with (aeval st a1') by
(subst a1'; rewrite ← fold_constants_aexp_sound; reflexivity).
replace (aeval st a2) with (aeval st a2') by
(subst a2'; rewrite ← fold_constants_aexp_sound; reflexivity).
destruct a1'; destruct a2'; try reflexivity.
(* The only interesting case is when both a1 and a2
become constants after folding *)
simpl. destruct (beq_nat n n0); reflexivity.
Case "BLe".
(* FILL IN HERE *) admit.
Case "BNot".
simpl. remember (fold_constants_bexp b) as b' eqn:Heqb'.
rewrite IHb.
destruct b'; reflexivity.
Case "BAnd".
simpl.
remember (fold_constants_bexp b1) as b1' eqn:Heqb1'.
remember (fold_constants_bexp b2) as b2' eqn:Heqb2'.
rewrite IHb1. rewrite IHb2.
destruct b1'; destruct b2'; reflexivity. Qed.
unfold btrans_sound. intros b. unfold bequiv. intros st.
bexp_cases (induction b) Case;
(* BTrue and BFalse are immediate *)
try reflexivity.
Case "BEq".
(* Doing induction when there are a lot of constructors makes
specifying variable names a chore, but Coq doesn't always
choose nice variable names. We can rename entries in the
context with the rename tactic: rename a into a1 will
change a to a1 in the current goal and context. *)
rename a into a1. rename a0 into a2. simpl.
remember (fold_constants_aexp a1) as a1' eqn:Heqa1'.
remember (fold_constants_aexp a2) as a2' eqn:Heqa2'.
replace (aeval st a1) with (aeval st a1') by
(subst a1'; rewrite ← fold_constants_aexp_sound; reflexivity).
replace (aeval st a2) with (aeval st a2') by
(subst a2'; rewrite ← fold_constants_aexp_sound; reflexivity).
destruct a1'; destruct a2'; try reflexivity.
(* The only interesting case is when both a1 and a2
become constants after folding *)
simpl. destruct (beq_nat n n0); reflexivity.
Case "BLe".
(* FILL IN HERE *) admit.
Case "BNot".
simpl. remember (fold_constants_bexp b) as b' eqn:Heqb'.
rewrite IHb.
destruct b'; reflexivity.
Case "BAnd".
simpl.
remember (fold_constants_bexp b1) as b1' eqn:Heqb1'.
remember (fold_constants_bexp b2) as b2' eqn:Heqb2'.
rewrite IHb1. rewrite IHb2.
destruct b1'; destruct b2'; reflexivity. Qed.
Theorem fold_constants_com_sound :
ctrans_sound fold_constants_com.
Proof.
unfold ctrans_sound. intros c.
com_cases (induction c) Case; simpl.
Case "SKIP". apply refl_cequiv.
Case "::=". apply CAss_congruence. apply fold_constants_aexp_sound.
Case ";;". apply CSeq_congruence; assumption.
Case "IFB".
assert (bequiv b (fold_constants_bexp b)).
SCase "Pf of assertion". apply fold_constants_bexp_sound.
destruct (fold_constants_bexp b) eqn:Heqb;
(* If the optimization doesn't eliminate the if, then the result
is easy to prove from the IH and fold_constants_bexp_sound *)
try (apply CIf_congruence; assumption).
SCase "b always true".
apply trans_cequiv with c1; try assumption.
apply IFB_true; assumption.
SCase "b always false".
apply trans_cequiv with c2; try assumption.
apply IFB_false; assumption.
Case "WHILE".
(* FILL IN HERE *) Admitted.
unfold ctrans_sound. intros c.
com_cases (induction c) Case; simpl.
Case "SKIP". apply refl_cequiv.
Case "::=". apply CAss_congruence. apply fold_constants_aexp_sound.
Case ";;". apply CSeq_congruence; assumption.
Case "IFB".
assert (bequiv b (fold_constants_bexp b)).
SCase "Pf of assertion". apply fold_constants_bexp_sound.
destruct (fold_constants_bexp b) eqn:Heqb;
(* If the optimization doesn't eliminate the if, then the result
is easy to prove from the IH and fold_constants_bexp_sound *)
try (apply CIf_congruence; assumption).
SCase "b always true".
apply trans_cequiv with c1; try assumption.
apply IFB_true; assumption.
SCase "b always false".
apply trans_cequiv with c2; try assumption.
apply IFB_false; assumption.
Case "WHILE".
(* FILL IN HERE *) Admitted.
Proving That Programs Are Not Equivalent
       c1  =  (X ::= 42 + 53;; 
Y ::= Y + X)
c2 = (X ::= 42 + 53;;
Y ::= Y + (42 + 53)) 
    Clearly, this particular c1 and c2 are equivalent.  Is this
    true in general? 
Y ::= Y + X)
c2 = (X ::= 42 + 53;;
Y ::= Y + (42 + 53))
Fixpoint subst_aexp (i : id) (u : aexp) (a : aexp) : aexp :=
match a with
| ANum n ⇒ ANum n
| AId i' ⇒ if eq_id_dec i i' then u else AId i'
| APlus a1 a2 ⇒ APlus (subst_aexp i u a1) (subst_aexp i u a2)
| AMinus a1 a2 ⇒ AMinus (subst_aexp i u a1) (subst_aexp i u a2)
| AMult a1 a2 ⇒ AMult (subst_aexp i u a1) (subst_aexp i u a2)
end.
Example subst_aexp_ex :
subst_aexp X (APlus (ANum 42) (ANum 53)) (APlus (AId Y) (AId X)) =
(APlus (AId Y) (APlus (ANum 42) (ANum 53))).
Proof. reflexivity. Qed.
And here is the property we are interested in, expressing the
    claim that commands c1 and c2 as described above are
    always equivalent.  
Definition subst_equiv_property := ∀i1 i2 a1 a2,
cequiv (i1 ::= a1;; i2 ::= a2)
(i1 ::= a1;; i2 ::= subst_aexp i1 a1 a2).
Sadly, the property does not always hold. 
 
    Theorem: It is not the case that, for all i1, i2, a1,
    and a2,
 
 
 
 
 
    By our assumption, we know that
 
 
 
         cequiv (i1 ::= a1;; i2 ::= a2)
(i1 ::= a1;; i2 ::= subst_aexp i1 a1 a2). 
    Proof: Suppose, for a contradiction, that for all i1, i2,
    a1, and a2, we have
(i1 ::= a1;; i2 ::= subst_aexp i1 a1 a2).
      cequiv (i1 ::= a1;; i2 ::= a2) 
(i1 ::= a1;; i2 ::= subst_aexp i1 a1 a2). 
    Consider the following program:
(i1 ::= a1;; i2 ::= subst_aexp i1 a1 a2).
         X ::= APlus (AId X) (ANum 1);; Y ::= AId X
 
    Note that
         (X ::= APlus (AId X) (ANum 1);; Y ::= AId X)
/ empty_state ⇓ st1, 
    where st1 = { X ↦ 1, Y ↦ 1 }.
/ empty_state ⇓ st1,
        cequiv (X ::= APlus (AId X) (ANum 1);; Y ::= AId X)
(X ::= APlus (AId X) (ANum 1);; Y ::= APlus (AId X) (ANum 1)) 
    so, by the definition of cequiv, we have
(X ::= APlus (AId X) (ANum 1);; Y ::= APlus (AId X) (ANum 1))
        (X ::= APlus (AId X) (ANum 1);; Y ::= APlus (AId X) (ANum 1))
/ empty_state ⇓ st1. 
    But we can also derive
/ empty_state ⇓ st1.
        (X ::= APlus (AId X) (ANum 1);; Y ::= APlus (AId X) (ANum 1))
/ empty_state ⇓ st2, 
    where st2 = { X ↦ 1, Y ↦ 2 }.  Note that st1 ≠ st2; this
    is a contradiction, since ceval is deterministic!  ☐ 
/ empty_state ⇓ st2,
Theorem subst_inequiv :
¬ subst_equiv_property.
Proof.
unfold subst_equiv_property.
intros Contra.
(* Here is the counterexample: assuming that subst_equiv_property
holds allows us to prove that these two programs are
equivalent... *)
remember (X ::= APlus (AId X) (ANum 1);;
Y ::= AId X)
as c1.
remember (X ::= APlus (AId X) (ANum 1);;
Y ::= APlus (AId X) (ANum 1))
as c2.
assert (cequiv c1 c2) by (subst; apply Contra).
(* ... allows us to show that the command c2 can terminate
in two different final states:
st1 = {X |-> 1, Y |-> 1}
st2 = {X |-> 1, Y |-> 2}. *)
remember (update (update empty_state X 1) Y 1) as st1.
remember (update (update empty_state X 1) Y 2) as st2.
assert (H1: c1 / empty_state ⇓ st1);
assert (H2: c2 / empty_state ⇓ st2);
try (subst;
apply E_Seq with (st' := (update empty_state X 1));
apply E_Ass; reflexivity).
apply H in H1.
(* Finally, we use the fact that evaluation is deterministic
to obtain a contradiction. *)
assert (Hcontra: st1 = st2)
by (apply (ceval_deterministic c2 empty_state); assumption).
assert (Hcontra': st1 Y = st2 Y)
by (rewrite Hcontra; reflexivity).
subst. inversion Hcontra'. Qed.
unfold subst_equiv_property.
intros Contra.
(* Here is the counterexample: assuming that subst_equiv_property
holds allows us to prove that these two programs are
equivalent... *)
remember (X ::= APlus (AId X) (ANum 1);;
Y ::= AId X)
as c1.
remember (X ::= APlus (AId X) (ANum 1);;
Y ::= APlus (AId X) (ANum 1))
as c2.
assert (cequiv c1 c2) by (subst; apply Contra).
(* ... allows us to show that the command c2 can terminate
in two different final states:
st1 = {X |-> 1, Y |-> 1}
st2 = {X |-> 1, Y |-> 2}. *)
remember (update (update empty_state X 1) Y 1) as st1.
remember (update (update empty_state X 1) Y 2) as st2.
assert (H1: c1 / empty_state ⇓ st1);
assert (H2: c2 / empty_state ⇓ st2);
try (subst;
apply E_Seq with (st' := (update empty_state X 1));
apply E_Ass; reflexivity).
apply H in H1.
(* Finally, we use the fact that evaluation is deterministic
to obtain a contradiction. *)
assert (Hcontra: st1 = st2)
by (apply (ceval_deterministic c2 empty_state); assumption).
assert (Hcontra': st1 Y = st2 Y)
by (rewrite Hcontra; reflexivity).
subst. inversion Hcontra'. Qed.