changeset 11:e5501b9c965d

Nats StackMachine done
author Adam Chlipala <adamc@hcoop.net>
date Mon, 01 Sep 2008 12:52:30 -0400
parents d8363f2a3bef
children ea400f692b07
files book/src/StackMachine.v
diffstat 1 files changed, 290 insertions(+), 2 deletions(-) [+]
line wrap: on
line diff
--- a/book/src/StackMachine.v	Mon Sep 01 11:30:27 2008 -0400
+++ b/book/src/StackMachine.v	Mon Sep 01 12:52:30 2008 -0400
@@ -16,7 +16,9 @@
 
 (** I will start off by jumping right in to a fully-worked set of examples, building certified compilers from increasingly complicated source languages to stack machines.  We will meet a few useful tactics and see how they can be used in manual proofs, and we will also see how easily these proofs can be automated instead.  I assume that you have installed Coq and Proof General.
 
-As always, you can step through the source file %\texttt{%#<tt>#StackMachine.v#</tt>#%}% for this chapter interactively in Proof General.  Alternatively, to get a feel for the whole lifecycle of creating a Coq development, you can enter the pieces of source code in this chapter in a new %\texttt{%#<tt>#.v#</tt>#%}% file in an Emacs buffer.  If you do the latter, include a line [Require Import List Tactics] at the start of the file, to match some code hidden from the chapter source, and be sure to run the Coq binary %\texttt{%#<tt>#coqtop#</tt>#%}% with the command-line argument %\texttt{%#<tt>#-I SRC#</tt>#%}%, where %\texttt{%#<tt>#SRC#</tt>#%}% is the path to a directory containing the source for this book.  In either case, if you have installed Proof General properly, it should start automatically when you visit a %\texttt{%#<tt>#.v#</tt>#%}% buffer in Emacs. *)
+As always, you can step through the source file %\texttt{%#<tt>#StackMachine.v#</tt>#%}% for this chapter interactively in Proof General.  Alternatively, to get a feel for the whole lifecycle of creating a Coq development, you can enter the pieces of source code in this chapter in a new %\texttt{%#<tt>#.v#</tt>#%}% file in an Emacs buffer.  If you do the latter, include a line [Require Import List Tactics] at the start of the file, to match some code hidden from the chapter source, and be sure to run the Coq binary %\texttt{%#<tt>#coqtop#</tt>#%}% with the command-line argument %\texttt{%#<tt>#-I SRC#</tt>#%}%, where %\texttt{%#<tt>#SRC#</tt>#%}% is the path to a directory containing the source for this book.  In either case, if you have installed Proof General properly, it should start automatically when you visit a %\texttt{%#<tt>#.v#</tt>#%}% buffer in Emacs.
+
+With Proof General, the portion of a buffer that Coq has processed is highlighted in some way, like being given a blue background.  You step through Coq source files by positioning the point at the position you want Coq to run to and pressing C-C C-RET.  This can be used both for normal step-by-step coding, by placing the point inside some command past the end of the highlighted region; and for undoing, by placing the point inside the highlighted region. *)
 
 
 (** * Arithmetic expressions over natural numbers *)
@@ -143,27 +145,269 @@
 
 (** ** Translation correctness *)
 
+(** We are ready to prove that our compiler is implemented correctly.  We can use a new vernacular command [Theorem] to start a correctness proof, in terms of the semantics we defined earlier: *)
+
+Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
+(* begin hide *)
+Abort.
+(* end hide *)
+
+(** Though a pencil-and-paper proof might clock out at this point, writing "by a routine induction on [e]," it turns out not to make sense to attack this proof directly.  We need to use the standard trick of %\textit{%#<i>#strengthening the induction hypothesis#</i>#%}%.  We do that by proving an auxiliary lemma:
+*)
+
 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
   progDenote p (expDenote e :: s).
+
+(** After the period in the [Lemma] command, we are in %\textit{%#<i>#the interactive proof-editing mode#</i>#%}%.  We find ourselves staring at this ominous screen of text:
+
+[[
+1 subgoal
+  
+ ============================
+ forall (e : exp) (s : stack) (p : list instr),
+   progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
+]]
+
+Coq seems to be restating the lemma for us.  What we are seeing is a limited case of a more general protocol for describing where we are in a proof.  We are told that we have a single subgoal.  In general, during a proof, we can have many pending subgoals, each of which is a logical proposition to prove.  Subgoals can be proved in any order, but it usually works best to prove them in the order that Coq chooses.
+
+Next in the output, we see our single subgoal described in full detail.  There is a double-dashed line, above which would be our free variables and hypotheses, if we had any.  Below the line is the conclusion, which, in general, is to be proved from the hypotheses.
+
+We manipulate the proof state by running commands called %\textit{%#<i>#tactics#</i>#%}%.  Let us start out by running one of the most important tactics:
+*)
+
   induction e.
 
+(** We declare that this proof will proceed by induction on the structure of the expression [e].  This swaps out our initial subgoal for two new subgoals, one for each case of the inductive proof:
+
+[[
+2 subgoals
+  
+ n : nat
+ ============================
+ forall (s : stack) (p : list instr),
+   progDenote (compile (Const n) ++ p) s =
+   progDenote p (expDenote (Const n) :: s)
+]]
+[[
+ subgoal 2 is:
+  forall (s : stack) (p : list instr),
+    progDenote (compile (Binop b e1 e2) ++ p) s =
+    progDenote p (expDenote (Binop b e1 e2) :: s)
+]]
+
+The first and current subgoal is displayed with the double-dashed line below free variables and hypotheses, while later subgoals are only summarized with their conclusions.  We see an example of a free variable in the first subgoal; [n] is a free variable of type [nat].  The conclusion is the original theorem statement where [e] has been replaced by [Const n].  In a similar manner, the second case has [e] replaced by a generalized invocation of the [Binop] expression constructor.  We can see that proving both cases corresponds to a standard proof by structural induction.
+
+We begin the first case with another very common tactic.
+*)
+
   intros.
+
+(** The current subgoal changes to:
+[[
+
+ n : nat
+ s : stack
+ p : list instr
+ ============================
+ progDenote (compile (Const n) ++ p) s =
+ progDenote p (expDenote (Const n) :: s)
+]]
+
+We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
+
+To progress further, we need to use the definitions of some of the functions appearing in the goal.  The [unfold] tactic replaces an identifier with its definition.
+*)
+
   unfold compile.
+
+(** [[
+
+ n : nat
+ s : stack
+ p : list instr
+ ============================
+ progDenote ((IConst n :: nil) ++ p) s =
+ progDenote p (expDenote (Const n) :: s)
+]] *)
+
   unfold expDenote.
+
+(** [[
+
+ n : nat
+ s : stack
+ p : list instr
+ ============================
+ progDenote ((IConst n :: nil) ++ p) s = progDenote p (n :: s)
+]]
+
+We only need to unfold the first occurrence of [progDenote] to prove the goal: *)
+
+  unfold progDenote at 1.
+
+(** [[
+
+ n : nat
+ s : stack
+ p : list instr
+ ============================
+  (fix progDenote (p0 : prog) (s0 : stack) {struct p0} : 
+    option stack :=
+      match p0 with
+      | nil => Some s0
+      | i :: p' =>
+          match instrDenote i s0 with
+          | Some s' => progDenote p' s'
+          | None => None (A:=stack)
+          end
+      end) ((IConst n :: nil) ++ p) s =
+   progDenote p (n :: s)
+]]
+
+This last [unfold] has left us with an anonymous fixpoint version of [progDenote], which will generally happen when unfolding recursive definitions.  Fortunately, in this case, we can eliminate such complications right away, since the structure of the argument [(IConst n :: nil) ++ p] is known, allowing us to simplify the internal pattern match with the [simpl] tactic:
+*)
+
   simpl.
+
+(** [[
+
+ n : nat
+ s : stack
+ p : list instr
+ ============================
+ (fix progDenote (p0 : prog) (s0 : stack) {struct p0} : 
+  option stack :=
+    match p0 with
+    | nil => Some s0
+    | i :: p' =>
+        match instrDenote i s0 with
+        | Some s' => progDenote p' s'
+        | None => None (A:=stack)
+        end
+    end) p (n :: s) = progDenote p (n :: s)
+]]
+
+Now we can unexpand the definition of [progDenote]:
+*)
+
+  fold progDenote.
+
+(** [[
+
+ n : nat
+ s : stack
+ p : list instr
+ ============================
+ progDenote p (n :: s) = progDenote p (n :: s)
+]]
+
+It looks like we are at the end of this case, since we have a trivial equality.  Indeed, a single tactic finishes us off:
+*)
+
   reflexivity.
 
+(** On to the second inductive case:
+
+[[
+
+  b : binop
+  e1 : exp
+  IHe1 : forall (s : stack) (p : list instr),
+         progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
+  e2 : exp
+  IHe2 : forall (s : stack) (p : list instr),
+         progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
+  ============================
+   forall (s : stack) (p : list instr),
+   progDenote (compile (Binop b e1 e2) ++ p) s =
+   progDenote p (expDenote (Binop b e1 e2) :: s)
+]]
+
+We see our first example of hypotheses above the double-dashed line.  They are the inductive hypotheses [IHe1] and [IHe2] corresponding to the subterms [e1] and [e2], respectively.
+
+We start out the same way as before, introducing new free variables and unfolding and folding the appropriate definitions.  The seemingly frivolous [unfold]/[fold] pairs are actually accomplishing useful work, because [unfold] will sometimes perform easy simplifications. *)
+
   intros.
   unfold compile.
   fold compile.
   unfold expDenote.
   fold expDenote.
+
+(** Now we arrive at a point where the tactics we have seen so far are insufficient:
+
+[[
+
+  b : binop
+  e1 : exp
+  IHe1 : forall (s : stack) (p : list instr),
+         progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
+  e2 : exp
+  IHe2 : forall (s : stack) (p : list instr),
+         progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
+  s : stack
+  p : list instr
+  ============================
+   progDenote ((compile e2 ++ compile e1 ++ IBinop b :: nil) ++ p) s =
+   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
+]]
+
+What we need is the associative law of list concatenation, available as a theorem [app_ass] in the standard library. *)
+
+Check app_ass.
+
+(** [[
+
+app_ass
+     : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
+]]
+
+We use it to perform a rewrite: *)
+
   rewrite app_ass.
+
+(** changing the conclusion to: [[
+
+   progDenote (compile e2 ++ (compile e1 ++ IBinop b :: nil) ++ p) s =
+   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
+]]
+
+Now we can notice that the lefthand side of the equality matches the lefthand side of the second inductive hypothesis, so we can rewrite with that hypothesis, too: *)
+
   rewrite IHe2.
+
+(** [[
+
+   progDenote ((compile e1 ++ IBinop b :: nil) ++ p) (expDenote e2 :: s) =
+   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
+]]
+
+The same process lets us apply the remaining hypothesis. *)
+
   rewrite app_ass.
   rewrite IHe1.
+
+(** [[
+
+   progDenote ((IBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
+   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
+]]
+
+Now we can apply a similar sequence of tactics to that that ended the proof of the first case.
+*)
+
+  unfold progDenote at 1.
   simpl.
+  fold progDenote.
   reflexivity.
+
+(** And the proof is completed, as indicated by the message:
+
+[[
+Proof completed.
+
+And there lies our first proof.  Already, even for simple theorems like this, the final proof script is unstructured and not very enlightening to readers.  If we extend this approach to more serious theorems, we arrive at the unreadable proof scripts that are the favorite complaints of opponents of tactic-based proving.  Fortunately, Coq has rich support for scripted automation, and we can take advantage of such a scripted tactic (defined elsewhere) to make short work of this lemma.  We abort the old proof attempt and start again.
+*)
+
 Abort.
 
 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
@@ -171,9 +415,53 @@
   induction e; crush.
 Qed.
 
+(** We need only to state the basic inductive proof scheme and call a tactic that automates the tedious reasoning in between.  In contrast to the period tactic terminator from our last proof, the semicolon tactic separator supports structured, compositional proofs.  The tactic [t1; t2] has the effect of running [t1] and then running [t2] on each remaining subgoal.  The semicolon is one of the most fundamental building blocks of effective proof automation.  The period terminator is very useful for exploratory proving, where you need to see intermediate proof states, but final proofs of any serious complexity should have just one period, terminating a single compound tactic that probably uses semicolons.
+
+The proof of our main theorem is now easy.  We prove it with four period-terminated tactics, though separating them with semicolons would work as well; the version here is easier to step through. *)
+
 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
-  intro.
+  intros.
+
+(** [[
+
+  e : exp
+  ============================
+   progDenote (compile e) nil = Some (expDenote e :: nil)
+]]
+
+At this point, we want to massage the lefthand side to match the statement of [compileCorrect'].  A theorem from the standard library is useful: *)
+
+Check app_nil_end.
+
+(** [[
+
+app_nil_end
+     : forall (A : Type) (l : list A), l = l ++ nil
+]] *)
+
   rewrite (app_nil_end (compile e)).
+
+(** This time, we explicitly specify the value of the variable [l] from the theorem statement, since multiple expressions of list type appear in the conclusion.  [rewrite] might choose the wrong place to rewrite if we did not specify which we want.
+
+[[
+
+  e : exp
+  ============================
+   progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
+]]
+
+Now we can apply the lemma. *)
+
   rewrite compileCorrect'.
+
+(** [[
+
+  e : exp
+  ============================
+   progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
+]]
+
+We are almost done.  The lefthand and righthand sides can be seen to match by simple symbolic evaluation.  That means we are in luck, because Coq identifies any pair of terms as equal whenever they normalize to the same result by symbolic evaluation.  By the definition of [progDenote], that is the case here, but we do not need to worry about such details.  A simple invocation of [reflexivity] does the normalization and checks that the two results are syntactically equal. *)
+
   reflexivity.
 Qed.