ImpSimple Imperative Programs
Z := X;
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)
| BNeq (a1 a2 : aexp)
| BLe (a1 a2 : aexp)
| BGt (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)
| BNeq (a1 a2 : aexp)
| BLe (a1 a2 : aexp)
| BGt (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)
| BNeq a1 a2 ⇒ negb ((aeval a1) =? (aeval a2))
| BLe a1 a2 ⇒ (aeval a1) <=? (aeval a2)
| BGt a1 a2 ⇒ negb ((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)
| BNeq a1 a2 ⇒ negb ((aeval a1) =? (aeval a2))
| BLe a1 a2 ⇒ (aeval a1) <=? (aeval a2)
| BGt a1 a2 ⇒ negb ((aeval a1) <=? (aeval a2))
| BNot b1 ⇒ negb (beval b1)
| BAnd b1 b2 ⇒ andb (beval b1) (beval b2)
end.
What does the following expression evaluate to?
aeval (APlus (ANum 3) (AMinus (ANum 4) (ANum 1)))
(1) true
(2) false
(3) 0
(4) 3
(5) 6
aeval (APlus (ANum 3) (AMinus (ANum 4) (ANum 1)))
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.
Example test_optimize_0plus:
optimize_0plus (APlus (ANum 2)
(APlus (ANum 0)
(APlus (ANum 0) (ANum 1))))
= APlus (ANum 2) (ANum 1).
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.
Example test_optimize_0plus:
optimize_0plus (APlus (ANum 2)
(APlus (ANum 0)
(APlus (ANum 0) (ANum 1))))
= APlus (ANum 2) (ANum 1).
Proof. reflexivity. Qed.
Theorem optimize_0plus_sound: ∀ a,
aeval (optimize_0plus a) = aeval 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 silly1 : ∀ (P : Prop), P → P.
Proof.
intros P HP.
try reflexivity. (* Plain reflexivity would have failed. *)
apply HP. (* We can still finish the proof in some other way. *)
Qed.
Theorem silly2 : ∀ ae, aeval ae = aeval ae.
Proof.
try reflexivity. (* This just does reflexivity. *)
Qed.
Proof.
intros P HP.
try reflexivity. (* Plain reflexivity would have failed. *)
apply HP. (* We can still finish the proof in some other way. *)
Qed.
Theorem silly2 : ∀ ae, aeval ae = aeval ae.
Proof.
try reflexivity. (* This just does reflexivity. *)
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.
repeat can loop forever.
Theorem repeat_loop : ∀ (m n : nat),
m + n = n + m.
Proof.
intros m n.
(* Uncomment the next line to see the infinite loop occur. You will
then need to interrupt Coq to make it listen to you again. (In
Proof General, C-c C-c does this.) *)
(* repeat rewrite Nat.add_comm. *)
Admitted.
m + n = n + m.
Proof.
intros m n.
(* Uncomment the next line to see the infinite loop occur. You will
then need to interrupt Coq to make it listen to you again. (In
Proof General, C-c C-c does this.) *)
(* repeat rewrite Nat.add_comm. *)
Admitted.
Defining New Tactics
- Ltac: scripting language for tactics (good for more sophisticated proof engineering)
- Tactic Notation for defining tactics with custom concrete syntax
- OCaml tactic scripting API (for wizards)
Ltac invert H :=
inversion H; subst; clear H.
inversion H; subst; clear H.
The lia 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 >), 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. lia.
Qed.
Example add_comm__lia : ∀ m n,
m + n = n + m.
Proof.
intros. lia.
Qed.
Example add_assoc__lia : ∀ m n p,
m + (n + p) = m + n + p.
Proof.
intros. lia.
Qed.
m + n ≤ n + o ∧ o + 3 = p + 3 →
m ≤ p.
Proof.
intros. lia.
Qed.
Example add_comm__lia : ∀ m n,
m + n = n + m.
Proof.
intros. lia.
Qed.
Example add_assoc__lia : ∀ m n p,
m + (n + p) = m + n + p.
Proof.
intros. lia.
Qed.
A Few More Handy Tactics
- clear H: Delete hypothesis H from the context.
- subst x: Given 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, solve the goal.
- contradiction: Try to find a hypothesis H in the 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 : nat) :
aevalR (ANum n) n
| E_APlus (e1 e2 : aexp) (n1 n2 : nat) :
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (APlus e1 e2) (n1 + n2)
| E_AMinus (e1 e2 : aexp) (n1 n2 : nat) :
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (AMinus e1 e2) (n1 - n2)
| E_AMult (e1 e2 : aexp) (n1 n2 : nat) :
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (AMult e1 e2) (n1 × n2).
| E_ANum (n : nat) :
aevalR (ANum n) n
| E_APlus (e1 e2 : aexp) (n1 n2 : nat) :
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (APlus e1 e2) (n1 + n2)
| E_AMinus (e1 e2 : aexp) (n1 n2 : nat) :
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (AMinus e1 e2) (n1 - n2)
| E_AMult (e1 e2 : aexp) (n1 n2 : nat) :
aevalR e1 n1 →
aevalR e2 n2 →
aevalR (AMult e1 e2) (n1 × n2).
Notation "e '==>' n"
:= (aevalR e n)
(at level 90, left associativity)
: type_scope.
:= (aevalR e n)
(at level 90, left associativity)
: 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) ...can be written like this as an inference rule:
e1 ==> n1 | |
e2 ==> n2 | (E_APlus) |
APlus e1 e2 ==> n1+n2 |
There is nothing very deep going on here:
- a group of inference rules corresponds to a single Inductive definition
- each rule's 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
(E_ANum) | |
ANum n ==> n |
e1 ==> n1 | |
e2 ==> n2 | (E_APlus) |
APlus e1 e2 ==> n1+n2 |
e1 ==> n1 | |
e2 ==> n2 | (E_AMinus) |
AMinus e1 e2 ==> n1-n2 |
e1 ==> n1 | |
e2 ==> n2 | (E_AMult) |
AMult e1 e2 ==> n1*n2 |
Equivalence of the Definitions
Theorem aevalR_iff_aeval : ∀ a n,
(a ==> n) ↔ aeval 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.
Again, we can make the proof quite a bit shorter using some
tacticals.
Theorem aevalR_iff_aeval' : ∀ a n,
(a ==> n) ↔ aeval a = n.
Proof.
(* WORK IN CLASS *) Admitted.
(a ==> n) ↔ aeval a = n.
Proof.
(* WORK IN CLASS *) Admitted.
Computational vs. Relational Definitions
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 *)
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) : (* <----- NEW *)
(a1 ==> n1) → (a2 ==> n2) → (n2 > 0) →
(mult n2 n3 = n1) → (ADiv a1 a2) ==> n3
where "a '==>' n" := (aevalR a n) : type_scope.
Notice that this evaluation relation corresponds to a partial
function: There are some inputs for which it does not specify an
output.
Inductive aexp : Type :=
| AAny (* <--- NEW *)
| ANum (n : nat)
| APlus (a1 a2 : aexp)
| AMinus (a1 a2 : aexp)
| AMult (a1 a2 : aexp).
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.
| 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
- Functional: take advantage of computation.
- Relational: more (easily) expressive.
- Best of both worlds: define both and prove them equivalent.
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
Definition example_aexp : aexp := <{ 3 + (X × 2) }>.
Definition example_bexp : bexp := <{ true && ~(X ≤ 4) }>.
Definition example_bexp : bexp := <{ true && ~(X ≤ 4) }>.
Fixpoint aeval (st : state) (* <--- NEW *)
(a : aexp) : nat :=
match a with
| ANum n ⇒ n
| AId x ⇒ st x (* <--- NEW *)
| <{a1 + a2}> ⇒ (aeval st a1) + (aeval st a2)
| <{a1 - a2}> ⇒ (aeval st a1) - (aeval st a2)
| <{a1 × a2}> ⇒ (aeval st a1) × (aeval st a2)
end.
Fixpoint beval (st : state) (* <--- NEW *)
(b : bexp) : bool :=
match b with
| <{true}> ⇒ true
| <{false}> ⇒ false
| <{a1 = a2}> ⇒ (aeval st a1) =? (aeval st a2)
| <{a1 ≠ a2}> ⇒ negb ((aeval st a1) =? (aeval st a2))
| <{a1 ≤ a2}> ⇒ (aeval st a1) <=? (aeval st a2)
| <{a1 > a2}> ⇒ negb ((aeval st a1) <=? (aeval st a2))
| <{~ b1}> ⇒ negb (beval st b1)
| <{b1 && b2}> ⇒ andb (beval st b1) (beval st b2)
end.
(a : aexp) : nat :=
match a with
| ANum n ⇒ n
| AId x ⇒ st x (* <--- NEW *)
| <{a1 + a2}> ⇒ (aeval st a1) + (aeval st a2)
| <{a1 - a2}> ⇒ (aeval st a1) - (aeval st a2)
| <{a1 × a2}> ⇒ (aeval st a1) × (aeval st a2)
end.
Fixpoint beval (st : state) (* <--- NEW *)
(b : bexp) : bool :=
match b with
| <{true}> ⇒ true
| <{false}> ⇒ false
| <{a1 = a2}> ⇒ (aeval st a1) =? (aeval st a2)
| <{a1 ≠ a2}> ⇒ negb ((aeval st a1) =? (aeval st a2))
| <{a1 ≤ a2}> ⇒ (aeval st a1) <=? (aeval st a2)
| <{a1 > a2}> ⇒ negb ((aeval st a1) <=? (aeval st a2))
| <{~ b1}> ⇒ negb (beval st b1)
| <{b1 && b2}> ⇒ andb (beval st b1) (beval st b2)
end.
We can use our notation for total maps in the specific case of states -- i.e., we write the empty state as (_ !-> 0).
Definition empty_st := (_ !-> 0).
Also, we can add a notation for a "singleton state" with just one
variable bound to a value.
Notation "x '!->' v" := (x !-> v ; empty_st) (at level 100).
Syntax
c := skip
| x := a
| c ; c
| if b then c else c end
| while b do c end
Inductive com : Type :=
| CSkip
| CAsgn (x : string) (a : aexp)
| CSeq (c1 c2 : com)
| CIf (b : bexp) (c1 c2 : com)
| CWhile (b : bexp) (c : com).
| CSkip
| CAsgn (x : string) (a : aexp)
| CSeq (c1 c2 : com)
| CIf (b : bexp) (c1 c2 : com)
| CWhile (b : bexp) (c : com).
As we did for expressions, we can use a few Notation declarations to make reading and writing Imp programs more convenient.
For example, here is the factorial function again, written as a formal Coq definition. When this command terminates, the variable Y will contain the factorial of the initial value of X.
Definition fact_in_coq : com :=
<{ Z := X;
Y := 1;
while Z ≠ 0 do
Y := Y × Z;
Z := Z - 1
end }>.
<{ Z := X;
Y := 1;
while Z ≠ 0 do
Y := Y × Z;
Z := Z - 1
end }>.
Definition plus2 : com :=
<{ X := X + 2 }>.
Definition XtimesYinZ : com :=
<{ Z := X × Y }>.
<{ X := X + 2 }>.
Definition XtimesYinZ : com :=
<{ Z := X × Y }>.
Definition subtract_slowly_body : com :=
<{ Z := Z - 1 ;
X := X - 1 }>.
Definition subtract_slowly : com :=
<{ while X ≠ 0 do
subtract_slowly_body
end }>.
Definition subtract_3_from_5_slowly : com :=
<{ X := 3 ;
Z := 5 ;
subtract_slowly }>.
<{ Z := Z - 1 ;
X := X - 1 }>.
Definition subtract_slowly : com :=
<{ while X ≠ 0 do
subtract_slowly_body
end }>.
Definition subtract_3_from_5_slowly : com :=
<{ X := 3 ;
Z := 5 ;
subtract_slowly }>.
Definition loop : com :=
<{ while true do
skip
end }>.
<{ while true do
skip
end }>.
Evaluating Commands
Evaluation as a Function (Failed Attempt)
Fixpoint ceval_fun_no_while (st : state) (c : com) : state :=
match c with
| <{ skip }> ⇒
st
| <{ x := a }> ⇒
(x !-> (aeval st a) ; st)
| <{ c1 ; c2 }> ⇒
let st' := ceval_fun_no_while st c1 in
ceval_fun_no_while st' c2
| <{ if b then c1 else c2 end}> ⇒
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.
match c with
| <{ skip }> ⇒
st
| <{ x := a }> ⇒
(x !-> (aeval st a) ; st)
| <{ c1 ; c2 }> ⇒
let st' := ceval_fun_no_while st c1 in
ceval_fun_no_while st' c2
| <{ if b then c1 else c2 end}> ⇒
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.
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 a = n | (E_Asgn) |
st =[ x := a ]=> (x !-> n ; st) |
st =[ c1 ]=> st' | |
st' =[ c2 ]=> st'' | (E_Seq) |
st =[ c1;c2 ]=> st'' |
beval st b = true | |
st =[ c1 ]=> st' | (E_IfTrue) |
st =[ if b then c1 else c2 end ]=> st' |
beval st b = false | |
st =[ c2 ]=> st' | (E_IfFalse) |
st =[ if b then c1 else c2 end ]=> 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'' |
Inductive ceval : com → state → state → Prop :=
| E_Skip : ∀ st,
st =[ skip ]=> st
| E_Asgn : ∀ st a n x,
aeval st a = n →
st =[ x := a ]=> (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 =[ if b then c1 else c2 end]=> st'
| E_IfFalse : ∀ st st' b c1 c2,
beval st b = false →
st =[ c2 ]=> st' →
st =[ if b then c1 else c2 end]=> 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').
Example ceval_example1:
empty_st =[
X := 2;
if (X ≤ 1)
then Y := 3
else Z := 4
end
]=> (Z !-> 4 ; X !-> 2).
Proof.
(* We must supply the intermediate state *)
apply E_Seq with (X !-> 2).
- (* assignment command *)
apply E_Asgn. reflexivity.
- (* if command *)
apply E_IfFalse.
+ reflexivity.
+ apply E_Asgn. reflexivity.
Qed.
empty_st =[
X := 2;
if (X ≤ 1)
then Y := 3
else Z := 4
end
]=> (Z !-> 4 ; X !-> 2).
Proof.
(* We must supply the intermediate state *)
apply E_Seq with (X !-> 2).
- (* assignment command *)
apply E_Asgn. reflexivity.
- (* if command *)
apply E_IfFalse.
+ reflexivity.
+ apply E_Asgn. reflexivity.
Qed.
What sorts of things might we want to prove using these
definitions?
Here are some simple examples...
Is the following proposition provable?
∀ (c : com) (st st' : state),
st =[ skip ; c ]=> st' →
st =[ c ]=> st' (1) Yes
(2) No
(3) Not sure
∀ (c : com) (st st' : state),
st =[ skip ; c ]=> st' →
st =[ c ]=> st' (1) Yes
Is the following proposition provable?
∀ (c1 c2 : com) (st st' : state),
st =[ c1 ; c2 ]=> st' →
st =[ c1 ]=> st →
st =[ c2 ]=> st' (1) Yes
(2) No
(3) Not sure
∀ (c1 c2 : com) (st st' : state),
st =[ c1 ; c2 ]=> st' →
st =[ c1 ]=> st →
st =[ c2 ]=> st' (1) Yes
Is the following proposition provable?
∀ (b : bexp) (c : com) (st st' : state),
st =[ if b then c else c end ]=> st' →
st =[ c ]=> st' (1) Yes
(2) No
(3) Not sure
∀ (b : bexp) (c : com) (st st' : state),
st =[ if b then c else c end ]=> st' →
st =[ c ]=> st' (1) Yes
Is the following proposition provable?
∀ b : bexp,
(∀ st, beval st b = true) →
∀ (c : com) (st : state),
~(∃ st', st =[ while b do c end ]=> st') (1) Yes
(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
Is the following proposition provable?
∀ (b : bexp) (c : com) (st : state),
~(∃ st', st =[ while b do c end ]=> st') →
∀ st'', beval st'' b = true (1) Yes
(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
Determinism of Evaluation
Theorem ceval_deterministic: ∀ c st st1 st2,
st =[ c ]=> st1 →
st =[ c ]=> st2 →
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_Asgn *) reflexivity.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
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.
- (* E_IfFalse, b evaluates to true (contradiction) *)
rewrite H in H5. discriminate.
- (* 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.
- (* E_WhileTrue, b evaluates to false (contradiction) *)
rewrite H in H4. discriminate.
- (* E_WhileTrue, b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
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_Asgn *) reflexivity.
- (* E_Seq *)
rewrite (IHE1_1 st'0 H1) in ×.
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.
- (* E_IfFalse, b evaluates to true (contradiction) *)
rewrite H in H5. discriminate.
- (* 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.
- (* E_WhileTrue, b evaluates to false (contradiction) *)
rewrite H in H4. discriminate.
- (* E_WhileTrue, b evaluates to true *)
rewrite (IHE1_1 st'0 H3) in ×.
apply IHE1_2. assumption. Qed.