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.
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. (* ################################################################# *)
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.
(* ================================================================= *)
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.
- The BNF is more informal -- for example, it gives some
suggestions about the surface syntax of expressions (like the
fact that the addition operation is written with an infix
+) while leaving other aspects of lexical analysis and
parsing (like the relative precedence of +, -, and
×, the use of parens to group subexpressions, etc.)
unspecified. Some additional information -- and human
intelligence -- would be required to turn this description
into a formal definition, e.g., for implementing a compiler.
- Conversely, the BNF version is lighter and easier to read.
Its informality makes it flexible, a big advantage in
situations like discussions at the blackboard, where
conveying general ideas is more important than getting every
detail nailed down precisely.
(* ================================================================= *)
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)) = 4reflexivity. Qed.aeval (APlus (ANum 2) (ANum 2)) = 4
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. (* ================================================================= *)
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)reflexivity. Qed.optimize_0plus (APlus (ANum 2) (APlus (ANum 0) (APlus (ANum 0) (ANum 1)))) = APlus (ANum 2) (ANum 1)
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 aforall a : aexp, aeval (optimize_0plus a) = aeval aa:aexpaeval (optimize_0plus a) = aeval an:nataeval (optimize_0plus (ANum n)) = aeval (ANum n)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (AMinus a1 a2)) = aeval (AMinus a1 a2)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (AMult a1 a2)) = aeval (AMult a1 a2)reflexivity.n:nataeval (optimize_0plus (ANum n)) = aeval (ANum n)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)a1, a2:aexpn:natEa1:a1 = ANum nIHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)a1, a2, a3, a4:aexpEa1:a1 = APlus a3 a4IHa1:aeval (optimize_0plus (APlus a3 a4)) = aeval (APlus a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (APlus a3 a4) a2)) = aeval (APlus (APlus a3 a4) a2)a1, a2, a3, a4:aexpEa1:a1 = AMinus a3 a4IHa1:aeval (optimize_0plus (AMinus a3 a4)) = aeval (AMinus a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (AMinus a3 a4) a2)) = aeval (APlus (AMinus a3 a4) a2)a1, a2, a3, a4:aexpEa1:a1 = AMult a3 a4IHa1:aeval (optimize_0plus (AMult a3 a4)) = aeval (AMult a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (AMult a3 a4) a2)) = aeval (APlus (AMult a3 a4) a2)a1, a2:aexpn:natEa1:a1 = ANum nIHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)a1, a2:aexpn:natEn:n = 0Ea1:a1 = ANum 0IHa1:aeval (optimize_0plus (ANum 0)) = aeval (ANum 0)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum 0) a2)) = aeval (APlus (ANum 0) a2)a1, a2:aexpn, n0:natEn:n = S n0Ea1:a1 = ANum (S n0)IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum (S n0)) a2)) = aeval (APlus (ANum (S n0)) a2)a1, a2:aexpn:natEn:n = 0Ea1:a1 = ANum 0IHa1:aeval (optimize_0plus (ANum 0)) = aeval (ANum 0)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum 0) a2)) = aeval (APlus (ANum 0) a2)apply IHa2.a1, a2:aexpn:natEn:n = 0Ea1:a1 = ANum 0IHa1:aeval (optimize_0plus (ANum 0)) = aeval (ANum 0)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a2) = aeval a2a1, a2:aexpn, n0:natEn:n = S n0Ea1:a1 = ANum (S n0)IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum (S n0)) a2)) = aeval (APlus (ANum (S n0)) a2)a1, a2:aexpn, n0:natEn:n = S n0Ea1:a1 = ANum (S n0)IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))IHa2:aeval (optimize_0plus a2) = aeval a2S (n0 + aeval (optimize_0plus a2)) = S (n0 + aeval a2)reflexivity.a1, a2:aexpn, n0:natEn:n = S n0Ea1:a1 = ANum (S n0)IHa1:aeval (optimize_0plus (ANum (S n0))) = aeval (ANum (S n0))IHa2:aeval (optimize_0plus a2) = aeval a2S (n0 + aeval a2) = S (n0 + aeval a2)a1, a2, a3, a4:aexpEa1:a1 = APlus a3 a4IHa1:aeval (optimize_0plus (APlus a3 a4)) = aeval (APlus a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (APlus a3 a4) a2)) = aeval (APlus (APlus a3 a4) a2)a1, a2, a3, a4:aexpEa1:a1 = APlus a3 a4IHa1:aeval (optimize_0plus (APlus a3 a4)) = aeval (APlus a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval 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 a2a1, a2, a3, a4:aexpEa1:a1 = APlus a3 a4IHa1:aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end = aeval a3 + aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval 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 a2a1, a2, a3, a4:aexpEa1:a1 = APlus a3 a4IHa1:aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end = aeval a3 + aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval a3 + aeval a4 + aeval (optimize_0plus a2) = aeval a3 + aeval a4 + aeval a2reflexivity.a1, a2, a3, a4:aexpEa1:a1 = APlus a3 a4IHa1:aeval match a3 with | ANum 0 => optimize_0plus a4 | _ => APlus (optimize_0plus a3) (optimize_0plus a4) end = aeval a3 + aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval a3 + aeval a4 + aeval a2 = aeval a3 + aeval a4 + aeval a2a1, a2, a3, a4:aexpEa1:a1 = AMinus a3 a4IHa1:aeval (optimize_0plus (AMinus a3 a4)) = aeval (AMinus a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (AMinus a3 a4) a2)) = aeval (APlus (AMinus a3 a4) a2)a1, a2, a3, a4:aexpEa1:a1 = AMinus a3 a4IHa1:aeval (optimize_0plus (AMinus a3 a4)) = aeval (AMinus a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a3) - aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 - aeval a4 + aeval a2a1, a2, a3, a4:aexpEa1:a1 = AMinus a3 a4IHa1:aeval (optimize_0plus a3) - aeval (optimize_0plus a4) = aeval a3 - aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a3) - aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 - aeval a4 + aeval a2a1, a2, a3, a4:aexpEa1:a1 = AMinus a3 a4IHa1:aeval (optimize_0plus a3) - aeval (optimize_0plus a4) = aeval a3 - aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval a3 - aeval a4 + aeval (optimize_0plus a2) = aeval a3 - aeval a4 + aeval a2reflexivity.a1, a2, a3, a4:aexpEa1:a1 = AMinus a3 a4IHa1:aeval (optimize_0plus a3) - aeval (optimize_0plus a4) = aeval a3 - aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval a3 - aeval a4 + aeval a2 = aeval a3 - aeval a4 + aeval a2a1, a2, a3, a4:aexpEa1:a1 = AMult a3 a4IHa1:aeval (optimize_0plus (AMult a3 a4)) = aeval (AMult a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (AMult a3 a4) a2)) = aeval (APlus (AMult a3 a4) a2)a1, a2, a3, a4:aexpEa1:a1 = AMult a3 a4IHa1:aeval (optimize_0plus (AMult a3 a4)) = aeval (AMult a3 a4)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a3) * aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 * aeval a4 + aeval a2a1, a2, a3, a4:aexpEa1:a1 = AMult a3 a4IHa1:aeval (optimize_0plus a3) * aeval (optimize_0plus a4) = aeval a3 * aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a3) * aeval (optimize_0plus a4) + aeval (optimize_0plus a2) = aeval a3 * aeval a4 + aeval a2a1, a2, a3, a4:aexpEa1:a1 = AMult a3 a4IHa1:aeval (optimize_0plus a3) * aeval (optimize_0plus a4) = aeval a3 * aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval a3 * aeval a4 + aeval (optimize_0plus a2) = aeval a3 * aeval a4 + aeval a2reflexivity.a1, a2, a3, a4:aexpEa1:a1 = AMult a3 a4IHa1:aeval (optimize_0plus a3) * aeval (optimize_0plus a4) = aeval a3 * aeval a4IHa2:aeval (optimize_0plus a2) = aeval a2aeval a3 * aeval a4 + aeval a2 = aeval a3 * aeval a4 + aeval a2a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (AMinus a1 a2)) = aeval (AMinus a1 a2)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a1) - aeval (optimize_0plus a2) = aeval a1 - aeval a2a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval a1 - aeval (optimize_0plus a2) = aeval a1 - aeval a2reflexivity.a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval a1 - aeval a2 = aeval a1 - aeval a2a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (AMult a1 a2)) = aeval (AMult a1 a2)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus a1) * aeval (optimize_0plus a2) = aeval a1 * aeval a2a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval a1 * aeval (optimize_0plus a2) = aeval a1 * aeval a2reflexivity. Qed. (* ################################################################# *)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval a1 * aeval a2 = aeval a1 * aeval a2
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 is Coq's term for tactics that take other tactics as
arguments -- "higher-order tactics," if you will.
(* ----------------------------------------------------------------- *)
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 aetry reflexivity. (* This just does [reflexivity]. *) Qed.forall ae : aexp, aeval ae = aeval aeforall P : Prop, P -> Pforall P : Prop, P -> PP:PropHP:PPapply HP. (* We can still finish the proof in some other way. *) Qed.P:PropHP:PP
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.
(* ----------------------------------------------------------------- *)
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) = trueforall n : nat, (0 <=? n) = truen:nat(0 <=? n) = true(* Leaves two subgoals, which are discharged identically... *)(0 <=? 0) = truen:nat(0 <=? S n) = true(0 <=? 0) = truereflexivity.true = truen:nat(0 <=? S n) = truereflexivity. Qed.n:nattrue = true
We can simplify this proof using the ; tactical:
forall n : nat, (0 <=? n) = trueforall 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.n:nat(0 <=? n) = true
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 aforall a : aexp, aeval (optimize_0plus a) = aeval aa:aexpaeval (optimize_0plus a) = aeval a(* ... but the remaining cases -- ANum and APlus -- are different: *)n:nataeval (optimize_0plus (ANum n)) = aeval (ANum n)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)reflexivity.n:nataeval (optimize_0plus (ANum n)) = aeval (ANum n)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 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:aexpn:natEa1:a1 = ANum nIHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)destruct n eqn:En; simpl; rewrite IHa2; reflexivity. Qed.a1, a2:aexpn:natEa1:a1 = ANum nIHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
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:
- Suppose a = ANum n for some n. We must show
- Suppose a = APlus a1 a2 for some a1 and a2. We must
show
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 aforall a : aexp, aeval (optimize_0plus a) = aeval aa:aexpaeval (optimize_0plus a) = aeval a(* The interesting case is when a = APlus a1 a2. *)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)a1, a2:aexpIHa1:aeval (optimize_0plus a1) = aeval a1IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus a1 a2)) = aeval (APlus a1 a2)n:nata2:aexpIHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)destruct n; simpl; rewrite IHa2; reflexivity. Qed. (* ----------------------------------------------------------------- *)n:nata2:aexpIHa1:aeval (optimize_0plus (ANum n)) = aeval (ANum n)IHa2:aeval (optimize_0plus a2) = aeval a2aeval (optimize_0plus (APlus (ANum n) a2)) = aeval (APlus (ANum n) a2)
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 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]repeat (try (left; reflexivity); right). Qed.In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]
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]repeat (right; try (left; reflexivity)). Qed.In 10 [1; 2; 3; 4; 5; 6; 7; 8; 9; 10]
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)
Admitted.optimize_0plus_b:bexp -> bexpb:bexpbexpforall b : bexp, beval (optimize_0plus_b b) = beval b(* FILL IN HERE *) Admitted.forall b : bexp, beval (optimize_0plus_b b) = beval b
☐
Exercise: 4 stars, standard, optional (optimize)
(* FILL IN HERE
[] *)
(* ================================================================= *)
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.
- The Tactic Notation idiom illustrated below gives a handy way
to define "shorthand tactics" that bundle several tactics into a
single command.
- For more sophisticated programming, Coq offers a built-in
language called Ltac with primitives that can examine and
modify the proof state. The details are a bit too complicated
to get into here (and it is generally agreed that Ltac is not
the most beautiful part of Coq's design!), but they can be found
in the reference manual and other books on Coq, and there are
many examples of Ltac definitions in the Coq standard library
that you can use as examples.
- There is also an OCaml API, which can be used to build tactics that access Coq's internal structures at a lower level, but this is seldom worth the trouble for ordinary Coq users.
Tactic Notation "simpl_and_try" tactic(c) :=
simpl;
try c.
This defines a new tactical called simpl_and_try that takes one
tactic c as an argument and is defined to be equivalent to the
tactic simpl; try c. Now writing "simpl_and_try reflexivity."
in a proof will be the same as writing "simpl; try reflexivity."
(* ================================================================= *)
The omega tactic 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.)
- 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 →,
forall m n o p : nat, m + n <= n + o /\ o + 3 = p + 3 -> m <= pforall m n o p : nat, m + n <= n + o /\ o + 3 = p + 3 -> m <= pomega. Qed.m, n, o, p:natH:m + n <= n + o /\ o + 3 = p + 3m <= p
(Note the From Coq Require Import omega.Omega. at the top of
the file.)
(* ================================================================= *)
Finally, here are some miscellaneous tactics that you may find
convenient.
We'll see examples of all of these as we go along.
- 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.
(* ################################################################# *)
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. (* ================================================================= *)
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
(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
(E_ANum) ANum n \\ n
(E_APlus) APlus e1 e2 \\ n1+n2
(E_AMinus) AMinus e1 e2 \\ n1-n2
(E_AMult) AMult e1 e2 \\ n1*n2
Exercise: 1 star, standard, optional (beval_rules)
(* FILL IN HERE *) (* Do not modify the following line: *) Definition manual_grade_for_beval_rules : option (nat*string) := None.
☐
(* ================================================================= *)
It is straightforward to prove that the relational and functional
definitions of evaluation agree:
forall (a : aexp) (n : nat), a \\ n <-> aeval a = nforall (a : aexp) (n : nat), a \\ n <-> aeval a = na:aexpn:nata \\ n -> aeval a = na:aexpn:nataeval a = n -> a \\ na:aexpn:nata \\ n -> aeval a = na:aexpn:natH:a \\ naeval a = nn:natn = ne1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2aeval e1 + aeval e2 = n1 + n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2aeval e1 - aeval e2 = n1 - n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2aeval e1 * aeval e2 = n1 * n2reflexivity.n:natn = ne1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2aeval e1 + aeval e2 = n1 + n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2n1 + aeval e2 = n1 + n2reflexivity.e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2n1 + n2 = n1 + n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2aeval e1 - aeval e2 = n1 - n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2n1 - aeval e2 = n1 - n2reflexivity.e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2n1 - n2 = n1 - n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2aeval e1 * aeval e2 = n1 * n2e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2n1 * aeval e2 = n1 * n2reflexivity.e1, e2:aexpn1, n2:natH:e1 \\ n1H0:e2 \\ n2IHaevalR1:aeval e1 = n1IHaevalR2:aeval e2 = n2n1 * n2 = n1 * n2a:aexpn:nataeval a = n -> a \\ na:aexpforall n : nat, aeval a = n -> a \\ nn0:natANum n0 \\ n0a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ nAPlus a1 a2 \\ aeval a1 + aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ nAMinus a1 a2 \\ aeval a1 - aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ nAMult a1 a2 \\ aeval a1 * aeval a2apply E_ANum.n0:natANum n0 \\ n0a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ nAPlus a1 a2 \\ aeval a1 + aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na1 \\ aeval a1a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ naeval a1 = aeval a1a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2reflexivity.a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ naeval a2 = aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ nAMinus a1 a2 \\ aeval a1 - aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na1 \\ aeval a1a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ naeval a1 = aeval a1a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2reflexivity.a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ naeval a2 = aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ nAMult a1 a2 \\ aeval a1 * aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na1 \\ aeval a1a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ naeval a1 = aeval a1a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ na2 \\ aeval a2reflexivity. Qed.a1, a2:aexpIHa1:forall n : nat, aeval a1 = n -> a1 \\ nIHa2:forall n : nat, aeval a2 = n -> a2 \\ naeval a2 = aeval a2
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(* WORKED IN CLASS *)forall (a : aexp) (n : nat), a \\ n <-> aeval a = na:aexpn:nata \\ n -> aeval a = na:aexpn:nataeval a = n -> a \\ nintros H; induction H; subst; reflexivity.a:aexpn:nata \\ n -> aeval a = na:aexpn:nataeval a = n -> a \\ ninduction a; simpl; intros; subst; constructor; try apply IHa1; try apply IHa2; reflexivity. Qed.a:aexpforall n : nat, aeval a = n -> a \\ n
Exercise: 3 stars, standard (bevalR)
Inductive bevalR: bexp -> bool -> Prop := (* FILL IN HERE *) .forall (b : bexp) (bv : bool), bevalR b bv <-> beval b = bv(* FILL IN HERE *) Admitted.forall (b : bexp) (bv : bool), bevalR b bv <-> beval b = bv
☐
End AExp. (* ================================================================= *)
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.
- 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.
(* ################################################################# *)
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.
(* ================================================================= *)
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. (* ================================================================= *)
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). (* ================================================================= *)
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.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.Unset Printing Coercions. (* ================================================================= *)
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) = 13reflexivity. Qed.aeval (X !-> 5) (3 + X * 2) = 13beval (X !-> 5) (true && (~ X <= 4)) = truereflexivity. Qed. (* ################################################################# *)beval (X !-> 5) (true && (~ X <= 4)) = true
Now we are ready define the syntax and behavior of Imp
commands (sometimes called statements).
(* ================================================================= *)
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. (* ================================================================= *)
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 (undo with Set Printing Notations)
- Set Printing Coercions (undo with Unset Printing Coercions)
- Set Printing All (undo with Unset Printing All)
Unset Printing Notations.Set Printing Notations. Set Printing Coercions.Unset Printing Coercions. (* ================================================================= *)
(* ----------------------------------------------------------------- *)
When faced with unknown notation, use Locate with a string
containing one of its symbols to see its possible
interpretations.
(* ----------------------------------------------------------------- *)
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.
(* ================================================================= *)
Assignment:
Definition plus2 : com := 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)%imp. Definition subtract_3_from_5_slowly : com := X ::= 3 ;; Z ::= 5 ;; subtract_slowly. (* ----------------------------------------------------------------- *)
Definition loop : com := WHILE true DO SKIP END. (* ################################################################# *)
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...
(* ================================================================= *)
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).
(* ================================================================= *)
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'".
(* ----------------------------------------------------------------- *)
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''
(E_Skip) st = SKIP => st
(E_Ass) st = x := a1 => (x !-> n ; st)
(E_Seq) st = c1;;c2 => st''
(E_IfTrue) st = TEST b1 THEN c1 ELSE c2 FI => st'
(E_IfFalse) st = TEST b1 THEN c1 ELSE c2 FI => st'
(E_WhileFalse) 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)(* We must supply the intermediate state *)empty_st =[ X ::= 2;; TEST X <= 1 THEN Y ::= 3 ELSE Z ::= 4 FI ]=> (Z !-> 4; X !-> 2)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)reflexivity.aeval empty_st 2 = 2(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)reflexivity. Qed.aeval (X !-> 2) 4 = 4
empty_st =[ X ::= 0;; Y ::= 1;; Z ::= 2 ]=> (Z !-> 2; Y !-> 1; X !-> 0)(* FILL IN HERE *) Admitted.empty_st =[ X ::= 0;; Y ::= 1;; Z ::= 2 ]=> (Z !-> 2; Y !-> 1; X !-> 0)
☐
Exercise: 3 stars, standard, optional (pup_to_n)
Admitted.com(X !-> 2) =[ pup_to_n ]=> (X !-> 0; Y !-> 3; X !-> 1; Y !-> 2; Y !-> 0; X !-> 2)(* FILL IN HERE *) Admitted.(X !-> 2) =[ pup_to_n ]=> (X !-> 0; Y !-> 3; X !-> 1; Y !-> 2; Y !-> 0; X !-> 2)
☐
(* ================================================================= *)
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 = st2forall (c : com) (st st1 st2 : state), st =[ c ]=> st1 -> st =[ c ]=> st2 -> st1 = st2c:comst, st1, st2:stateE1:st =[ c ]=> st1E2:st =[ c ]=> st2st1 = st2c:comst, st1:stateE1:st =[ c ]=> st1forall st2 : state, st =[ c ]=> st2 -> st1 = st2st2:stateE2:st2 =[ SKIP ]=> st2st2 = st2st:statea1:aexpx:stringE2:st =[ x ::= a1 ]=> (x !-> aeval st a1; st)(x !-> aeval st a1; st) = (x !-> aeval st a1; st)c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2st'0:stateH1:st =[ c1 ]=> st'0H4:st'0 =[ c2 ]=> st2st'' = st2st, st':stateb:bexpc1, c2:comH:beval st b = trueE1:st =[ c1 ]=> st'IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = trueH6:st =[ c1 ]=> st2st' = st2st, st':stateb:bexpc1, c2:comH:beval st b = trueE1:st =[ c1 ]=> st'IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = falseH6:st =[ c2 ]=> st2st' = st2st, st':stateb:bexpc1, c2:comH:beval st b = falseE1:st =[ c2 ]=> st'IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = trueH6:st =[ c1 ]=> st2st' = st2st, st':stateb:bexpc1, c2:comH:beval st b = falseE1:st =[ c2 ]=> st'IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = falseH6:st =[ c2 ]=> st2st' = st2b:bexpc:comst2:stateE2:st2 =[ WHILE b DO c END ]=> st2H, H4:beval st2 b = falsest2 = st2b:bexpst:statec:comH:beval st b = falsest2:stateE2:st =[ WHILE b DO c END ]=> st2st':stateH2:beval st b = trueH3:st =[ c ]=> st'H6:st' =[ WHILE b DO c END ]=> st2st = st2st', st'':stateb:bexpc:comst2:stateE2:st2 =[ WHILE b DO c END ]=> st2E1_1:st2 =[ c ]=> st'H:beval st2 b = trueE1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st2 =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0H4:beval st2 b = falsest'' = st2st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2st'0:stateH2:beval st b = trueH3:st =[ c ]=> st'0H6:st'0 =[ WHILE b DO c END ]=> st2st'' = st2reflexivity.st2:stateE2:st2 =[ SKIP ]=> st2st2 = st2reflexivity.st:statea1:aexpx:stringE2:st =[ x ::= a1 ]=> (x !-> aeval st a1; st)(x !-> aeval st a1; st) = (x !-> aeval st a1; st)c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2st'0:stateH1:st =[ c1 ]=> st'0H4:st'0 =[ c2 ]=> st2st'' = st2c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2st'0:stateH1:st =[ c1 ]=> st'0H4:st'0 =[ c2 ]=> st2st' = st'0c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2st'0:stateH1:st =[ c1 ]=> st'0H4:st'0 =[ c2 ]=> st2EQ1:st' = st'0st'' = st2apply IHE1_1; assumption.c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2st'0:stateH1:st =[ c1 ]=> st'0H4:st'0 =[ c2 ]=> st2st' = st'0c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2st'0:stateH1:st =[ c1 ]=> st'0H4:st'0 =[ c2 ]=> st2EQ1:st' = st'0st'' = st2c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2H4:st' =[ c2 ]=> st2H1:st =[ c1 ]=> st'st'' = st2assumption.c1, c2:comst, st', st'':stateE1_1:st =[ c1 ]=> st'E1_2:st' =[ c2 ]=> st''IHE1_1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ c2 ]=> st0 -> st'' = st0st2:stateE2:st =[ c1;; c2 ]=> st2H4:st' =[ c2 ]=> st2H1:st =[ c1 ]=> st'st' =[ c2 ]=> st2st, st':stateb:bexpc1, c2:comH:beval st b = trueE1:st =[ c1 ]=> st'IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = trueH6:st =[ c1 ]=> st2st' = st2assumption.st, st':stateb:bexpc1, c2:comH:beval st b = trueE1:st =[ c1 ]=> st'IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = trueH6:st =[ c1 ]=> st2st =[ c1 ]=> st2st, st':stateb:bexpc1, c2:comH:beval st b = trueE1:st =[ c1 ]=> st'IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = falseH6:st =[ c2 ]=> st2st' = st2discriminate H5.st, st':stateb:bexpc1, c2:comH:beval st b = trueE1:st =[ c1 ]=> st'IHE1:forall st0 : state, st =[ c1 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:true = falseH6:st =[ c2 ]=> st2st' = st2st, st':stateb:bexpc1, c2:comH:beval st b = falseE1:st =[ c2 ]=> st'IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = trueH6:st =[ c1 ]=> st2st' = st2discriminate H5.st, st':stateb:bexpc1, c2:comH:beval st b = falseE1:st =[ c2 ]=> st'IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:false = trueH6:st =[ c1 ]=> st2st' = st2st, st':stateb:bexpc1, c2:comH:beval st b = falseE1:st =[ c2 ]=> st'IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = falseH6:st =[ c2 ]=> st2st' = st2assumption.st, st':stateb:bexpc1, c2:comH:beval st b = falseE1:st =[ c2 ]=> st'IHE1:forall st0 : state, st =[ c2 ]=> st0 -> st' = st0st2:stateE2:st =[ TEST b THEN c1 ELSE c2 FI ]=> st2H5:beval st b = falseH6:st =[ c2 ]=> st2st =[ c2 ]=> st2reflexivity.b:bexpc:comst2:stateE2:st2 =[ WHILE b DO c END ]=> st2H, H4:beval st2 b = falsest2 = st2b:bexpst:statec:comH:beval st b = falsest2:stateE2:st =[ WHILE b DO c END ]=> st2st':stateH2:beval st b = trueH3:st =[ c ]=> st'H6:st' =[ WHILE b DO c END ]=> st2st = st2discriminate H2.b:bexpst:statec:comH:beval st b = falsest2:stateE2:st =[ WHILE b DO c END ]=> st2st':stateH2:false = trueH3:st =[ c ]=> st'H6:st' =[ WHILE b DO c END ]=> st2st = st2st', st'':stateb:bexpc:comst2:stateE2:st2 =[ WHILE b DO c END ]=> st2E1_1:st2 =[ c ]=> st'H:beval st2 b = trueE1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st2 =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0H4:beval st2 b = falsest'' = st2discriminate H4.st', st'':stateb:bexpc:comst2:stateE2:st2 =[ WHILE b DO c END ]=> st2E1_1:st2 =[ c ]=> st'H:beval st2 b = trueE1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st2 =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0H4:true = falsest'' = st2st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2st'0:stateH2:beval st b = trueH3:st =[ c ]=> st'0H6:st'0 =[ WHILE b DO c END ]=> st2st'' = st2st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2st'0:stateH2:beval st b = trueH3:st =[ c ]=> st'0H6:st'0 =[ WHILE b DO c END ]=> st2st' = st'0st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2st'0:stateH2:beval st b = trueH3:st =[ c ]=> st'0H6:st'0 =[ WHILE b DO c END ]=> st2EQ1:st' = st'0st'' = st2apply IHE1_1; assumption.st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2st'0:stateH2:beval st b = trueH3:st =[ c ]=> st'0H6:st'0 =[ WHILE b DO c END ]=> st2st' = st'0st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2st'0:stateH2:beval st b = trueH3:st =[ c ]=> st'0H6:st'0 =[ WHILE b DO c END ]=> st2EQ1:st' = st'0st'' = st2st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2H2:beval st b = trueH6:st' =[ WHILE b DO c END ]=> st2H3:st =[ c ]=> st'st'' = st2assumption. Qed. (* ################################################################# *)st, st', st'':stateb:bexpc:comH:beval st b = trueE1_1:st =[ c ]=> st'E1_2:st' =[ WHILE b DO c END ]=> st''IHE1_1:forall st0 : state, st =[ c ]=> st0 -> st' = st0IHE1_2:forall st0 : state, st' =[ WHILE b DO c END ]=> st0 -> st'' = st0st2:stateE2:st =[ WHILE b DO c END ]=> st2H2:beval st b = trueH6:st' =[ WHILE b DO c END ]=> st2H3:st =[ c ]=> st'st' =[ WHILE b DO c END ]=> st2
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 + 2forall (st : string -> nat) (n : nat) (st' : state), st X = n -> st =[ plus2 ]=> st' -> st' X = n + 2st:string -> natn:natst':stateHX:st X = nHeval: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 -> natn:natst':stateHX:st X = nHeval:st =[ plus2 ]=> st'st0:statea1:aexpn0:natx:stringH3:aeval st (X + 2) = n0H:x = XH1:a1 = (X + 2)%impH0:st0 = stH2:(X !-> n0; st) = st'(X !-> n0; st) X = n + 2st:string -> natHeval:st =[ plus2 ]=> (X !-> aeval st (X + 2); st)(X !-> aeval st (X + 2); st) X = st X + 2st:string -> nat(X !-> aeval st (X + 2); st) X = st X + 2apply t_update_eq. Qed.st:string -> nat(X !-> st X + 2; st) X = st X + 2
Exercise: 3 stars, standard, recommended (XtimesYinZ_spec)
(* FILL IN HERE *) (* Do not modify the following line: *) Definition manual_grade_for_XtimesYinZ_spec : option (nat*string) := None.
☐
forall st st' : state, ~ st =[ loop ]=> st'forall st st' : state, ~ st =[ loop ]=> st'st, st':statecontra:st =[ loop ]=> st'Falsest, st':statecontra:st =[ WHILE true DO SKIP END ]=> st'Falsest, st':stateloopdef:comHeqloopdef:loopdef = (WHILE true DO SKIP END)%impcontra: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.
☐
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(* FILL IN HERE *) Admitted.forall c : com, no_whiles c = true <-> no_whilesR c
☐
Exercise: 4 stars, standard (no_whiles_terminating)
(* FILL IN HERE *) (* Do not modify the following line: *) Definition manual_grade_for_no_whiles_terminating : option (nat*string) := None.
☐
(* ################################################################# *)
Exercise: 3 stars, standard (stack_compiler)
- SPush n: Push the number n on the stack.
- SLoad x: Load the identifier x from the store and push it on the stack
- SPlus: Pop the two top numbers from the stack, add them, and push the result onto the stack.
- SMinus: Similar, but subtract.
- SMult: Similar, but multiply.
Inductive sinstr : Type :=
| SPush (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.
Admitted.s_execute:state -> list nat -> list sinstr -> list natst:statestack:list natprog:list sinstrlist nat(* FILL IN HERE *) 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]
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.
Admitted.s_compile:aexp -> list sinstre:aexplist sinstr
After you've defined s_compile, prove the following to test
that it works.
(* FILL IN HERE *) Admitted.s_compile (X - 2 * Y) = [SLoad X; SPush 2; SLoad Y; SMult; SMinus]
☐
Exercise: 4 stars, advanced (stack_compiler_correct)
forall (st : state) (e : aexp), s_execute st [] (s_compile e) = [aeval st e](* FILL IN HERE *) Admitted.forall (st : state) (e : aexp), s_execute st [] (s_compile e) = [aeval st e]
☐
Exercise: 3 stars, standard, optional (short_circuit)
(* FILL IN HERE [] *) Module BreakImp.
Exercise: 4 stars, advanced (break_imp)
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:
- If the command is SKIP, then the state doesn't change and
execution of any enclosing loop can continue normally.
- If the command is BREAK, the state stays unchanged but we
signal a SBreak.
- If the command is an assignment, then we update the binding for
that variable in the state accordingly and signal that execution
can continue normally.
- If the command is of the form TEST b THEN c1 ELSE c2 FI, then
the state is updated as in the original semantics of Imp, except
that we also propagate the signal from the execution of
whichever branch was taken.
- If the command is a sequence c1 ;; c2, we first execute
c1. If this yields a SBreak, we skip the execution of c2
and propagate the SBreak signal to the surrounding context;
the resulting state is the same as the one obtained by
executing c1 alone. Otherwise, we execute c2 on the state
obtained after executing c1, and propagate the signal
generated there.
- Finally, for a loop of the form WHILE b DO c END, the semantics is almost the same as before. The only difference is that, when b evaluates to true, we execute c and check the signal that it raises. If that signal is SContinue, then the execution proceeds as in the original semantics. Otherwise, we stop the execution of the loop, and the resulting state is the same as the one resulting from the execution of the current iteration. In either case, since BREAK only terminates the innermost loop, WHILE signals SContinue.
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'(* FILL IN HERE *) Admitted.forall (c : com) (st st' : state) (s : result), st =[ BREAK;; c ]=> st' / s -> st = st'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) (s : result), st =[ WHILE b DO c END ]=> st' / s -> s = SContinueforall (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.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), st =[ WHILE b DO c END ]=> st' / SContinue -> beval st' b = true -> exists st'' : state, st'' =[ c ]=> st' / SBreak(* FILL IN HERE *) Admitted.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 (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.forall (c : com) (st st1 st2 : state) (s1 s2 : result), st =[ c ]=> st1 / s1 -> st =[ c ]=> st2 / s2 -> st1 = st2 /\ s1 = s2
☐
End BreakImp.
Exercise: 4 stars, standard, optional (add_for_loop)
(* FILL IN HERE
[] *)
(* Wed Jan 9 12:02:46 EST 2019 *)