ImpSimple Imperative Programs
     Z ::= X;; Y ::= 1;; WHILE not (Z = 0) DO Y ::= Y * Z;; Z ::= Z -
1 END 
1 END
(* IMPORTS *)
Set Warnings "-notation-overridden,-parsing".
Require Import Coq.Bool.Bool.
Require Import Coq.Arith.Arith.
Require Import Coq.Arith.EqNat.
Require Import Coq.omega.Omega.
Require Import Coq.Lists.List.
Require Import Coq.omega.Omega.
Import ListNotations.
Require Import Maps.
(* /IMPORTS *)
Arithmetic and Boolean Expressions
Module AExp.
These two definitions specify the abstract syntax of
    arithmetic and boolean expressions. 
Inductive aexp : Type :=
| ANum : nat → aexp
| APlus : aexp → aexp → aexp
| AMinus : aexp → aexp → aexp
| AMult : aexp → aexp → aexp.
Inductive bexp : Type :=
| BTrue : bexp
| BFalse : bexp
| BEq : aexp → aexp → bexp
| BLe : aexp → aexp → bexp
| BNot : bexp → bexp
| BAnd : bexp → bexp → bexp.
In this chapter, we'll elide the translation from the
    concrete syntax that a programmer would actually write to these
    abstract syntax trees — the process that, for example, would
    translate the string "1+2*3" to the AST
 
 
 For comparison, here's a conventional BNF (Backus-Naur Form)
    grammar defining the same abstract syntax:
 
 
 Compared to the Coq version above...
 
 
    It's good to be comfortable with both sorts of notations: informal
    ones for communicating between humans and formal ones for carrying
    out implementations and proofs. 
      APlus (ANum 1) (AMult (ANum 2) (ANum 3)).
 
    The optional chapter ImpParser develops a simple
    implementation of a lexical analyzer and parser that can perform
    this translation.  You do not need to understand that chapter to
    understand this one, but if you haven't taken a course where these
    techniques are covered (e.g., a compilers course) you may want to
    skim it. 
    a ::= nat
| a + a
| a - a
| a * a
b ::= true
| false
| a = a
| a ≤ a
| not b
| b and b 
| a + a
| a - a
| a * a
b ::= true
| false
| a = a
| a ≤ a
| not b
| b and b
-  The BNF is more informal — for example, it gives some
         suggestions about the surface syntax of expressions (like the
         fact that the addition operation is written + and is an
         infix symbol) while leaving other aspects of lexical analysis
         and parsing (like the relative precedence of +, -, and
         *, the use of parens to explicitly group subexpressions,
         etc.) unspecified.  Some additional information (and human
         intelligence) would be required to turn this description into
         a formal definition, for example when implementing a
         compiler.
The Coq version consistently omits all this information and concentrates on the abstract syntax only.
-  On the other hand, the BNF version is lighter and easier to
         read.  Its informality makes it flexible, a big advantage in
         situations like discussions at the blackboard, where
         conveying general ideas is more important than getting every
         detail nailed down precisely.
Indeed, there are dozens of BNF-like notations and people switch freely among them, usually without bothering to say which form of BNF they're using because there is no need to: a rough-and-ready informal understanding is all that's important.
Fixpoint aeval (a : aexp) : nat :=
match a with
| ANum n ⇒ n
| APlus a1 a2 ⇒ (aeval a1) + (aeval a2)
| AMinus a1 a2 ⇒ (aeval a1) - (aeval a2)
| AMult a1 a2 ⇒ (aeval a1) * (aeval a2)
end.
Example test_aeval1:
aeval (APlus (ANum 2) (ANum 2)) = 4.
Proof. reflexivity. Qed.
Similarly, evaluating a boolean expression yields a boolean. 
Fixpoint beval (b : bexp) : bool :=
match b with
| BTrue ⇒ true
| BFalse ⇒ false
| BEq a1 a2 ⇒ beq_nat (aeval a1) (aeval a2)
| BLe a1 a2 ⇒ leb (aeval a1) (aeval a2)
| BNot b1 ⇒ negb (beval b1)
| BAnd b1 b2 ⇒ andb (beval b1) (beval b2)
end.
Optimization
Fixpoint optimize_0plus (a:aexp) : aexp :=
match a with
| ANum n ⇒
ANum n
| APlus (ANum 0) e2 ⇒
optimize_0plus e2
| APlus e1 e2 ⇒
APlus (optimize_0plus e1) (optimize_0plus e2)
| AMinus e1 e2 ⇒
AMinus (optimize_0plus e1) (optimize_0plus e2)
| AMult e1 e2 ⇒
AMult (optimize_0plus e1) (optimize_0plus e2)
end.
To make sure our optimization is doing the right thing we
    can test it on some examples and see if the output looks OK. 
Example test_optimize_0plus:
optimize_0plus (APlus (ANum 2)
(APlus (ANum 0)
(APlus (ANum 0) (ANum 1))))
= APlus (ANum 2) (ANum 1).
Proof. reflexivity. Qed.
But if we want to be sure the optimization is correct —
    i.e., that evaluating an optimized expression gives the same
    result as the original — we should prove it. 
Theorem optimize_0plus_sound: ∀ a,
aeval (optimize_0plus a) = aeval a.
Proof.
intros a. induction a.
- (* ANum *) reflexivity.
- (* APlus *) destruct a1.
+ (* a1 = ANum n *) destruct n.
* (* n = 0 *) simpl. apply IHa2.
* (* n <> 0 *) simpl. rewrite IHa2. reflexivity.
+ (* a1 = APlus a1_1 a1_2 *)
simpl. simpl in IHa1. rewrite IHa1.
rewrite IHa2. reflexivity.
+ (* a1 = AMinus a1_1 a1_2 *)
simpl. simpl in IHa1. rewrite IHa1.
rewrite IHa2. reflexivity.
+ (* a1 = AMult a1_1 a1_2 *)
simpl. simpl in IHa1. rewrite IHa1.
rewrite IHa2. reflexivity.
- (* AMinus *)
simpl. rewrite IHa1. rewrite IHa2. reflexivity.
- (* AMult *)
simpl. rewrite IHa1. rewrite IHa2. reflexivity. Qed.
Coq Automation
Tacticals
The try Tactical
Theorem silly1 : ∀ ae, aeval ae = aeval ae.
Proof. try reflexivity. (* this just does reflexivity *) Qed.
Theorem silly2 : ∀ (P : Prop), P → P.
Proof.
intros P HP.
try reflexivity. (* just reflexivity would have failed *)
apply HP. (* we can still finish the proof in some other way *)
Qed.
There is no real reason to use try in completely manual
    proofs like these, but it is very useful for doing automated
    proofs in conjunction with the ; tactical, which we show
    next. 
The ; Tactical (Simple Form)
Lemma foo : ∀ n, leb 0 n = true.
Proof.
intros.
destruct n.
(* Leaves two subgoals, which are discharged identically... *)
- (* n=0 *) simpl. reflexivity.
- (* n=Sn' *) simpl. reflexivity.
Qed.
We can simplify this proof using the ; tactical: 
Lemma foo' : ∀ n, leb 0 n = true.
Proof.
intros.
(* destruct the current goal *)
destruct n;
(* then simpl each resulting subgoal *)
simpl;
(* and do reflexivity on each resulting subgoal *)
reflexivity.
Qed.
Using try and ; together, we can get rid of the repetition in
    the proof that was bothering us a little while ago. 
Theorem optimize_0plus_sound': ∀ a,
aeval (optimize_0plus a) = aeval a.
Proof.
intros a.
induction a;
(* Most cases follow directly by the IH... *)
try (simpl; rewrite IHa1; rewrite IHa2; reflexivity).
(* ... but the remaining cases -- ANum and APlus --
are different: *)
- (* ANum *) reflexivity.
- (* APlus *)
destruct a1;
(* Again, most cases follow directly by the IH: *)
try (simpl; simpl in IHa1; rewrite IHa1;
rewrite IHa2; reflexivity).
(* The interesting case, on which the try...
does nothing, is when e1 = ANum n. In this
case, we have to destruct n (to see whether
the optimization applies) and rewrite with the
induction hypothesis. *)
+ (* a1 = ANum n *) destruct n;
simpl; rewrite IHa2; reflexivity. Qed.
Coq experts often use this "...; try... " idiom after a tactic
    like induction to take care of many similar cases all at once.
    Naturally, this practice has an analog in informal proofs.  For
    example, here is an informal proof of the optimization theorem
    that matches the structure of the formal one:
 
    Theorem: For all arithmetic expressions a,
 
 
 
 However, this proof can still be improved: the first case (for
    a = ANum n) is very trivial — even more trivial than the cases
    that we said simply followed from the IH — yet we have chosen to
    write it out in full.  It would be better and clearer to drop it
    and just say, at the top, "Most cases are either immediate or
    direct from the IH.  The only interesting case is the one for
    APlus..."  We can make the same improvement in our formal proof
    too.  Here's how it looks: 
       aeval (optimize_0plus a) = aeval a.
 
    Proof: By induction on a.  Most cases follow directly from the
    IH.  The remaining cases are as follows:
-  Suppose a = ANum n for some n.  We must show
aeval (optimize_0plus (ANum n)) = aeval (ANum n).This is immediate from the definition of optimize_0plus.
-  Suppose a = APlus a1 a2 for some a1 and a2.  We must
        show
aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2).Consider the possible forms of a1. For most of them, optimize_0plus simply calls itself recursively for the subexpressions and rebuilds a new expression of the same form as a1; in these cases, the result follows directly from the IH.The interesting case is when a1 = ANum n for some n. If n = ANum 0, thenoptimize_0plus (APlus a1 a2) = optimize_0plus a2and the IH for a2 is exactly what we need. On the other hand, if n = S n' for some n', then again optimize_0plus simply calls itself recursively, and the result follows from the IH. ☐
Theorem optimize_0plus_sound'': ∀ a,
aeval (optimize_0plus a) = aeval a.
Proof.
intros a.
induction a;
(* Most cases follow directly by the IH *)
try (simpl; rewrite IHa1; rewrite IHa2; reflexivity);
(* ... or are immediate by definition *)
try reflexivity.
(* The interesting case is when a = APlus a1 a2. *)
- (* APlus *)
destruct a1; try (simpl; simpl in IHa1; rewrite IHa1;
rewrite IHa2; reflexivity).
+ (* a1 = ANum n *) destruct n;
simpl; rewrite IHa2; reflexivity. Qed.
The ; Tactical (General Form)
      T; [T1 | T2 | ... | Tn]
 
   is a tactic that first performs T and then performs T1 on the
   first subgoal generated by T, performs T2 on the second
   subgoal, etc.
      T; [T' | T' | ... | T']
 
The repeat Tactical
Theorem In10 : In 10 [1;2;3;4;5;6;7;8;9;10].
Proof.
repeat (try (left; reflexivity); right).
Qed.
The tactic repeat T never fails: if the tactic T doesn't apply
    to the original goal, then repeat still succeeds without changing
    the original goal (i.e., it repeats zero times). 
Theorem In10' : In 10 [1;2;3;4;5;6;7;8;9;10].
Proof.
repeat (left; reflexivity).
repeat (right; try (left; reflexivity)).
Qed.
The tactic repeat T also does not have any upper bound on the
    number of times it applies T.  If T is a tactic that always
    succeeds, then repeat T will loop forever (e.g., repeat simpl
    loops forever, since simpl always succeeds).  While evaluation
    in Coq's term language, Gallina, is guaranteed to terminate,
    tactic evaluation is not!  This does not affect Coq's logical
    consistency, however, since the job of repeat and other tactics
    is to guide Coq in constructing proofs; if the construction
    process diverges, this simply means that we have failed to
    construct a proof, not that we have constructed a wrong one. 
 
Exercise: 3 stars (optimize_0plus_b)
Since the optimize_0plus transformation doesn't change the value of aexps, we should be able to apply it to all the aexps that appear in a bexp without changing the bexp's value. Write a function which performs that transformation on bexps, and prove it is sound. Use the tacticals we've just seen to make the proof as elegant as possible.Fixpoint optimize_0plus_b (b : bexp) : bexp
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Theorem optimize_0plus_b_sound : ∀ b,
beval (optimize_0plus_b b) = beval b.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 4 stars, optional (optimizer)
Design exercise: The optimization implemented by our optimize_0plus function is only one of many possible optimizations on arithmetic and boolean expressions. Write a more sophisticated optimizer and prove it correct. (You will probably find it easiest to start small — add just a single, simple optimization and prove it correct — and build up to something more interesting incrementially.)☐
Defining New Tactic Notations
-  The Tactic Notation idiom illustrated below gives a handy way to
      define "shorthand tactics" that bundle several tactics into a
      single command.
-  For more sophisticated programming, Coq offers a built-in
      programming language called Ltac with primitives that can
      examine and modify the proof state.  The details are a bit too
      complicated to get into here (and it is generally agreed that
      Ltac is not the most beautiful part of Coq's design!), but they
      can be found in the reference manual and other books on Coq, and
      there are many examples of Ltac definitions in the Coq standard
      library that you can use as examples.
- There is also an OCaml API, which can be used to build tactics that access Coq's internal structures at a lower level, but this is seldom worth the trouble for ordinary Coq users.
Tactic Notation "simpl_and_try" tactic(c) :=
simpl;
try c.
This defines a new tactical called simpl_and_try that takes one
    tactic c as an argument and is defined to be equivalent to the
    tactic simpl; try c.  Now writing "simpl_and_try reflexivity."
    in a proof will be the same as writing "simpl; try
    reflexivity." 
The omega Tactic
-  numeric constants, addition (+ and S), subtraction (-
        and pred), and multiplication by constants (this is what
        makes it Presburger arithmetic),
-  equality (= and ≠) and inequality (≤), and
- the logical connectives ∧, ∨, ¬, and →,
Example silly_presburger_example : ∀ m n o p,
m + n ≤ n + o ∧ o + 3 = p + 3 →
m ≤ p.
Proof.
intros. omega.
Qed.
(Note the Require Import Coq.omega.Omega. at the top of the file.) 
A Few More Handy Tactics
-  clear H: Delete hypothesis H from the context.
-  subst x: Find an assumption x = e or e = x in the
       context, replace x with e throughout the context and
       current goal, and clear the assumption.
-  subst: Substitute away all assumptions of the form x = e
       or e = x.
-  rename... into...: Change the name of a hypothesis in the
       proof context.  For example, if the context includes a variable
       named x, then rename x into y will change all occurrences
       of x to y.
-  assumption: Try to find a hypothesis H in the context that
       exactly matches the goal; if one is found, behave like apply
       H.
-  contradiction: Try to find a hypothesis H in the current
       context that is logically equivalent to False.  If one is
       found, solve the goal.
- constructor: Try to find a constructor c (from some Inductive definition in the current environment) that can be applied to solve the current goal. If one is found, behave like apply c.
Evaluation as a Relation
Module aevalR_first_try.
Inductive aevalR : aexp → nat → Prop :=
| E_ANum : ∀ (n: nat),
aevalR (ANum n) n
| E_APlus : ∀ (e1 e2: aexp) (n1 n2: nat),
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (APlus e1 e2) (n1 + n2)
| E_AMinus: ∀ (e1 e2: aexp) (n1 n2: nat),
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (AMinus e1 e2) (n1 - n2)
| E_AMult : ∀ (e1 e2: aexp) (n1 n2: nat),
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (AMult e1 e2) (n1 * n2).
It will be convenient to have an infix notation for
    aevalR.  We'll write e \\ n to mean that arithmetic expression
    e evaluates to value n. 
Notation "e '\\' n"
:= (aevalR e n)
(at level 50, left associativity)
: type_scope.
End aevalR_first_try.
In fact, Coq provides a way to use this notation in the
    definition of aevalR itself.  This reduces confusion by avoiding
    situations where we're working on a proof involving statements in
    the form e \\ n but we have to refer back to a definition
    written using the form aevalR e n.
 
    We do this by first "reserving" the notation, then giving the
    definition together with a declaration of what the notation
    means. 
Reserved Notation "e '\\' n" (at level 50, left associativity).
Inductive aevalR : aexp → nat → Prop :=
| E_ANum : ∀ (n:nat),
(ANum n) \\ n
| E_APlus : ∀ (e1 e2: aexp) (n1 n2 : nat),
(e1 \\ n1) → (e2 \\ n2) → (APlus e1 e2) \\ (n1 + n2)
| E_AMinus : ∀ (e1 e2: aexp) (n1 n2 : nat),
(e1 \\ n1) → (e2 \\ n2) → (AMinus e1 e2) \\ (n1 - n2)
| E_AMult : ∀ (e1 e2: aexp) (n1 n2 : nat),
(e1 \\ n1) → (e2 \\ n2) → (AMult e1 e2) \\ (n1 * n2)
where "e '\\' n" := (aevalR e n) : type_scope.
Inference Rule Notation
      | E_APlus : ∀ (e1 e2: aexp) (n1 n2: nat),
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (APlus e1 e2) (n1 + n2) 
    ...would be written like this as an inference rule:
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (APlus e1 e2) (n1 + n2)
| e1 \\ n1 | |
| e2 \\ n2 | (E_APlus) | 
| APlus e1 e2 \\ n1+n2 | 
| (E_ANum) | |
| ANum n \\ n | 
| e1 \\ n1 | |
| e2 \\ n2 | (E_APlus) | 
| APlus e1 e2 \\ n1+n2 | 
| e1 \\ n1 | |
| e2 \\ n2 | (E_AMinus) | 
| AMinus e1 e2 \\ n1-n2 | 
| e1 \\ n1 | |
| e2 \\ n2 | (E_AMult) | 
| AMult e1 e2 \\ n1*n2 | 
Equivalence of the Definitions
Theorem aeval_iff_aevalR : ∀ a n,
(a \\ n) ↔ aeval a = n.
Proof.
split.
- (* -> *)
intros H.
induction H; simpl.
+ (* E_ANum *)
reflexivity.
+ (* E_APlus *)
rewrite IHaevalR1. rewrite IHaevalR2. reflexivity.
+ (* E_AMinus *)
rewrite IHaevalR1. rewrite IHaevalR2. reflexivity.
+ (* E_AMult *)
rewrite IHaevalR1. rewrite IHaevalR2. reflexivity.
- (* <- *)
generalize dependent n.
induction a;
simpl; intros; subst.
+ (* ANum *)
apply E_ANum.
+ (* APlus *)
apply E_APlus.
apply IHa1. reflexivity.
apply IHa2. reflexivity.
+ (* AMinus *)
apply E_AMinus.
apply IHa1. reflexivity.
apply IHa2. reflexivity.
+ (* AMult *)
apply E_AMult.
apply IHa1. reflexivity.
apply IHa2. reflexivity.
Qed.
split.
- (* -> *)
intros H.
induction H; simpl.
+ (* E_ANum *)
reflexivity.
+ (* E_APlus *)
rewrite IHaevalR1. rewrite IHaevalR2. reflexivity.
+ (* E_AMinus *)
rewrite IHaevalR1. rewrite IHaevalR2. reflexivity.
+ (* E_AMult *)
rewrite IHaevalR1. rewrite IHaevalR2. reflexivity.
- (* <- *)
generalize dependent n.
induction a;
simpl; intros; subst.
+ (* ANum *)
apply E_ANum.
+ (* APlus *)
apply E_APlus.
apply IHa1. reflexivity.
apply IHa2. reflexivity.
+ (* AMinus *)
apply E_AMinus.
apply IHa1. reflexivity.
apply IHa2. reflexivity.
+ (* AMult *)
apply E_AMult.
apply IHa1. reflexivity.
apply IHa2. reflexivity.
Qed.
We can make the proof quite a bit shorter by making more
    use of tacticals. 
Theorem aeval_iff_aevalR' : ∀ a n,
(a \\ n) ↔ aeval a = n.
Proof.
(* WORKED IN CLASS *)
split.
- (* -> *)
intros H; induction H; subst; reflexivity.
- (* <- *)
generalize dependent n.
induction a; simpl; intros; subst; constructor;
try apply IHa1; try apply IHa2; reflexivity.
Qed.
Exercise: 3 stars (bevalR)
Write a relation bevalR in the same style as aevalR, and prove that it is equivalent to beval.Inductive bevalR: bexp → bool → Prop :=
(* FILL IN HERE *)
.
Lemma beval_iff_bevalR : ∀ b bv,
bevalR b bv ↔ beval b = bv.
Proof.
(* FILL IN HERE *) Admitted.
End AExp.
Computational vs. Relational Definitions
Module aevalR_division.
For example, suppose that we wanted to extend the arithmetic
    operations by considering also a division operation:
Inductive aexp : Type :=
| ANum : nat → aexp
| APlus : aexp → aexp → aexp
| AMinus : aexp → aexp → aexp
| AMult : aexp → aexp → aexp
| ADiv : aexp → aexp → aexp. (* <--- new *)
Extending the definition of aeval to handle this new operation
    would not be straightforward (what should we return as the result
    of ADiv (ANum 5) (ANum 0)?).  But extending aevalR is
    straightforward. 
Reserved Notation "e '\\' n"
(at level 50, left associativity).
Inductive aevalR : aexp → nat → Prop :=
| E_ANum : ∀ (n:nat),
(ANum n) \\ n
| E_APlus : ∀ (a1 a2: aexp) (n1 n2 : nat),
(a1 \\ n1) → (a2 \\ n2) → (APlus a1 a2) \\ (n1 + n2)
| E_AMinus : ∀ (a1 a2: aexp) (n1 n2 : nat),
(a1 \\ n1) → (a2 \\ n2) → (AMinus a1 a2) \\ (n1 - n2)
| E_AMult : ∀ (a1 a2: aexp) (n1 n2 : nat),
(a1 \\ n1) → (a2 \\ n2) → (AMult a1 a2) \\ (n1 * n2)
| E_ADiv : ∀ (a1 a2: aexp) (n1 n2 n3: nat),
(a1 \\ n1) → (a2 \\ n2) → (n2 > 0) →
(mult n2 n3 = n1) → (ADiv a1 a2) \\ n3
where "a '\\' n" := (aevalR a n) : type_scope.
End aevalR_division.
Module aevalR_extended.
Suppose, instead, that we want to extend the arithmetic operations
    by a nondeterministic number generator any that, when evaluated,
    may yield any number.  (Note that this is not the same as making a
    probabilistic choice among all possible numbers — we're not
    specifying any particular distribution of results, but just saying
    what results are possible.) 
Reserved Notation "e '\\' n" (at level 50, left associativity).
Inductive aexp : Type :=
| AAny : aexp (* <--- NEW *)
| ANum : nat → aexp
| APlus : aexp → aexp → aexp
| AMinus : aexp → aexp → aexp
| AMult : aexp → aexp → aexp.
Again, extending aeval would be tricky, since now evaluation is
    not a deterministic function from expressions to numbers, but
    extending aevalR is no problem: 
Inductive aevalR : aexp → nat → Prop :=
| E_Any : ∀ (n:nat),
AAny \\ n (* <--- new *)
| E_ANum : ∀ (n:nat),
(ANum n) \\ n
| E_APlus : ∀ (a1 a2: aexp) (n1 n2 : nat),
(a1 \\ n1) → (a2 \\ n2) → (APlus a1 a2) \\ (n1 + n2)
| E_AMinus : ∀ (a1 a2: aexp) (n1 n2 : nat),
(a1 \\ n1) → (a2 \\ n2) → (AMinus a1 a2) \\ (n1 - n2)
| E_AMult : ∀ (a1 a2: aexp) (n1 n2 : nat),
(a1 \\ n1) → (a2 \\ n2) → (AMult a1 a2) \\ (n1 * n2)
where "a '\\' n" := (aevalR a n) : type_scope.
End aevalR_extended.
At this point you maybe wondering: which style should I use by
    default?  The examples above show that relational definitions are
    fundamentally more powerful than functional ones.  For situations
    like these, where the thing being defined is not easy to express
    as a function, or indeed where it is not a function, there is no
    choice.  But what about when both styles are workable?
 
    One point in favor of relational definitions is that some people
    feel they are more elegant and easier to understand.  Another is
    that Coq automatically generates nice inversion and induction
    principles from Inductive definitions.
 
    On the other hand, functional definitions can often be more
    convenient:
 
 
    Furthermore, functions can be directly "extracted" to executable
    code in OCaml or Haskell.
 
    Ultimately, the choice often comes down to either the specifics of
    a particular situation or simply a question of taste.  Indeed, in
    large Coq developments it is common to see a definition given in
    both functional and relational styles, plus a lemma stating that
    the two coincide, allowing further proofs to switch from one point
    of view to the other at will.
- Functions are by definition deterministic and defined on all arguments; for a relation we have to show these properties explicitly if we need them.
- With functions we can also take advantage of Coq's computation mechanism to simplify expressions during proofs.
Expressions With Variables
States
Definition state := total_map nat.
Definition empty_state : state :=
t_empty 0.
Syntax
Inductive aexp : Type :=
| ANum : nat → aexp
| AId : id → aexp (* <----- NEW *)
| APlus : aexp → aexp → aexp
| AMinus : aexp → aexp → aexp
| AMult : aexp → aexp → aexp.
Defining a few variable names as notational shorthands will make
    examples easier to read: 
Definition W : id := Id "W".
Definition X : id := Id "X".
Definition Y : id := Id "Y".
Definition Z : id := Id "Z".
(This convention for naming program variables (X, Y,
    Z) clashes a bit with our earlier use of uppercase letters for
    types.  Since we're not using polymorphism heavily in the chapters
    devoped to Imp, this overloading should not cause confusion.) 
 
 The definition of bexps is unchanged (except for using the new
    aexps): 
Inductive bexp : Type :=
| BTrue : bexp
| BFalse : bexp
| BEq : aexp → aexp → bexp
| BLe : aexp → aexp → bexp
| BNot : bexp → bexp
| BAnd : bexp → bexp → bexp.
Evaluation
Fixpoint aeval (st : state) (a : aexp) : nat :=
match a with
| ANum n ⇒ n
| AId x ⇒ st x (* <----- NEW *)
| APlus a1 a2 ⇒ (aeval st a1) + (aeval st a2)
| AMinus a1 a2 ⇒ (aeval st a1) - (aeval st a2)
| AMult a1 a2 ⇒ (aeval st a1) * (aeval st a2)
end.
Fixpoint beval (st : state) (b : bexp) : bool :=
match b with
| BTrue ⇒ true
| BFalse ⇒ false
| BEq a1 a2 ⇒ beq_nat (aeval st a1) (aeval st a2)
| BLe a1 a2 ⇒ leb (aeval st a1) (aeval st a2)
| BNot b1 ⇒ negb (beval st b1)
| BAnd b1 b2 ⇒ andb (beval st b1) (beval st b2)
end.
Example aexp1 :
aeval (t_update empty_state X 5)
(APlus (ANum 3) (AMult (AId X) (ANum 2)))
= 13.
Proof. reflexivity. Qed.
Example bexp1 :
beval (t_update empty_state X 5)
(BAnd BTrue (BNot (BLe (AId X) (ANum 4))))
= true.
Proof. reflexivity. Qed.
Commands
Syntax
     c ::= SKIP | x ::= a | c ;; c | IFB b THEN c ELSE c FI
| WHILE b DO c END 
| WHILE b DO c END
     Z ::= X;;
Y ::= 1;;
WHILE not (Z = 0) DO
Y ::= Y * Z;;
Z ::= Z - 1
END 
   When this command terminates, the variable Y will contain the
   factorial of the initial value of X. 
Y ::= 1;;
WHILE not (Z = 0) DO
Y ::= Y * Z;;
Z ::= Z - 1
END
Inductive com : Type :=
| CSkip : com
| CAss : id → aexp → com
| CSeq : com → com → com
| CIf : bexp → com → com → com
| CWhile : bexp → com → com.
As usual, we can use a few Notation declarations to make things
    more readable.  To avoid conflicts with Coq's built-in notations,
    we keep this light — in particular, we don't introduce any
    notations for aexps and bexps to avoid confusion with the
    numeric and boolean operators we've already defined. 
Notation "'SKIP'" :=
CSkip.
Notation "x '::=' a" :=
(CAss x a) (at level 60).
Notation "c1 ;; c2" :=
(CSeq c1 c2) (at level 80, right associativity).
Notation "'WHILE' b 'DO' c 'END'" :=
(CWhile b c) (at level 80, right associativity).
Notation "'IFB' c1 'THEN' c2 'ELSE' c3 'FI'" :=
(CIf c1 c2 c3) (at level 80, right associativity).
For example, here is the factorial function again, written as a
    formal definition to Coq: 
Definition fact_in_coq : com :=
Z ::= AId X;;
Y ::= ANum 1;;
WHILE BNot (BEq (AId Z) (ANum 0)) DO
Y ::= AMult (AId Y) (AId Z);;
Z ::= AMinus (AId Z) (ANum 1)
END.
Definition plus2 : com :=
X ::= (APlus (AId X) (ANum 2)).
Definition XtimesYinZ : com :=
Z ::= (AMult (AId X) (AId Y)).
Definition subtract_slowly_body : com :=
Z ::= AMinus (AId Z) (ANum 1) ;;
X ::= AMinus (AId X) (ANum 1).
Definition subtract_slowly : com :=
WHILE BNot (BEq (AId X) (ANum 0)) DO
subtract_slowly_body
END.
Definition subtract_3_from_5_slowly : com :=
X ::= ANum 3 ;;
Z ::= ANum 5 ;;
subtract_slowly.
Definition loop : com :=
WHILE BTrue DO
SKIP
END.
Evaluating Commands
Evaluation as a Function (Failed Attempt)
Fixpoint ceval_fun_no_while (st : state) (c : com)
: state :=
match c with
| SKIP ⇒
st
| x ::= a1 ⇒
t_update st x (aeval st a1)
| c1 ;; c2 ⇒
let st' := ceval_fun_no_while st c1 in
ceval_fun_no_while st' c2
| IFB b THEN c1 ELSE c2 FI ⇒
if (beval st b)
then ceval_fun_no_while st c1
else ceval_fun_no_while st c2
| WHILE b DO c END ⇒
st (* bogus *)
end.
In a traditional functional programming language like OCaml or
    Haskell we could add the WHILE case as follows:
 
    Thus, because it doesn't terminate on all inputs,
    of ceval_fun cannot be written in Coq — at least not without
    additional tricks and workarounds (see chapter ImpCEvalFun
    if you're curious about what those might be). 
  Fixpoint ceval_fun (st : state) (c : com) : state :=
    match c with
      ...
      | WHILE b DO c END =>
          if (beval st b)
            then ceval_fun st (c; WHILE b DO c END)
            else st
    end.
    Coq 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, it doesn't always
    terminate: for example, the full version of the ceval_fun
    function applied to the loop program above 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!) 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
    (loop_false 0 would be a proof of False), which
    would be a disaster for Coq's logical consistency.
Evaluation as a Relation
Operational Semantics
| (E_Skip) | |
| SKIP / st \\ st | 
| aeval st a1 = n | (E_Ass) | 
| x := a1 / st \\ (t_update st x n) | 
| c1 / st \\ st' | |
| c2 / st' \\ st'' | (E_Seq) | 
| c1;;c2 / st \\ st'' | 
| beval st b1 = true | |
| c1 / st \\ st' | (E_IfTrue) | 
| IF b1 THEN c1 ELSE c2 FI / st \\ st' | 
| beval st b1 = false | |
| c2 / st \\ st' | (E_IfFalse) | 
| IF b1 THEN c1 ELSE c2 FI / st \\ st' | 
| beval st b = false | (E_WhileFalse) | 
| WHILE b DO c END / st \\ st | 
| beval st b = true | |
| c / st \\ st' | |
| WHILE b DO c END / st' \\ st'' | (E_WhileTrue) | 
| WHILE b DO c END / st \\ st'' | 
Reserved Notation "c1 '/' st '\\' st'"
(at level 40, st at level 39).
Inductive ceval : com → state → state → Prop :=
| E_Skip : ∀ st,
SKIP / st \\ st
| E_Ass : ∀ st a1 n x,
aeval st a1 = n →
(x ::= a1) / st \\ (t_update st x n)
| E_Seq : ∀ c1 c2 st st' st'',
c1 / st \\ st' →
c2 / st' \\ st'' →
(c1 ;; c2) / st \\ st''
| E_IfTrue : ∀ st st' b c1 c2,
beval st b = true →
c1 / st \\ st' →
(IFB b THEN c1 ELSE c2 FI) / st \\ st'
| E_IfFalse : ∀ st st' b c1 c2,
beval st b = false →
c2 / st \\ st' →
(IFB b THEN c1 ELSE c2 FI) / st \\ st'
| E_WhileFalse : ∀ b st c,
beval st b = false →
(WHILE b DO c END) / st \\ st
| E_WhileTrue : ∀ st st' st'' b c,
beval st b = true →
c / st \\ st' →
(WHILE b DO c END) / st' \\ st'' →
(WHILE b DO c END) / st \\ st''
where "c1 '/' st '\\' st'" := (ceval c1 st st').
The cost of defining evaluation as a relation instead of a
    function is that we now need to construct proofs that some
    program evaluates to some result state, rather than just letting
    Coq's computation mechanism do it for us. 
Example ceval_example1:
(X ::= ANum 2;;
IFB BLe (AId X) (ANum 1)
THEN Y ::= ANum 3
ELSE Z ::= ANum 4
FI)
/ empty_state
\\ (t_update (t_update empty_state X 2) Z 4).
Proof.
(* We must supply the intermediate state *)
apply E_Seq with (t_update empty_state X 2).
- (* assignment command *)
apply E_Ass. reflexivity.
- (* if command *)
apply E_IfFalse.
reflexivity.
apply E_Ass. reflexivity. Qed.
Example ceval_example2:
(X ::= ANum 0;; Y ::= ANum 1;; Z ::= ANum 2) / empty_state \\
(t_update (t_update (t_update empty_state X 0) Y 1) Z 2).
Proof.
(* FILL IN HERE *) Admitted.
☐ 
(X ::= ANum 0;; Y ::= ANum 1;; Z ::= ANum 2) / empty_state \\
(t_update (t_update (t_update empty_state X 0) Y 1) Z 2).
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, advanced (pup_to_n)
Write an Imp program that sums the numbers from 1 to X (inclusive: 1 + 2 + ... + X) in the variable Y. Prove that this program executes as intended for X = 2 (this is trickier than you might expect).Definition pup_to_n : com
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Theorem pup_to_2_ceval :
pup_to_n / (t_update empty_state X 2) \\
t_update (t_update (t_update (t_update (t_update (t_update empty_state
X 2) Y 0) Y 2) X 1) Y 3) X 0.
Proof.
(* FILL IN HERE *) Admitted.
Determinism of Evaluation
Theorem ceval_deterministic: ∀ c st st1 st2,
c / st \\ st1 →
c / st \\ st2 →
st1 = st2.
Proof.
intros c st st1 st2 E1 E2.
generalize dependent st2.
induction E1;
intros st2 E2; inversion E2; subst.
- (* E_Skip *) reflexivity.
- (* E_Ass *) reflexivity.
- (* E_Seq *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; assumption. }
subst st'0.
apply IHE1_2. assumption.
- (* E_IfTrue, b1 evaluates to true *)
apply IHE1. assumption.
- (* E_IfTrue, b1 evaluates to false (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_IfFalse, b1 evaluates to true (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_IfFalse, b1 evaluates to false *)
apply IHE1. assumption.
- (* E_WhileFalse, b1 evaluates to false *)
reflexivity.
- (* E_WhileFalse, b1 evaluates to true (contradiction) *)
rewrite H in H2. inversion H2.
- (* E_WhileTrue, b1 evaluates to false (contradiction) *)
rewrite H in H4. inversion H4.
- (* E_WhileTrue, b1 evaluates to true *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; assumption. }
subst st'0.
apply IHE1_2. assumption. Qed.
intros c st st1 st2 E1 E2.
generalize dependent st2.
induction E1;
intros st2 E2; inversion E2; subst.
- (* E_Skip *) reflexivity.
- (* E_Ass *) reflexivity.
- (* E_Seq *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; assumption. }
subst st'0.
apply IHE1_2. assumption.
- (* E_IfTrue, b1 evaluates to true *)
apply IHE1. assumption.
- (* E_IfTrue, b1 evaluates to false (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_IfFalse, b1 evaluates to true (contradiction) *)
rewrite H in H5. inversion H5.
- (* E_IfFalse, b1 evaluates to false *)
apply IHE1. assumption.
- (* E_WhileFalse, b1 evaluates to false *)
reflexivity.
- (* E_WhileFalse, b1 evaluates to true (contradiction) *)
rewrite H in H2. inversion H2.
- (* E_WhileTrue, b1 evaluates to false (contradiction) *)
rewrite H in H4. inversion H4.
- (* E_WhileTrue, b1 evaluates to true *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; assumption. }
subst st'0.
apply IHE1_2. assumption. Qed.
Reasoning About Imp Programs
Theorem plus2_spec : ∀ st n st',
st X = n →
plus2 / st \\ st' →
st' X = n + 2.
Proof.
intros st n st' HX Heval.
Inverting Heval essentially forces Coq to expand one step of
      the ceval computation — in this case revealing that st'
      must be st extended with the new value of X, since plus2
      is an assignment 
inversion Heval. subst. clear Heval. simpl.
apply t_update_eq. Qed.
(* FILL IN HERE *)
Theorem loop_never_stops : ∀ st st',
~(loop / st \\ st').
Proof.
intros st st' contra. unfold loop in contra.
remember (WHILE BTrue DO SKIP END) as loopdef
eqn:Heqloopdef.
~(loop / st \\ st').
Proof.
intros st st' contra. unfold loop in contra.
remember (WHILE BTrue DO SKIP END) as loopdef
eqn:Heqloopdef.
Proceed by induction on the assumed derivation showing that
      loopdef terminates.  Most of the cases are immediately
      contradictory (and so can be solved in one step with
      inversion). 
(* FILL IN HERE *) Admitted.
Fixpoint no_whiles (c : com) : bool :=
match c with
| SKIP ⇒
true
| _ ::= _ ⇒
true
| c1 ;; c2 ⇒
andb (no_whiles c1) (no_whiles c2)
| IFB _ THEN ct ELSE cf FI ⇒
andb (no_whiles ct) (no_whiles cf)
| WHILE _ DO _ END ⇒
false
end.
This predicate yields true just on programs that have no while
    loops.  Using Inductive, write a property no_whilesR such that
    no_whilesR c is provable exactly when c is a program with no
    while loops.  Then prove its equivalence with no_whiles. 
Inductive no_whilesR: com → Prop :=
(* FILL IN HERE *)
.
Theorem no_whiles_eqv:
∀ c, no_whiles c = true ↔ no_whilesR c.
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 4 starsM (no_whiles_terminating)
Imp programs that don't involve while loops always terminate. State and prove a theorem no_whiles_terminating that says this. Use either no_whiles or no_whilesR, as you prefer.(* FILL IN HERE *)
Additional Exercises
Exercise: 3 stars (stack_compiler)
HP Calculators, programming languages like Forth and Postscript and abstract machines like the Java Virtual Machine all evaluate arithmetic expressions using a stack. For instance, the expression
      (2*3)+(3*(4-2))
   would be entered as
      2 3 * 3 4 2 - * +
   and evaluated like this (where we show the program being evaluated
   on the right and the contents of the stack on the left):
      [ ]           |    2 3 * 3 4 2 - * +
      [2]           |    3 * 3 4 2 - * +
      [3, 2]        |    * 3 4 2 - * +
      [6]           |    3 4 2 - * +
      [3, 6]        |    4 2 - * +
      [4, 3, 6]     |    2 - * +
      [2, 4, 3, 6]  |    - * +
      [2, 3, 6]     |    * +
      [6, 6]        |    +
      [12]          |
  The task of this exercise is to write a small compiler that
  translates aexps into stack machine instructions.
- SPush n: Push the number n on the stack.
- SLoad x: Load the identifier x from the store and push it on the stack
- SPlus: Pop the two top numbers from the stack, add them, and push the result onto the stack.
- SMinus: Similar, but subtract.
- SMult: Similar, but multiply.
Inductive sinstr : Type :=
| SPush : nat → sinstr
| SLoad : id → sinstr
| SPlus : sinstr
| SMinus : sinstr
| SMult : sinstr.
Write a function to evaluate programs in the stack language. It
    should take as input a state, a stack represented as a list of
    numbers (top stack item is the head of the list), and a program
    represented as a list of instructions, and it should return the
    stack after executing the program.  Test your function on the
    examples below.
 
    Note that the specification leaves unspecified what to do when
    encountering an SPlus, SMinus, or SMult instruction if the
    stack contains less than two elements.  In a sense, it is
    immaterial what we do, since our compiler will never emit such a
    malformed program. 
Fixpoint s_execute (st : state) (stack : list nat)
(prog : list sinstr)
: list nat
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
Example s_execute1 :
s_execute empty_state []
[SPush 5; SPush 3; SPush 1; SMinus]
= [2; 5].
(* FILL IN HERE *) Admitted.
Example s_execute2 :
s_execute (t_update empty_state X 3) [3;4]
[SPush 4; SLoad X; SMult; SPlus]
= [15; 4].
(* FILL IN HERE *) Admitted.
Next, write a function that compiles an aexp into a stack
    machine program. The effect of running the program should be the
    same as pushing the value of the expression on the stack. 
Fixpoint s_compile (e : aexp) : list sinstr
(* REPLACE THIS LINE WITH ":= _your_definition_ ." *). Admitted.
After you've defined s_compile, prove the following to test
    that it works. 
Example s_compile1 :
s_compile (AMinus (AId X) (AMult (ANum 2) (AId Y)))
= [SLoad X; SPush 2; SLoad Y; SMult; SMinus].
(* FILL IN HERE *) Admitted.
Exercise: 4 stars, advanced (stack_compiler_correct)
Now we'll prove the correctness of the compiler implemented in the previous exercise. Remember that the specification left unspecified what to do when encountering an SPlus, SMinus, or SMult instruction if the stack contains less than two elements. (In order to make your correctness proof easier you might find it helpful to go back and change your implementation!)Theorem s_compile_correct : ∀ (st : state) (e : aexp),
s_execute st [] (s_compile e) = [ aeval st e ].
Proof.
(* FILL IN HERE *) Admitted.
Exercise: 3 stars, optional (short_circuit)
Most modern programming languages use a "short-circuit" evaluation rule for boolean and: to evaluate BAnd b1 b2, first evaluate b1. If it evaluates to false, then the entire BAnd expression evaluates to false immediately, without evaluating b2. Otherwise, b2 is evaluated to determine the result of the BAnd expression.(* FILL IN HERE *)
Module BreakImp.
Exercise: 4 stars, advanced (break_imp)
Inductive com : Type :=
| CSkip : com
| CBreak : com (* <-- new *)
| CAss : id → aexp → com
| CSeq : com → com → com
| CIf : bexp → com → com → com
| CWhile : bexp → com → com.
Notation "'SKIP'" :=
CSkip.
Notation "'BREAK'" :=
CBreak.
Notation "x '::=' a" :=
(CAss x a) (at level 60).
Notation "c1 ;; c2" :=
(CSeq c1 c2) (at level 80, right associativity).
Notation "'WHILE' b 'DO' c 'END'" :=
(CWhile b c) (at level 80, right associativity).
Notation "'IFB' c1 'THEN' c2 'ELSE' c3 'FI'" :=
(CIf c1 c2 c3) (at level 80, right associativity).
Next, we need to define the behavior of BREAK.  Informally,
    whenever BREAK is executed in a sequence of commands, it stops
    the execution of that sequence and signals that the innermost
    enclosing loop should terminate.  (If there aren't any
    enclosing loops, then the whole program simply terminates.)  The
    final state should be the same as the one in which the BREAK
    statement was executed.
 
    One important point is what to do when there are multiple loops
    enclosing a given BREAK. In those cases, BREAK should only
    terminate the innermost loop. Thus, after executing the
    following...
 
 
    One way of expressing this behavior is to add another parameter to
    the evaluation relation that specifies whether evaluation of a
    command executes a BREAK statement: 
       X ::= 0;;
Y ::= 1;;
WHILE 0 ≠ Y DO
WHILE TRUE DO
BREAK
END;;
X ::= 1;;
Y ::= Y - 1
END 
    ... the value of X should be 1, and not 0.
Y ::= 1;;
WHILE 0 ≠ Y DO
WHILE TRUE DO
BREAK
END;;
X ::= 1;;
Y ::= Y - 1
END
Inductive result : Type :=
| SContinue : result
| SBreak : result.
Reserved Notation "c1 '/' st '\\' s '/' st'"
(at level 40, st, s at level 39).
Intuitively, c / st \\ s / st' means that, if c is started in
    state st, then it terminates in state st' and either signals
    that the innermost surrounding loop (or the whole program) should
    exit immediately (s = SBreak) or that execution should continue
    normally (s = SContinue).
 
    The definition of the "c / st \\ s / st'" relation is very
    similar to the one we gave above for the regular evaluation
    relation (c / st \\ st') — we just need to handle the
    termination signals appropriately:
 
 
 Based on the above description, complete the definition of the
    ceval relation. 
-  If the command is SKIP, then the state doesn't change and
      execution of any enclosing loop can continue normally.
-  If the command is BREAK, the state stays unchanged but we
      signal a SBreak.
-  If the command is an assignment, then we update the binding for
      that variable in the state accordingly and signal that execution
      can continue normally.
-  If the command is of the form IFB b THEN c1 ELSE c2 FI, then
      the state is updated as in the original semantics of Imp, except
      that we also propagate the signal from the execution of
      whichever branch was taken.
-  If the command is a sequence c1 ;; c2, we first execute
      c1.  If this yields a SBreak, we skip the execution of c2
      and propagate the SBreak signal to the surrounding context;
      the resulting state is the same as the one obtained by
      executing c1 alone. Otherwise, we execute c2 on the state
      obtained after executing c1, and propagate the signal
      generated there.
- Finally, for a loop of the form WHILE b DO c END, the semantics is almost the same as before. The only difference is that, when b evaluates to true, we execute c and check the signal that it raises. If that signal is SContinue, then the execution proceeds as in the original semantics. Otherwise, we stop the execution of the loop, and the resulting state is the same as the one resulting from the execution of the current iteration. In either case, since BREAK only terminates the innermost loop, WHILE signals SContinue.
Inductive ceval : com → state → result → state → Prop :=
| E_Skip : ∀ st,
CSkip / st \\ SContinue / st
(* FILL IN HERE *)
where "c1 '/' st '\\' s '/' st'" := (ceval c1 st s st').
Now prove the following properties of your definition of ceval: 
Theorem break_ignore : ∀ c st st' s,
(BREAK;; c) / st \\ s / st' →
st = st'.
Proof.
(* FILL IN HERE *) Admitted.
Theorem while_continue : ∀ b c st st' s,
(WHILE b DO c END) / st \\ s / st' →
s = SContinue.
Proof.
(* FILL IN HERE *) Admitted.
Theorem while_stops_on_break : ∀ b c st st',
beval st b = true →
c / st \\ SBreak / st' →
(WHILE b DO c END) / st \\ SContinue / st'.
Proof.
(* FILL IN HERE *) Admitted.
Theorem while_break_true : ∀ b c st st',
(WHILE b DO c END) / st \\ SContinue / st' →
beval st' b = true →
∃ st'', c / st'' \\ SBreak / st'.
Proof.
(* FILL IN HERE *) Admitted.
☐ 
(WHILE b DO c END) / st \\ SContinue / st' →
beval st' b = true →
∃ st'', c / st'' \\ SBreak / st'.
Proof.
(* FILL IN HERE *) Admitted.
Theorem ceval_deterministic: ∀ (c:com) st st1 st2 s1 s2,
c / st \\ s1 / st1 →
c / st \\ s2 / st2 →
st1 = st2 ∧ s1 = s2.
Proof.
(* FILL IN HERE *) Admitted.
☐ 
c / st \\ s1 / st1 →
c / st \\ s2 / st2 →
st1 = st2 ∧ s1 = s2.
Proof.
(* FILL IN HERE *) Admitted.
End BreakImp.
Exercise: 4 stars, optional (add_for_loop)
Add C-style for loops to the language of commands, update the ceval definition to define the semantics of for loops, and add cases for for loops as needed so that all the proofs in this file are accepted by Coq.(* FILL IN HERE *)
(* $Date: 2017-07-04 07:52:33 -0400 (Tue, 04 Jul 2017) $ *)
