ImpSimple Imperative Programs
Z ::= X;;
Y ::= 1;;
WHILE ~(Z = 0) DO
Y ::= Y * Z;;
Z ::= Z - 1
END
Y ::= 1;;
WHILE ~(Z = 0) DO
Y ::= Y * Z;;
Z ::= Z - 1
END
Arithmetic and Boolean Expressions
Abstract syntax trees for 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).
| 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).
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.
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.
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.
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.
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.
Theorem optimize_0plus_sound: ∀a,
aeval (optimize_0plus a) = aeval a.
Proof.
intros a. induction a.
- (* ANum *) reflexivity.
- (* APlus *) destruct a1 eqn:Ea1.
+ (* a1 = ANum n *) destruct n eqn:En.
* (* 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.
intros a. induction a.
- (* ANum *) reflexivity.
- (* APlus *) destruct a1 eqn:Ea1.
+ (* a1 = ANum n *) destruct n eqn:En.
* (* 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 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.
Proof.
intros P HP.
try reflexivity. (* Just reflexivity would have failed. *)
apply HP. (* We can still finish the proof in some other way. *)
Qed.
The ; Tactical (Simple Form)
Lemma foo : ∀n, 0 <=? n = true.
Proof.
intros.
destruct n.
(* Leaves two subgoals, which are discharged identically... *)
- (* n=0 *) simpl. reflexivity.
- (* n=Sn' *) simpl. reflexivity.
Qed.
Proof.
intros.
destruct n.
(* Leaves two subgoals, which are discharged identically... *)
- (* n=0 *) simpl. reflexivity.
- (* n=Sn' *) simpl. reflexivity.
Qed.
Lemma foo' : ∀n, 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.
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 eqn:Ea1;
(* 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 eqn:En;
simpl; rewrite IHa2; reflexivity. Qed.
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 eqn:Ea1;
(* 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 eqn:En;
simpl; rewrite IHa2; reflexivity. Qed.
The repeat Tactical
Theorem In10 : In 10 [1;2;3;4;5;6;7;8;9;10].
Proof.
repeat (try (left; reflexivity); right).
Qed.
Proof.
repeat (try (left; reflexivity); right).
Qed.
Defining New Tactic Notations
- Tactic Notation: syntax extension for tactics (good for simple macros)
- Ltac: scripting language for tactics (good for more sophisticated proof engineering)
- OCaml tactic scripting API (for wizards)
Tactic Notation "simpl_and_try" tactic(c) :=
simpl;
try c.
simpl;
try c.
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 ordering (≤), 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.
m + n ≤ n + o ∧ o + 3 = p + 3 →
m ≤ p.
Proof.
intros. omega.
Qed.
A Few More Handy Tactics
- clear H: Delete hypothesis H from the context.
- subst x: For a variable 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 (where x is a variable).
- 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
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).
| 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).
Notation "e '\\' n"
:= (aevalR e n)
(at level 50, left associativity)
: type_scope.
:= (aevalR e n)
(at level 50, left associativity)
: type_scope.
If we "reserve" the notation in advance, we can even use it
in the definition:
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.
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 |
There is nothing very deep going on here:
- the rule name corresponds to a constructor name
- above the line are premises
- below the line is conclusion
- metavariables like e1 and n1 are implicitly universally quantified
- the whole collection of rules is implicitly wrapped in an Inductive (sometimes we write this slightly more explicitly, as "...closed under these rules...")
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.
Computational vs. Relational Definitions
Adding division
Inductive aexp : Type :=
| ANum (n : nat)
| APlus (a1 a2 : aexp)
| AMinus (a1 a2 : aexp)
| AMult (a1 a2 : aexp)
| ADiv (a1 a2 : aexp). (* <--- NEW *)
| 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.
Adding division, relationally
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.
Adding Nondeterminism
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).
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...
Adding nondeterminism, relationally
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.
Tradeoffs
On the other hand, functional definitions can often be more convenient:
- 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.
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
States
Definition state := total_map nat.
Syntax
Inductive aexp : Type :=
| ANum (n : nat)
| AId (x : string) (* <--- NEW *)
| APlus (a1 a2 : aexp)
| AMinus (a1 a2 : aexp)
| AMult (a1 a2 : aexp).
| 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".
Definition X : string := "X".
Definition Y : string := "Y".
Definition Z : string := "Z".
Notations
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 interpretations 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.
Bind Scope imp_scope with aexp.
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.
Coercion ANum : nat >-> aexp.
Definition bool_to_bexp (b : bool) : bexp :=
if b then BTrue else BFalse.
Coercion bool_to_bexp : bool >-> bexp.
Bind Scope imp_scope with aexp.
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.
We can now write 3 + (X * 2) instead of APlus 3 (AMult X 2),
and true && ~(X ≤ 4) instead of BAnd true (BNot (BLe X 4)).
Definition example_aexp : aexp := 3 + (X * 2).
Definition example_bexp : bexp := true && ~(X ≤ 4).
Definition example_bexp : bexp := true && ~(X ≤ 4).
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.
Print example_bexp.
(* ===> example_bexp = bool_to_bexp true && ~ (AId X <= ANum 4) *)
Unset Printing Coercions.
Print example_bexp.
(* ===> example_bexp = bool_to_bexp true && ~ (AId X <= ANum 4) *)
Unset Printing Coercions.
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.
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 "x '!->' v" := (t_update empty_st x v) (at level 100).
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.
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.
Imp Factorial in Coq
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.
(Z ::= X;;
Y ::= 1;;
WHILE ~(Z = 0) DO
Y ::= Y * Z;;
Z ::= Z - 1
END)%imp.
Evaluating Commands
Evaluation as a Function (Failed Attempt)
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.
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.
Nontermination leads to Inconsistency
Consider the following "proof object":Fixpoint loop_false (n : nat) : False := loop_false n.
Evaluation as a Relation
Operational Semantics
(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'' |
Reserved Notation "st '=[' c ']⇒' st'"
(at level 40).
Inductive ceval : com → state → state → Prop :=
| E_Skip : ∀st,
st =[ SKIP ]⇒ st
| E_Ass : ∀st a1 n x,
aeval st a1 = n →
st =[ x ::= a1 ]⇒ (x !-> n ; st)
| E_Seq : ∀c1 c2 st st' st'',
st =[ c1 ]⇒ st' →
st' =[ c2 ]⇒ st'' →
st =[ c1 ;; c2 ]⇒ st''
| E_IfTrue : ∀st st' b c1 c2,
beval st b = true →
st =[ c1 ]⇒ st' →
st =[ TEST b THEN c1 ELSE c2 FI ]⇒ st'
| E_IfFalse : ∀st st' b c1 c2,
beval st b = false →
st =[ c2 ]⇒ st' →
st =[ TEST b THEN c1 ELSE c2 FI ]⇒ st'
| E_WhileFalse : ∀b st c,
beval st b = false →
st =[ WHILE b DO c END ]⇒ st
| E_WhileTrue : ∀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.
Example ceval_example1:
empty_st =[
X ::= 2;;
TEST X ≤ 1
THEN Y ::= 3
ELSE Z ::= 4
FI
]⇒ (Z !-> 4 ; X !-> 2).
Proof.
(* We must supply the intermediate state *)
apply E_Seq with (X !-> 2).
- (* assignment command *)
apply E_Ass. reflexivity.
- (* if command *)
apply E_IfFalse.
reflexivity.
apply E_Ass. reflexivity.
Qed.
empty_st =[
X ::= 2;;
TEST X ≤ 1
THEN Y ::= 3
ELSE Z ::= 4
FI
]⇒ (Z !-> 4 ; X !-> 2).
Proof.
(* We must supply the intermediate state *)
apply E_Seq with (X !-> 2).
- (* assignment command *)
apply E_Ass. reflexivity.
- (* if command *)
apply E_IfFalse.
reflexivity.
apply E_Ass. reflexivity.
Qed.
Is the following proposition provable?
(2) No
(3) Not sure
∀(c : com) (st st' : state),
st =[ SKIP ;; c ]⇒ st' →
st =[ c ]⇒ st'
(1) Yes
st =[ SKIP ;; c ]⇒ st' →
st =[ c ]⇒ st'
Is the following proposition provable?
(2) No
(3) Not sure
∀(c1 c2 : com) (st st' : state),
st =[ c1;;c2 ]⇒ st' →
st =[ c1 ]⇒ st →
st =[ c2 ]⇒ st'
(1) Yes
st =[ c1;;c2 ]⇒ st' →
st =[ c1 ]⇒ st →
st =[ c2 ]⇒ st'
Is the following proposition provable?
(2) No
(3) Not sure
∀(b : bexp) (c : com) (st st' : state),
st =[ TEST b THEN c ELSE c FI ]⇒ st' →
st =[ c ]⇒ st'
(1) Yes
st =[ TEST b THEN c ELSE c FI ]⇒ st' →
st =[ c ]⇒ st'
Is the following proposition provable?
(2) No
(3) Not sure
∀b : bexp,
(∀st, beval st b = true) →
∀(c : com) (st : state),
~(∃st', st =[ WHILE b DO c END ]⇒ st')
(1) Yes
(∀st, beval st b = true) →
∀(c : com) (st : state),
~(∃st', st =[ WHILE b DO c END ]⇒ st')
Is the following proposition provable?
(2) No
(3) Not sure
∀(b : bexp) (c : com) (st : state),
~(∃st', st =[ WHILE b DO c END ]⇒ st') →
∀st'', beval st'' b = true
(1) Yes
~(∃st', st =[ WHILE b DO c END ]⇒ st') →
∀st'', beval st'' b = true
Theorem ceval_deterministic: ∀c st st1 st2,
st =[ c ]⇒ st1 →
st =[ c ]⇒ 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, b evaluates to true *)
apply IHE1. assumption.
- (* E_IfTrue, b evaluates to false (contradiction) *)
rewrite H in H5. discriminate H5.
- (* E_IfFalse, b evaluates to true (contradiction) *)
rewrite H in H5. discriminate H5.
- (* E_IfFalse, b evaluates to false *)
apply IHE1. assumption.
- (* E_WhileFalse, b evaluates to false *)
reflexivity.
- (* E_WhileFalse, b evaluates to true (contradiction) *)
rewrite H in H2. discriminate H2.
- (* E_WhileTrue, b evaluates to false (contradiction) *)
rewrite H in H4. discriminate H4.
- (* E_WhileTrue, b 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, b evaluates to true *)
apply IHE1. assumption.
- (* E_IfTrue, b evaluates to false (contradiction) *)
rewrite H in H5. discriminate H5.
- (* E_IfFalse, b evaluates to true (contradiction) *)
rewrite H in H5. discriminate H5.
- (* E_IfFalse, b evaluates to false *)
apply IHE1. assumption.
- (* E_WhileFalse, b evaluates to false *)
reflexivity.
- (* E_WhileFalse, b evaluates to true (contradiction) *)
rewrite H in H2. discriminate H2.
- (* E_WhileTrue, b evaluates to false (contradiction) *)
rewrite H in H4. discriminate H4.
- (* E_WhileTrue, b evaluates to true *)
assert (st' = st'0) as EQ1.
{ (* Proof of assertion *) apply IHE1_1; assumption. }
subst st'0.
apply IHE1_2. assumption. Qed.