Built with Alectryon, running Coq+SerAPI v8.10.0+0.7.0. Coq sources are in this panel; goals and messages will appear in the other. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus.

Imp: Simple Imperative Programs

In this chapter, we take a more serious look at how to use Coq to study other things. Our case study is a simple imperative programming language called Imp, embodying a tiny core fragment of conventional mainstream languages such as C and Java. Here is a familiar mathematical function written in Imp.
Z ::= X;; Y ::= 1;; WHILE ~(Z = 0) DO Y ::= Y * Z;; Z ::= Z - 1 END
We concentrate here on defining the syntax and semantics of Imp; later chapters in Programming Language Foundations (Software Foundations, volume 2) develop a theory of program equivalence and introduce Hoare Logic, a widely used logic for reasoning about imperative programs.
Set Warnings "-notation-overridden,-parsing".
From Coq Require Import Bool.Bool.
From Coq Require Import Init.Nat.
From Coq Require Import Arith.Arith.
From Coq Require Import Arith.EqNat.
From Coq Require Import omega.Omega.
From Coq Require Import Lists.List.
From Coq Require Import Strings.String.
Import ListNotations.

From LF Require Import Maps.

(* ################################################################# *)

Arithmetic and Boolean Expressions

We'll present Imp in three parts: first a core language of arithmetic and boolean expressions, then an extension of these expressions with variables, and finally a language of commands including assignment, conditions, sequencing, and loops.
(* ================================================================= *)

Syntax

Module AExp.
These two definitions specify the abstract syntax of arithmetic and boolean expressions.
Inductive aexp : Type :=
  | ANum (n : nat)
  | APlus (a1 a2 : aexp)
  | AMinus (a1 a2 : aexp)
  | AMult (a1 a2 : aexp).

Inductive bexp : Type :=
  | BTrue
  | BFalse
  | BEq (a1 a2 : aexp)
  | BLe (a1 a2 : aexp)
  | BNot (b : bexp)
  | BAnd (b1 b2 : bexp).
In this chapter, we'll mostly 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
APlus (ANum 1) (AMult (ANum 2) (ANum 3)).
The optional chapter ImpParser develops a simple 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 already taken a course where these techniques are covered (e.g., a compilers course) you may want to skim it.
For comparison, here's a conventional BNF (Backus-Naur Form) grammar defining the same abstract syntax:
a ::= nat | a + a | a - a | a * a
b ::= true | false | a = a | a <= a | ~ b | b && b
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.
(* ================================================================= *)

Evaluation

Evaluating an arithmetic expression produces a number.
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.


aeval (APlus (ANum 2) (ANum 2)) = 4

aeval (APlus (ANum 2) (ANum 2)) = 4
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   => (aeval a1) =? (aeval a2)
  | BLe a1 a2   => (aeval a1) <=? (aeval a2)
  | BNot b1     => negb (beval b1)
  | BAnd b1 b2  => andb (beval b1) (beval b2)
  end.

(* ================================================================= *)

Optimization

We haven't defined very much yet, but we can already get some mileage out of the definitions. Suppose we define a function that takes an arithmetic expression and slightly simplifies it, changing every occurrence of 0 + e (i.e., (APlus (ANum 0) e) into just e.
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.

optimize_0plus (APlus (ANum 2) (APlus (ANum 0) (APlus (ANum 0) (ANum 1)))) = APlus (ANum 2) (ANum 1)

optimize_0plus (APlus (ANum 2) (APlus (ANum 0) (APlus (ANum 0) (ANum 1)))) = APlus (ANum 2) (ANum 1)
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.

forall a : aexp, aeval (optimize_0plus a) = aeval a

forall a : aexp, aeval (optimize_0plus a) = aeval a
a:aexp

aeval (optimize_0plus a) = aeval a
n:nat

aeval (optimize_0plus (ANum n)) = aeval (ANum n)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (AMinus a1 a2)) = aeval (AMinus a1 a2)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (AMult a1 a2)) = aeval (AMult a1 a2)
n:nat

aeval (optimize_0plus (ANum n)) = aeval (ANum n)
reflexivity.
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)
a1, a2:aexp
n:nat
Ea1:a1 = ANum n
IHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
a1, a2, a3, a4:aexp
Ea1:a1 = APlus a3 a4
IHa1:aeval (optimize_0plus (APlus a3 a4)) = aeval (APlus a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (APlus (APlus a3 a4) a2)) = aeval (APlus (APlus a3 a4) a2)
a1, a2, a3, a4:aexp
Ea1:a1 = AMinus a3 a4
IHa1:aeval (optimize_0plus (AMinus a3 a4)) = aeval (AMinus a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (APlus (AMinus a3 a4) a2)) = aeval (APlus (AMinus a3 a4) a2)
a1, a2, a3, a4:aexp
Ea1:a1 = AMult a3 a4
IHa1:aeval (optimize_0plus (AMult a3 a4)) = aeval (AMult a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (APlus (AMult a3 a4) a2)) = aeval (APlus (AMult a3 a4) a2)
a1, a2:aexp
n:nat
Ea1:a1 = ANum n
IHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
a1, a2:aexp
n:nat
En:n = 0
Ea1:a1 = ANum 0
IHa1:aeval (optimize_0plus (ANum 0)) = aeval (ANum 0)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum 0) a2)) = aeval (APlus (ANum 0) a2)
a1, a2:aexp
n, n0:nat
En:n = S n0
Ea1:a1 = ANum (S n0)
IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (APlus (ANum (S n0)) a2)) = aeval (APlus (ANum (S n0)) a2)
a1, a2:aexp
n:nat
En:n = 0
Ea1:a1 = ANum 0
IHa1:aeval (optimize_0plus (ANum 0)) = aeval (ANum 0)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum 0) a2)) = aeval (APlus (ANum 0) a2)
a1, a2:aexp
n:nat
En:n = 0
Ea1:a1 = ANum 0
IHa1:aeval (optimize_0plus (ANum 0)) = aeval (ANum 0)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a2) = aeval a2
apply IHa2.
a1, a2:aexp
n, n0:nat
En:n = S n0
Ea1:a1 = ANum (S n0)
IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum (S n0)) a2)) = aeval (APlus (ANum (S n0)) a2)
a1, a2:aexp
n, n0:nat
En:n = S n0
Ea1:a1 = ANum (S n0)
IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))
IHa2:aeval (optimize_0plus a2) = aeval a2

S (n0 + aeval (optimize_0plus a2)) = S (n0 + aeval a2)
a1, a2:aexp
n, n0:nat
En:n = S n0
Ea1:a1 = ANum (S n0)
IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))
IHa2:aeval (optimize_0plus a2) = aeval a2

S (n0 + aeval a2) = S (n0 + aeval a2)
reflexivity.
a1, a2, a3, a4:aexp
Ea1:a1 = APlus a3 a4
IHa1:aeval (optimize_0plus (APlus a3 a4)) = aeval (APlus a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (APlus a3 a4) a2)) = aeval (APlus (APlus a3 a4) a2)
a1, a2, a3, a4:aexp
Ea1:a1 = APlus a3 a4
IHa1:aeval (optimize_0plus (APlus a3 a4)) = aeval (APlus a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end + aeval (optimize_0plus a2) = aeval a3 + aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = APlus a3 a4
IHa1:aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end = aeval a3 + aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end + aeval (optimize_0plus a2) = aeval a3 + aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = APlus a3 a4
IHa1:aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end = aeval a3 + aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a3 + aeval a4 + aeval (optimize_0plus a2) = aeval a3 + aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = APlus a3 a4
IHa1:aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end = aeval a3 + aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a3 + aeval a4 + aeval a2 = aeval a3 + aeval a4 + aeval a2
reflexivity.
a1, a2, a3, a4:aexp
Ea1:a1 = AMinus a3 a4
IHa1:aeval (optimize_0plus (AMinus a3 a4)) = aeval (AMinus a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (AMinus a3 a4) a2)) = aeval (APlus (AMinus a3 a4) a2)
a1, a2, a3, a4:aexp
Ea1:a1 = AMinus a3 a4
IHa1:aeval (optimize_0plus (AMinus a3 a4)) = aeval (AMinus a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a3) - aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 - aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = AMinus a3 a4
IHa1:aeval (optimize_0plus a3) - aeval (optimize_0plus a4) = aeval a3 - aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a3) - aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 - aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = AMinus a3 a4
IHa1:aeval (optimize_0plus a3) - aeval (optimize_0plus a4) = aeval a3 - aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a3 - aeval a4 + aeval (optimize_0plus a2) = aeval a3 - aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = AMinus a3 a4
IHa1:aeval (optimize_0plus a3) - aeval (optimize_0plus a4) = aeval a3 - aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a3 - aeval a4 + aeval a2 = aeval a3 - aeval a4 + aeval a2
reflexivity.
a1, a2, a3, a4:aexp
Ea1:a1 = AMult a3 a4
IHa1:aeval (optimize_0plus (AMult a3 a4)) = aeval (AMult a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (AMult a3 a4) a2)) = aeval (APlus (AMult a3 a4) a2)
a1, a2, a3, a4:aexp
Ea1:a1 = AMult a3 a4
IHa1:aeval (optimize_0plus (AMult a3 a4)) = aeval (AMult a3 a4)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a3) * aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 * aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = AMult a3 a4
IHa1:aeval (optimize_0plus a3) * aeval (optimize_0plus a4) = aeval a3 * aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a3) * aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 * aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = AMult a3 a4
IHa1:aeval (optimize_0plus a3) * aeval (optimize_0plus a4) = aeval a3 * aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a3 * aeval a4 + aeval (optimize_0plus a2) = aeval a3 * aeval a4 + aeval a2
a1, a2, a3, a4:aexp
Ea1:a1 = AMult a3 a4
IHa1:aeval (optimize_0plus a3) * aeval (optimize_0plus a4) = aeval a3 * aeval a4
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a3 * aeval a4 + aeval a2 = aeval a3 * aeval a4 + aeval a2
reflexivity.
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (AMinus a1 a2)) = aeval (AMinus a1 a2)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a1) - aeval (optimize_0plus a2) = aeval a1 - aeval a2
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a1 - aeval (optimize_0plus a2) = aeval a1 - aeval a2
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a1 - aeval a2 = aeval a1 - aeval a2
reflexivity.
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (AMult a1 a2)) = aeval (AMult a1 a2)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus a1) * aeval (optimize_0plus a2) = aeval a1 * aeval a2
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a1 * aeval (optimize_0plus a2) = aeval a1 * aeval a2
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval a1 * aeval a2 = aeval a1 * aeval a2
reflexivity. Qed. (* ################################################################# *)

Coq Automation

The amount of repetition in this last proof is a little annoying. And if either the language of arithmetic expressions or the optimization being proved sound were significantly more complex, it would start to be a real problem.
So far, we've been doing all our proofs using just a small handful of Coq's tactics and completely ignoring its powerful facilities for constructing parts of proofs automatically. This section introduces some of these facilities, and we will see more over the next several chapters. Getting used to them will take some energy -- Coq's automation is a power tool -- but it will allow us to scale up our efforts to more complex definitions and more interesting properties without becoming overwhelmed by boring, repetitive, low-level details.
(* ================================================================= *)

Tacticals

Tacticals is Coq's term for tactics that take other tactics as arguments -- "higher-order tactics," if you will.
(* ----------------------------------------------------------------- *)

The try Tactical

If T is a tactic, then try T is a tactic that is just like T except that, if T fails, try T successfully does nothing at all (rather than failing).

forall ae : aexp, aeval ae = aeval ae

forall ae : aexp, aeval ae = aeval ae
try reflexivity. (* This just does [reflexivity]. *) Qed.

forall P : Prop, P -> P

forall P : Prop, P -> P
P:Prop
HP:P

P
P:Prop
HP:P

P
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)

In its most common form, the ; tactical takes two tactics as arguments. The compound tactic T;T' first performs T and then performs T' on each subgoal generated by T.
For example, consider the following trivial lemma:

forall n : nat, (0 <=? n) = true

forall n : nat, (0 <=? n) = true
n:nat

(0 <=? n) = true

(0 <=? 0) = true
n:nat
(0 <=? S n) = true
(* Leaves two subgoals, which are discharged identically... *)

(0 <=? 0) = true

true = true
reflexivity.
n:nat

(0 <=? S n) = true
n:nat

true = true
reflexivity. Qed.
We can simplify this proof using the ; tactical:

forall n : nat, (0 <=? n) = true

forall n : nat, (0 <=? n) = true
n:nat

(0 <=? n) = true
(* [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.

forall a : aexp, aeval (optimize_0plus a) = aeval a

forall a : aexp, aeval (optimize_0plus a) = aeval a
a:aexp

aeval (optimize_0plus a) = aeval a
n:nat

aeval (optimize_0plus (ANum n)) = aeval (ANum n)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2
aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)
(* ... but the remaining cases -- ANum and APlus -- are different: *)
n:nat

aeval (optimize_0plus (ANum n)) = aeval (ANum n)
reflexivity.
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)
a1, a2:aexp
n:nat
Ea1:a1 = ANum n
IHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
(* 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, a2:aexp
n:nat
Ea1:a1 = ANum n
IHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
destruct n eqn:En; 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,
aeval (optimize_0plus a) = aeval a.
Proof: By induction on a. Most cases follow directly from the IH. The remaining cases are as follows:
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:

forall a : aexp, aeval (optimize_0plus a) = aeval a

forall a : aexp, aeval (optimize_0plus a) = aeval a
a:aexp

aeval (optimize_0plus a) = aeval a
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)
(* The interesting case is when a = APlus a1 a2. *)
a1, a2:aexp
IHa1:aeval (optimize_0plus a1) = aeval a1
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)
n:nat
a2:aexp
IHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
n:nat
a2:aexp
IHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)
IHa2:aeval (optimize_0plus a2) = aeval a2

aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
destruct n; simpl; rewrite IHa2; reflexivity. Qed. (* ----------------------------------------------------------------- *)

The ; Tactical (General Form)

The ; tactical also has a more general form than the simple T;T' we've seen above. If T, T1, ..., Tn are tactics, then
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.
So T;T' is just special notation for the case when all of the Ti's are the same tactic; i.e., T;T' is shorthand for:
T; T' | T' | ... | T'
(* ----------------------------------------------------------------- *)

The repeat Tactical

The repeat tactical takes another tactic and keeps applying this tactic until it fails. Here is an example showing that 10 is in a long list using repeat.

In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]

In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]
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).

In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]

In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]

In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]
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, 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 (i.e., it does not terminate), this simply means that we have failed to construct a proof, not that we have constructed a wrong one.

Exercise: 3 stars, standard (optimize_0plus_b_sound)

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 that performs this transformation on bexps and prove it is sound. Use the tacticals we've just seen to make the proof as elegant as possible.
optimize_0plus_b:bexp -> bexp
b:bexp

bexp
Admitted.

forall b : bexp, beval (optimize_0plus_b b) = beval b

forall b : bexp, beval (optimize_0plus_b b) = beval b
(* FILL IN HERE *) Admitted.

Exercise: 4 stars, standard, optional (optimize)

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 its correctness proof -- and build up to something more interesting incrementially.)
(* FILL IN HERE 

    [] *)

(* ================================================================= *)

Defining New Tactic Notations

Coq also provides several ways of "programming" tactic scripts.
The Tactic Notation mechanism is the easiest to come to grips with, and it offers plenty of power for many purposes. Here's an example.
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

The omega tactic implements a decision procedure for a subset of first-order logic called Presburger arithmetic. It is based on the Omega algorithm invented by William Pugh Pugh 1991 (in Bib.v).
If the goal is a universally quantified formula made out of
then invoking omega will either solve the goal or fail, meaning that the goal is actually false. (If the goal is not of this form, omega will also fail.)

forall m n o p : nat, m + n <= n + o /\ o + 3 = p + 3 -> m <= p

forall m n o p : nat, m + n <= n + o /\ o + 3 = p + 3 -> m <= p
m, n, o, p:nat
H:m + n <= n + o /\ o + 3 = p + 3

m <= p
omega. Qed.
(Note the From Coq Require Import omega.Omega. at the top of the file.)
(* ================================================================= *)

A Few More Handy Tactics

Finally, here are some miscellaneous tactics that you may find convenient.
We'll see examples of all of these as we go along.
(* ################################################################# *)

Evaluation as a Relation

We have presented aeval and beval as functions defined by Fixpoints. Another way to think about evaluation -- one that we will see is often more flexible -- is as a relation between expressions and their values. This leads naturally to Inductive definitions like the following one for arithmetic expressions...
Module aevalR_first_try.

Inductive aevalR : aexp -> nat -> Prop :=
  | E_ANum n :
      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).

Module TooHardToRead.

(* A small notational aside. We would previously have written the
   definition of [aevalR] like this, with explicit names for the
   hypotheses in each case: *)

Inductive aevalR : aexp -> nat -> Prop :=
  | E_ANum n :
      aevalR (ANum n) n
  | E_APlus (e1 e2: aexp) (n1 n2: nat)
      (H1 : aevalR e1 n1)
      (H2 : aevalR e2 n2) :
      aevalR (APlus e1 e2) (n1 + n2)
  | E_AMinus (e1 e2: aexp) (n1 n2: nat)
      (H1 : aevalR e1 n1)
      (H2 : aevalR e2 n2) :
      aevalR (AMinus e1 e2) (n1 - n2)
  | E_AMult (e1 e2: aexp) (n1 n2: nat)
      (H1 : aevalR e1 n1)
      (H2 : aevalR e2 n2) :
      aevalR (AMult e1 e2) (n1 * n2).
Instead, we've chosen to leave the hypotheses anonymous, just giving their types. This style gives us less control over the names that Coq chooses during proofs involving aevalR, but it makes the definition itself quite a bit lighter.
End TooHardToRead.
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 90, 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

In informal discussions, it is convenient to write the rules for aevalR and similar relations in the more readable graphical form of inference rules, where the premises above the line justify the conclusion below the line (we have already seen them in the IndProp chapter).
For example, the constructor E_APlus...
| E_APlus : forall (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:
e1 \\ n1 e2 \\ n2
(E_APlus) APlus e1 e2 \\ n1+n2
Formally, there is nothing deep about inference rules: they are just implications. You can read the rule name on the right as the name of the constructor and read each of the linebreaks between the premises above the line (as well as the line itself) as . All the variables mentioned in the rule (e1, n1, etc.) are implicitly bound by universal quantifiers at the beginning. (Such variables are often called metavariables to distinguish them from the variables of the language we are defining. At the moment, our arithmetic expressions don't include variables, but we'll soon be adding them.) The whole collection of rules is understood as being wrapped in an Inductive declaration. In informal prose, this is either elided or else indicated by saying something like "Let aevalR be the smallest relation closed under the following rules...".
For example, \\ is the smallest relation closed under these rules:
(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

Exercise: 1 star, standard, optional (beval_rules)

Here, again, is the Coq definition of the beval function:
Fixpoint beval (e : bexp) : bool := match e with | BTrue => true | BFalse => false | BEq a1 a2 => (aeval a1) =? (aeval a2) | BLe a1 a2 => (aeval a1) <=? (aeval a2) | BNot b1 => negb (beval b1) | BAnd b1 b2 => andb (beval b1) (beval b2) end.
Write out a corresponding definition of boolean evaluation as a relation (in inference rule notation).
(* FILL IN HERE *)

(* Do not modify the following line: *)
Definition manual_grade_for_beval_rules : option (nat*string) := None.
(* ================================================================= *)

Equivalence of the Definitions

It is straightforward to prove that the relational and functional definitions of evaluation agree:

forall (a : aexp) (n : nat), a \\ n <-> aeval a = n

forall (a : aexp) (n : nat), a \\ n <-> aeval a = n
a:aexp
n:nat

a \\ n -> aeval a = n
a:aexp
n:nat
aeval a = n -> a \\ n
a:aexp
n:nat

a \\ n -> aeval a = n
a:aexp
n:nat
H:a \\ n

aeval a = n
n:nat

n = n
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2
aeval e1 + aeval e2 = n1 + n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2
aeval e1 - aeval e2 = n1 - n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2
aeval e1 * aeval e2 = n1 * n2
n:nat

n = n
reflexivity.
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

aeval e1 + aeval e2 = n1 + n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

n1 + aeval e2 = n1 + n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

n1 + n2 = n1 + n2
reflexivity.
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

aeval e1 - aeval e2 = n1 - n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

n1 - aeval e2 = n1 - n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

n1 - n2 = n1 - n2
reflexivity.
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

aeval e1 * aeval e2 = n1 * n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

n1 * aeval e2 = n1 * n2
e1, e2:aexp
n1, n2:nat
H:e1 \\ n1
H0:e2 \\ n2
IHaevalR1:aeval e1 = n1
IHaevalR2:aeval e2 = n2

n1 * n2 = n1 * n2
reflexivity.
a:aexp
n:nat

aeval a = n -> a \\ n
a:aexp

forall n : nat, aeval a = n -> a \\ n
n0:nat

ANum n0 \\ n0
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
APlus a1 a2 \\ aeval a1 + aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
AMinus a1 a2 \\ aeval a1 - aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
AMult a1 a2 \\ aeval a1 * aeval a2
n0:nat

ANum n0 \\ n0
apply E_ANum.
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

APlus a1 a2 \\ aeval a1 + aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

a1 \\ aeval a1
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

aeval a1 = aeval a1
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

aeval a2 = aeval a2
reflexivity.
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

AMinus a1 a2 \\ aeval a1 - aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

a1 \\ aeval a1
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

aeval a1 = aeval a1
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

aeval a2 = aeval a2
reflexivity.
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

AMult a1 a2 \\ aeval a1 * aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

a1 \\ aeval a1
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

aeval a1 = aeval a1
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n
a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

a2 \\ aeval a2
a1, a2:aexp
IHa1:forall n : nat, aeval a1 = n -> a1 \\ n
IHa2:forall n : nat, aeval a2 = n -> a2 \\ n

aeval a2 = aeval a2
reflexivity. Qed.
We can make the proof quite a bit shorter by making more use of tacticals.

forall (a : aexp) (n : nat), a \\ n <-> aeval a = n

forall (a : aexp) (n : nat), a \\ n <-> aeval a = n
(* WORKED IN CLASS *)
a:aexp
n:nat

a \\ n -> aeval a = n
a:aexp
n:nat
aeval a = n -> a \\ n
a:aexp
n:nat

a \\ n -> aeval a = n
intros H; induction H; subst; reflexivity.
a:aexp
n:nat

aeval a = n -> a \\ n
a:aexp

forall n : nat, aeval a = n -> a \\ n
induction a; simpl; intros; subst; constructor; try apply IHa1; try apply IHa2; reflexivity. Qed.

Exercise: 3 stars, standard (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 *)
.


forall (b : bexp) (bv : bool), bevalR b bv <-> beval b = bv

forall (b : bexp) (bv : bool), bevalR b bv <-> beval b = bv
(* FILL IN HERE *) Admitted.
End AExp.

(* ================================================================= *)

Computational vs. Relational Definitions

For the definitions of evaluation for arithmetic and boolean expressions, the choice of whether to use functional or relational definitions is mainly a matter of taste: either way works.
However, there are circumstances where relational definitions of evaluation work much better than functional ones.
Module aevalR_division.
For example, suppose that we wanted to extend the arithmetic operations with division:
Inductive aexp : Type :=
  | ANum (n : nat)
  | APlus (a1 a2 : aexp)
  | AMinus (a1 a2 : aexp)
  | AMult (a1 a2 : aexp)
  | ADiv (a1 a2 : 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 90, 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.
Or suppose 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 probability distribution for the results, just saying what results are possible.)
Reserved Notation "e '\\' n" (at level 90, left associativity).

Inductive aexp : Type :=
  | AAny                           (* <--- NEW *)
  | ANum (n : nat)
  | APlus (a1 a2 : aexp)
  | AMinus (a1 a2 : aexp)
  | AMult (a1 a2 : 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? In the examples we've just seen, relational definitions turned out to be more useful 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 real choice. But what about when both styles are workable?
One point in favor of relational definitions is that they can be 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" from Gallina 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.
(* ################################################################# *)

Expressions With Variables

Back to defining Imp. The next thing we need to do is to enrich our arithmetic and boolean expressions with variables. To keep things simple, we'll assume that all variables are global and that they only hold numbers.
(* ================================================================= *)

States

Since we'll want to look variables up to find out their current values, we'll reuse maps from the Maps chapter, and strings will be used to represent variables in Imp.
A machine state (or just state) represents the current values of all variables at some point in the execution of a program.
For simplicity, we assume that the state is defined for all variables, even though any given program is only going to mention a finite number of them. The state captures all of the information stored in memory. For Imp programs, because each variable stores a natural number, we can represent the state as a mapping from strings to nat, and will use 0 as default value in the store. For more complex programming languages, the state might have more structure.
Definition state := total_map nat.

(* ================================================================= *)

Syntax

We can add variables to the arithmetic expressions we had before by simply adding one more constructor:
Inductive aexp : Type :=
  | ANum (n : nat)
  | AId (x : string)              (* <--- NEW *)
  | APlus (a1 a2 : aexp)
  | AMinus (a1 a2 : aexp)
  | AMult (a1 a2 : aexp).
Defining a few variable names as notational shorthands will make examples easier to read:
Definition W : string := "W".
Definition X : string := "X".
Definition Y : string := "Y".
Definition Z : string := "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 developed to Imp, this overloading should not cause confusion.)
The definition of bexps is unchanged (except that it now refers to the new aexps):
Inductive bexp : Type :=
  | BTrue
  | BFalse
  | BEq (a1 a2 : aexp)
  | BLe (a1 a2 : aexp)
  | BNot (b : bexp)
  | BAnd (b1 b2 : bexp).

(* ================================================================= *)

Notations

To make Imp programs easier to read and write, we introduce some notations and implicit coercions.
You do not need to understand exactly what these declarations do. Briefly, though, the Coercion declaration in Coq stipulates that a function (or constructor) can be implicitly used by the type system to coerce a value of the input type to a value of the output type. For instance, the coercion declaration for AId allows us to use plain strings when an aexp is expected; the string will implicitly be wrapped with AId.
The notations below are declared in specific notation scopes, in order to avoid conflicts with other interpretations of the same symbols. Again, it is not necessary to understand the details, but it is important to recognize that we are defining new intepretations for some familiar operators like +, -, ×, =., [<=], etc.
Coercion AId : string >-> aexp.
Coercion ANum : nat >-> aexp.

Definition bool_to_bexp (b : bool) : bexp :=
  if b then BTrue else BFalse.
Coercion bool_to_bexp : bool >-> bexp.

Declaring a scope implicitly is deprecated; use in advance an explicit "Declare Scope imp_scope.". [undeclared-scope,deprecated]
Bind Scope imp_scope with bexp. Delimit Scope imp_scope with imp. Notation "x + y" := (APlus x y) (at level 50, left associativity) : imp_scope. Notation "x - y" := (AMinus x y) (at level 50, left associativity) : imp_scope. Notation "x * y" := (AMult x y) (at level 40, left associativity) : imp_scope. Notation "x <= y" := (BLe x y) (at level 70, no associativity) : imp_scope. Notation "x = y" := (BEq x y) (at level 70, no associativity) : imp_scope. Notation "x && y" := (BAnd x y) (at level 40, left associativity) : imp_scope. Notation "'~' b" := (BNot b) (at level 75, right associativity) : imp_scope. Definition example_aexp := (3 + (X * 2))%imp : aexp. Definition example_bexp := (true && ~(X <= 4))%imp : bexp.
One downside of these coercions is that they can make it a little harder for humans to calculate the types of expressions. If you get confused, try doing Set Printing Coercions to see exactly what is going on.
Set Printing Coercions.

example_bexp = (bool_to_bexp true && (~ AId X <= ANum 4))%imp : bexp
Unset Printing Coercions. (* ================================================================= *)

Evaluation

The arith and boolean evaluators are extended to handle variables in the obvious way, taking a state as an extra argument:
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   => (aeval st a1) =? (aeval st a2)
  | BLe a1 a2   => (aeval st a1) <=? (aeval st a2)
  | BNot b1     => negb (beval st b1)
  | BAnd b1 b2  => andb (beval st b1) (beval st b2)
  end.
We specialize our notation for total maps to the specific case of states, i.e. using (_ !-> 0) as empty state.
Definition empty_st := (_ !-> 0).
Now we can add a notation for a "singleton state" with just one variable bound to a value.
Notation "a '!->' x" := (t_update empty_st a x) (at level 100).


aeval (X !-> 5) (3 + X * 2) = 13

aeval (X !-> 5) (3 + X * 2) = 13
reflexivity. Qed.

beval (X !-> 5) (true && (~ X <= 4)) = true

beval (X !-> 5) (true && (~ X <= 4)) = true
reflexivity. Qed. (* ################################################################# *)

Commands

Now we are ready define the syntax and behavior of Imp commands (sometimes called statements).
(* ================================================================= *)

Syntax

Informally, commands c are described by the following BNF grammar.
c ::= SKIP | x ::= a | c ;; c | TEST b THEN c ELSE c FI | WHILE b DO c END
(We choose this slightly awkward concrete syntax for the sake of being able to define Imp syntax using Coq's notation mechanism. In particular, we use TEST to avoid conflicting with the if and IF notations from the standard library.) For example, here's factorial in Imp:
Z ::= X;; Y ::= 1;; WHILE ~(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.
Here is the formal definition of the abstract syntax of commands:
Inductive com : Type :=
  | CSkip
  | CAss (x : string) (a : aexp)
  | CSeq (c1 c2 : com)
  | CIf (b : bexp) (c1 c2 : com)
  | CWhile (b : bexp) (c : com).
As for expressions, we can use a few Notation declarations to make reading and writing Imp programs more convenient.
Bind Scope imp_scope with com.
Notation "'SKIP'" :=
   CSkip : imp_scope.
Notation "x '::=' a" :=
  (CAss x a) (at level 60) : imp_scope.
Notation "c1 ;; c2" :=
  (CSeq c1 c2) (at level 80, right associativity) : imp_scope.
Notation "'WHILE' b 'DO' c 'END'" :=
  (CWhile b c) (at level 80, right associativity) : imp_scope.
Notation "'TEST' c1 'THEN' c2 'ELSE' c3 'FI'" :=
  (CIf c1 c2 c3) (at level 80, right associativity) : imp_scope.
For example, here is the factorial function again, written as a formal definition to Coq:
Definition fact_in_coq : com :=
  (Z ::= X;;
  Y ::= 1;;
  WHILE ~(Z = 0) DO
    Y ::= Y * Z;;
    Z ::= Z - 1
  END)%imp.

(* ================================================================= *)

Desugaring notations

Coq offers a rich set of features to manage the increasing complexity of the objects we work with, such as coercions and notations. However, their heavy usage can make for quite overwhelming syntax. It is often instructive to "turn off" those features to get a more elementary picture of things, using the following commands:
These commands can also be used in the middle of a proof, to elaborate the current goal and context.
Unset Printing Notations.
fact_in_coq = CSeq (CAss Z X) (CSeq (CAss Y (S O)) (CWhile (BNot (BEq Z O)) (CSeq (CAss Y (AMult Y Z)) (CAss Z (AMinus Z (S O)))))) : com
Set Printing Notations. Set Printing Coercions.
fact_in_coq = (Z ::= AId X;; Y ::= ANum 1;; WHILE ~ AId Z = ANum 0 DO Y ::= AId Y * AId Z;; Z ::= AId Z - ANum 1 END)%imp : com
Unset Printing Coercions. (* ================================================================= *)

The Locate command

(* ----------------------------------------------------------------- *)

Finding notations

When faced with unknown notation, use Locate with a string containing one of its symbols to see its possible interpretations.
Notation "x && y" := andb x y : bool_scope (default interpretation) "x && y" := BAnd x y : imp_scope
Notation "c1 ;; c2" := CSeq c1 c2 : imp_scope
Notation "'WHILE' b 'DO' c 'END'" := CWhile b c : imp_scope "'_' '!->' v" := t_empty v (default interpretation)
(* ----------------------------------------------------------------- *)

Finding identifiers

When used with an identifier, the command Locate prints the full path to every value in scope with the same name. This is useful to troubleshoot problems due to variable shadowing.
Inductive SerTop.aexp Inductive SerTop.AExp.aexp (shorter name to refer to it in current context is AExp.aexp) Inductive SerTop.aevalR_division.aexp (shorter name to refer to it in current context is aevalR_division.aexp) Inductive SerTop.aevalR_extended.aexp (shorter name to refer to it in current context is aevalR_extended.aexp)
(* ================================================================= *)

More Examples

Assignment:
Definition plus2 : com :=
  X ::= X + 2.

Definition XtimesYinZ : com :=
  Z ::= X * Y.

Definition subtract_slowly_body : com :=
  Z ::= Z - 1 ;;
  X ::= X - 1.

(* ----------------------------------------------------------------- *)

Loops

Definition subtract_slowly : com :=
  (WHILE ~(X = 0) DO
    subtract_slowly_body
  END)%imp.

Definition subtract_3_from_5_slowly : com :=
  X ::= 3 ;;
  Z ::= 5 ;;
  subtract_slowly.

(* ----------------------------------------------------------------- *)

An infinite loop:

Definition loop : com :=
  WHILE true DO
    SKIP
  END.

(* ################################################################# *)

Evaluating Commands

Next we need to define what it means to evaluate an Imp command. The fact that WHILE loops don't necessarily terminate makes defining an evaluation function tricky...
(* ================================================================= *)

Evaluation as a Function (Failed Attempt)

Here's an attempt at defining an evaluation function for commands, omitting the WHILE case.
The following declaration is needed to be able to use the notations in match patterns.
Open Scope imp_scope.
Fixpoint ceval_fun_no_while (st : state) (c : com)
                          : state :=
  match c with
    | SKIP =>
        st
    | x ::= a1 =>
        (x !-> (aeval st a1) ; st)
    | c1 ;; c2 =>
        let st' := ceval_fun_no_while st c1 in
        ceval_fun_no_while st' c2
    | TEST 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.
Close Scope imp_scope.
In a traditional functional programming language like OCaml or Haskell we could add the WHILE case as follows:
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.
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).
(* ================================================================= *)

Evaluation as a Relation

Here's a better way: define ceval as a relation rather than a function -- i.e., define it in Prop instead of Type, as we did for aevalR above.
This is an important change. Besides freeing us from awkward workarounds, it gives us a lot more flexibility in the definition. For example, if we add nondeterministic features like any to the language, we want the definition of evaluation to be nondeterministic -- i.e., not only will it not be total, it will not even be a function!
We'll use the notation st =[ c ]=> st' for the ceval relation: st =[ c ]=> st' means that executing program c in a starting state st results in an ending state st'. This can be pronounced "c takes state st to st'".
(* ----------------------------------------------------------------- *)

Operational Semantics

Here is an informal definition of evaluation, presented as inference rules for readability:
(E_Skip) st = SKIP => st
aeval st a1 = n
(E_Ass) st = x := a1 => (x !-> n ; st)
st = c1 => st' st' = c2 => st''
(E_Seq) st = c1;;c2 => st''
beval st b1 = true st = c1 => st'
(E_IfTrue) st = TEST b1 THEN c1 ELSE c2 FI => st'
beval st b1 = false st = c2 => st'
(E_IfFalse) st = TEST b1 THEN c1 ELSE c2 FI => st'
beval st b = false
(E_WhileFalse) st = WHILE b DO c END => st
beval st b = true st = c => st' st' = WHILE b DO c END => st''
(E_WhileTrue) st = WHILE b DO c END => st''
Here is the formal definition. Make sure you understand how it corresponds to the inference rules.
Reserved Notation "st '=[' c ']=>' st'"
                  (at level 40).

Inductive ceval : com -> state -> state -> Prop :=
  | E_Skip : forall st,
      st =[ SKIP ]=> st
  | E_Ass  : forall st a1 n x,
      aeval st a1 = n ->
      st =[ x ::= a1 ]=> (x !-> n ; st)
  | E_Seq : forall c1 c2 st st' st'',
      st  =[ c1 ]=> st'  ->
      st' =[ c2 ]=> st'' ->
      st  =[ c1 ;; c2 ]=> st''
  | E_IfTrue : forall st st' b c1 c2,
      beval st b = true ->
      st =[ c1 ]=> st' ->
      st =[ TEST b THEN c1 ELSE c2 FI ]=> st'
  | E_IfFalse : forall st st' b c1 c2,
      beval st b = false ->
      st =[ c2 ]=> st' ->
      st =[ TEST b THEN c1 ELSE c2 FI ]=> st'
  | E_WhileFalse : forall b st c,
      beval st b = false ->
      st =[ WHILE b DO c END ]=> st
  | E_WhileTrue : forall st st' st'' b c,
      beval st b = true ->
      st  =[ c ]=> st' ->
      st' =[ WHILE b DO c END ]=> st'' ->
      st  =[ WHILE b DO c END ]=> st''

  where "st =[ c ]=> st'" := (ceval c 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.

empty_st =[ X ::= 2;; TEST X <= 1 THEN Y ::= 3 ELSE Z ::= 4 FI ]=> (Z !-> 4; X !-> 2)

empty_st =[ X ::= 2;; TEST X <= 1 THEN Y ::= 3 ELSE Z ::= 4 FI ]=> (Z !-> 4; X !-> 2)
(* We must supply the intermediate state *)

empty_st =[ X ::= 2 ]=> (X !-> 2)

(X !-> 2) =[ TEST X <= 1 THEN Y ::= 3 ELSE Z ::= 4 FI ]=> (Z !-> 4; X !-> 2)

empty_st =[ X ::= 2 ]=> (X !-> 2)

aeval empty_st 2 = 2
reflexivity.

(X !-> 2) =[ TEST X <= 1 THEN Y ::= 3 ELSE Z ::= 4 FI ]=> (Z !-> 4; X !-> 2)

beval (X !-> 2) (X <= 1) = false

(X !-> 2) =[ Z ::= 4 ]=> (Z !-> 4; X !-> 2)

(X !-> 2) =[ Z ::= 4 ]=> (Z !-> 4; X !-> 2)

aeval (X !-> 2) 4 = 4
reflexivity. Qed.

Exercise: 2 stars, standard (ceval_example2)


empty_st =[ X ::= 0;; Y ::= 1;; Z ::= 2 ]=> (Z !-> 2; Y !-> 1; X !-> 0)

empty_st =[ X ::= 0;; Y ::= 1;; Z ::= 2 ]=> (Z !-> 2; Y !-> 1; X !-> 0)
(* FILL IN HERE *) Admitted.

Exercise: 3 stars, standard, optional (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).

com
Admitted.

(X !-> 2) =[ pup_to_n ]=> (X !-> 0; Y !-> 3; X !-> 1; Y !-> 2; Y !-> 0; X !-> 2)

(X !-> 2) =[ pup_to_n ]=> (X !-> 0; Y !-> 3; X !-> 1; Y !-> 2; Y !-> 0; X !-> 2)
(* FILL IN HERE *) Admitted.
(* ================================================================= *)

Determinism of Evaluation

Changing from a computational to a relational definition of evaluation is a good move because it frees us from the artificial requirement that evaluation should be a total function. But it also raises a question: Is the second definition of evaluation really a partial function? Or is it possible that, beginning from the same state st, we could evaluate some command c in different ways to reach two different output states st' and st''?
In fact, this cannot happen: ceval is a partial function:

forall (c : com) (st st1 st2 : state), st =[ c ]=> st1 -> st =[ c ]=> st2 -> st1 = st2

forall (c : com) (st st1 st2 : state), st =[ c ]=> st1 -> st =[ c ]=> st2 -> st1 = st2
c:com
st, st1, st2:state
E1:st =[ c ]=> st1
E2:st =[ c ]=> st2

st1 = st2
c:com
st, st1:state
E1:st =[ c ]=> st1

forall st2 : state, st =[ c ]=> st2 -> st1 = st2
st2:state
E2:st2 =[ SKIP ]=> st2

st2 = st2
st:state
a1:aexp
x:string
E2:st =[ x ::= a1 ]=> (x !-> aeval st a1; st)
(x !-> aeval st a1; st) = (x !-> aeval st a1; st)
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
st'0:state
H1:st =[ c1 ]=> st'0
H4:st'0 =[ c2 ]=> st2
st'' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = true
E1:st =[ c1 ]=> st'
IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = true
H6:st =[ c1 ]=> st2
st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = true
E1:st =[ c1 ]=> st'
IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = false
H6:st =[ c2 ]=> st2
st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = false
E1:st =[ c2 ]=> st'
IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = true
H6:st =[ c1 ]=> st2
st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = false
E1:st =[ c2 ]=> st'
IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = false
H6:st =[ c2 ]=> st2
st' = st2
b:bexp
c:com
st2:state
E2:st2 =[ WHILE b DO c END ]=> st2
H, H4:beval st2 b = false
st2 = st2
b:bexp
st:state
c:com
H:beval st b = false
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st':state
H2:beval st b = true
H3:st =[ c ]=> st'
H6:st' =[ WHILE b DO c END ]=> st2
st = st2
st', st'':state
b:bexp
c:com
st2:state
E2:st2 =[ WHILE b DO c END ]=> st2
E1_1:st2 =[ c ]=> st'
H:beval st2 b = true
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st2 =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
H4:beval st2 b = false
st'' = st2
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st'0:state
H2:beval st b = true
H3:st =[ c ]=> st'0
H6:st'0 =[ WHILE b DO c END ]=> st2
st'' = st2
st2:state
E2:st2 =[ SKIP ]=> st2

st2 = st2
reflexivity.
st:state
a1:aexp
x:string
E2:st =[ x ::= a1 ]=> (x !-> aeval st a1; st)

(x !-> aeval st a1; st) = (x !-> aeval st a1; st)
reflexivity.
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
st'0:state
H1:st =[ c1 ]=> st'0
H4:st'0 =[ c2 ]=> st2

st'' = st2
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
st'0:state
H1:st =[ c1 ]=> st'0
H4:st'0 =[ c2 ]=> st2

st' = st'0
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
st'0:state
H1:st =[ c1 ]=> st'0
H4:st'0 =[ c2 ]=> st2
EQ1:st' = st'0
st'' = st2
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
st'0:state
H1:st =[ c1 ]=> st'0
H4:st'0 =[ c2 ]=> st2

st' = st'0
apply IHE1_1; assumption.
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
st'0:state
H1:st =[ c1 ]=> st'0
H4:st'0 =[ c2 ]=> st2
EQ1:st' = st'0

st'' = st2
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
H4:st' =[ c2 ]=> st2
H1:st =[ c1 ]=> st'

st'' = st2
c1, c2:com
st, st', st'':state
E1_1:st =[ c1 ]=> st'
E1_2:st' =[ c2 ]=> st''
IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0
st2:state
E2:st =[ c1;; c2 ]=> st2
H4:st' =[ c2 ]=> st2
H1:st =[ c1 ]=> st'

st' =[ c2 ]=> st2
assumption.
st, st':state
b:bexp
c1, c2:com
H:beval st b = true
E1:st =[ c1 ]=> st'
IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = true
H6:st =[ c1 ]=> st2

st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = true
E1:st =[ c1 ]=> st'
IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = true
H6:st =[ c1 ]=> st2

st =[ c1 ]=> st2
assumption.
st, st':state
b:bexp
c1, c2:com
H:beval st b = true
E1:st =[ c1 ]=> st'
IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = false
H6:st =[ c2 ]=> st2

st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = true
E1:st =[ c1 ]=> st'
IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:true = false
H6:st =[ c2 ]=> st2

st' = st2
discriminate H5.
st, st':state
b:bexp
c1, c2:com
H:beval st b = false
E1:st =[ c2 ]=> st'
IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = true
H6:st =[ c1 ]=> st2

st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = false
E1:st =[ c2 ]=> st'
IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:false = true
H6:st =[ c1 ]=> st2

st' = st2
discriminate H5.
st, st':state
b:bexp
c1, c2:com
H:beval st b = false
E1:st =[ c2 ]=> st'
IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = false
H6:st =[ c2 ]=> st2

st' = st2
st, st':state
b:bexp
c1, c2:com
H:beval st b = false
E1:st =[ c2 ]=> st'
IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0
st2:state
E2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2
H5:beval st b = false
H6:st =[ c2 ]=> st2

st =[ c2 ]=> st2
assumption.
b:bexp
c:com
st2:state
E2:st2 =[ WHILE b DO c END ]=> st2
H, H4:beval st2 b = false

st2 = st2
reflexivity.
b:bexp
st:state
c:com
H:beval st b = false
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st':state
H2:beval st b = true
H3:st =[ c ]=> st'
H6:st' =[ WHILE b DO c END ]=> st2

st = st2
b:bexp
st:state
c:com
H:beval st b = false
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st':state
H2:false = true
H3:st =[ c ]=> st'
H6:st' =[ WHILE b DO c END ]=> st2

st = st2
discriminate H2.
st', st'':state
b:bexp
c:com
st2:state
E2:st2 =[ WHILE b DO c END ]=> st2
E1_1:st2 =[ c ]=> st'
H:beval st2 b = true
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st2 =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
H4:beval st2 b = false

st'' = st2
st', st'':state
b:bexp
c:com
st2:state
E2:st2 =[ WHILE b DO c END ]=> st2
E1_1:st2 =[ c ]=> st'
H:beval st2 b = true
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st2 =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
H4:true = false

st'' = st2
discriminate H4.
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st'0:state
H2:beval st b = true
H3:st =[ c ]=> st'0
H6:st'0 =[ WHILE b DO c END ]=> st2

st'' = st2
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st'0:state
H2:beval st b = true
H3:st =[ c ]=> st'0
H6:st'0 =[ WHILE b DO c END ]=> st2

st' = st'0
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st'0:state
H2:beval st b = true
H3:st =[ c ]=> st'0
H6:st'0 =[ WHILE b DO c END ]=> st2
EQ1:st' = st'0
st'' = st2
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st'0:state
H2:beval st b = true
H3:st =[ c ]=> st'0
H6:st'0 =[ WHILE b DO c END ]=> st2

st' = st'0
apply IHE1_1; assumption.
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
st'0:state
H2:beval st b = true
H3:st =[ c ]=> st'0
H6:st'0 =[ WHILE b DO c END ]=> st2
EQ1:st' = st'0

st'' = st2
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
H2:beval st b = true
H6:st' =[ WHILE b DO c END ]=> st2
H3:st =[ c ]=> st'

st'' = st2
st, st', st'':state
b:bexp
c:com
H:beval st b = true
E1_1:st =[ c ]=> st'
E1_2:st' =[ WHILE b DO c END ]=> st''
IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0
IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0
st2:state
E2:st =[ WHILE b DO c END ]=> st2
H2:beval st b = true
H6:st' =[ WHILE b DO c END ]=> st2
H3:st =[ c ]=> st'

st' =[ WHILE b DO c END ]=> st2
assumption. Qed. (* ################################################################# *)

Reasoning About Imp Programs

We'll get deeper into more systematic and powerful techniques for reasoning about Imp programs in Programming Language Foundations, but we can get some distance just working with the bare definitions. This section explores some examples.

forall (st : string -> nat) (n : nat) (st' : state), st X = n -> st =[ plus2 ]=> st' -> st' X = n + 2

forall (st : string -> nat) (n : nat) (st' : state), st X = n -> st =[ plus2 ]=> st' -> st' X = n + 2
st:string -> nat
n:nat
st':state
HX:st X = n
Heval:st =[ plus2 ]=> st'

st' X = n + 2
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.
  
st:string -> nat
n:nat
st':state
HX:st X = n
Heval:st =[ plus2 ]=> st'
st0:state
a1:aexp
n0:nat
x:string
H3:aeval st (X + 2) = n0
H:x = X
H1:a1 = (X + 2)%imp
H0:st0 = st
H2:(X !-> n0; st) = st'

(X !-> n0; st) X = n + 2
st:string -> nat
Heval:st =[ plus2 ]=> (X !-> aeval st (X + 2); st)

(X !-> aeval st (X + 2); st) X = st X + 2
st:string -> nat

(X !-> aeval st (X + 2); st) X = st X + 2
st:string -> nat

(X !-> st X + 2; st) X = st X + 2
apply t_update_eq. Qed.

Exercise: 3 stars, standard, recommended (XtimesYinZ_spec)

State and prove a specification of XtimesYinZ.
(* FILL IN HERE *)

(* Do not modify the following line: *)
Definition manual_grade_for_XtimesYinZ_spec : option (nat*string) := None.

Exercise: 3 stars, standard, recommended (loop_never_stops)


forall st st' : state, ~ st =[ loop ]=> st'

forall st st' : state, ~ st =[ loop ]=> st'
st, st':state
contra:st =[ loop ]=> st'

False
st, st':state
contra:st =[ WHILE true DO SKIP END ]=> st'

False
st, st':state
loopdef:com
Heqloopdef:loopdef = (WHILE true DO SKIP END)%imp
contra:st =[ loopdef ]=> st'

False
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 discriminate).
  (* FILL IN HERE *) Admitted.

Exercise: 3 stars, standard (no_whiles_eqv)

Consider the following function:
Open Scope imp_scope.
Fixpoint no_whiles (c : com) : bool :=
  match c with
  | SKIP =>
      true
  | _ ::= _ =>
      true
  | c1 ;; c2 =>
      andb (no_whiles c1) (no_whiles c2)
  | TEST _ THEN ct ELSE cf FI =>
      andb (no_whiles ct) (no_whiles cf)
  | WHILE _ DO _ END  =>
      false
  end.
Close Scope imp_scope.
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 *)
.


forall c : com, no_whiles c = true <-> no_whilesR c

forall c : com, no_whiles c = true <-> no_whilesR c
(* FILL IN HERE *) Admitted.

Exercise: 4 stars, standard (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 *)

(* Do not modify the following line: *)
Definition manual_grade_for_no_whiles_terminating : option (nat*string) := None.
(* ################################################################# *)

Additional Exercises

Exercise: 3 stars, standard (stack_compiler)

Old 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 written 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 goal of this exercise is to write a small compiler that translates aexps into stack machine instructions.
The instruction set for our stack language will consist of the following instructions:
Inductive sinstr : Type :=
| SPush (n : nat)
| SLoad (x : string)
| SPlus
| SMinus
| SMult.
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.
s_execute:state -> list nat -> list sinstr -> list nat
st:state
stack:list nat
prog:list sinstr

list nat
Admitted.

s_execute empty_st [] [SPush 5; SPush 3; SPush 1; SMinus] = [2; 5]
(* FILL IN HERE *) Admitted.

s_execute (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.
s_compile:aexp -> list sinstr
e:aexp

list sinstr
Admitted.
After you've defined s_compile, prove the following to test that it works.

s_compile (X - 2 * 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!)
Prove the following theorem. You will need to start by stating a more general lemma to get a usable induction hypothesis; the main theorem will then be a simple corollary of this lemma.

forall (st : state) (e : aexp), s_execute st [] (s_compile e) = [aeval st e]

forall (st : state) (e : aexp), s_execute st [] (s_compile e) = [aeval st e]
(* FILL IN HERE *) Admitted.

Exercise: 3 stars, standard, 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.
Write an alternate version of beval that performs short-circuit evaluation of BAnd in this manner, and prove that it is equivalent to beval. (N.b. This is only true because expression evaluation in Imp is rather simple. In a bigger language where evaluating an expression might diverge, the short-circuiting BAnd would not be equivalent to the original, since it would make more programs terminate.)
(* FILL IN HERE 

    [] *)

Module BreakImp.

Exercise: 4 stars, advanced (break_imp)

Imperative languages like C and Java often include a break or similar statement for interrupting the execution of loops. In this exercise we consider how to add break to Imp. First, we need to enrich the language of commands with an additional case.
Inductive com : Type :=
  | CSkip
  | CBreak                        (* <--- NEW *)
  | CAss (x : string) (a : aexp)
  | CSeq (c1 c2 : com)
  | CIf (b : bexp) (c1 c2 : com)
  | CWhile (b : bexp) (c : 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 "'TEST' 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...
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.
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:
Inductive result : Type :=
  | SContinue
  | SBreak.

Reserved Notation "st '=[' c ']=>' st' '/' s"
         (at level 40, st' at next level).
Intuitively, st =[ c ]=> st' / s 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 "st =[ c ]=> st' / s" relation is very similar to the one we gave above for the regular evaluation relation (st =[ c ]=> st') -- we just need to handle the termination signals appropriately:
Based on the above description, complete the definition of the ceval relation.
Inductive ceval : com -> state -> result -> state -> Prop :=
  | E_Skip : forall st,
      st =[ CSkip ]=> st / SContinue
  (* FILL IN HERE *)

  where "st '=[' c ']=>' st' '/' s" := (ceval c st s st').
Now prove the following properties of your definition of ceval:

forall (c : com) (st st' : state) (s : result), st =[ BREAK;; c ]=> st' / s -> st = st'

forall (c : com) (st st' : state) (s : result), st =[ BREAK;; c ]=> st' / s -> st = st'
(* FILL IN HERE *) Admitted.

forall (b : bexp) (c : com) (st st' : state) (s : result), st =[ WHILE b DO c END ]=> st' / s -> s = SContinue

forall (b : bexp) (c : com) (st st' : state) (s : result), st =[ WHILE b DO c END ]=> st' / s -> s = SContinue
(* FILL IN HERE *) Admitted.

forall (b : bexp) (c : com) (st st' : state), beval st b = true -> st =[ c ]=> st' / SBreak -> st =[ WHILE b DO c END ]=> st' / SContinue

forall (b : bexp) (c : com) (st st' : state), beval st b = true -> st =[ c ]=> st' / SBreak -> st =[ WHILE b DO c END ]=> st' / SContinue
(* FILL IN HERE *) Admitted.

Exercise: 3 stars, advanced, optional (while_break_true)


forall (b : bexp) (c : com) (st st' : state), st =[ WHILE b DO c END ]=> st' / SContinue -> beval st' b = true -> exists st'' : state, st'' =[ c ]=> st' / SBreak

forall (b : bexp) (c : com) (st st' : state), st =[ WHILE b DO c END ]=> st' / SContinue -> beval st' b = true -> exists st'' : state, st'' =[ c ]=> st' / SBreak
(* FILL IN HERE *) Admitted.

Exercise: 4 stars, advanced, optional (ceval_deterministic)


forall (c : com) (st st1 st2 : state) (s1 s2 : result), st =[ c ]=> st1 / s1 -> st =[ c ]=> st2 / s2 -> st1 = st2 /\ s1 = s2

forall (c : com) (st st1 st2 : state) (s1 s2 : result), st =[ c ]=> st1 / s1 -> st =[ c ]=> st2 / s2 -> st1 = st2 /\ s1 = s2
(* FILL IN HERE *) Admitted.
End BreakImp.

Exercise: 4 stars, standard, 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.
A for loop should be parameterized by (a) a statement executed initially, (b) a test that is run on each iteration of the loop to determine whether the loop should continue, (c) a statement executed at the end of each loop iteration, and (d) a statement that makes up the body of the loop. (You don't need to worry about making up a concrete Notation for for loops, but feel free to play with this too if you like.)
(* FILL IN HERE 

    [] *)


(* Wed Jan 9 12:02:46 EST 2019 *)