StackMachine.v 47.3 KB
Newer Older
1
(* Copyright (c) 2008-2012, Adam Chlipala
2 3 4 5 6 7 8 9 10
 * 
 * This work is licensed under a
 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
 * Unported License.
 * The license text is available at:
 *   http://creativecommons.org/licenses/by-nc-nd/3.0/
 *)


11 12
(** %\chapter{Some Quick Examples}% *)

13
(** 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.  This chapter is not meant to give full explanations of the features that are employed.  Rather, it is meant more as an advertisement of what is possible.  Later chapters will introduce all of the concepts in bottom-up fashion.  In other words, it is expected that most readers will not understand what exactly is going on here, but I hope this demo will whet your appetite for the remaining chapters!
14

15
As always, you can step through the source file <<StackMachine.v>> 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 <<.v>> file in an Emacs buffer.  If you do the latter, include these two lines at the start of the file. *)
16

17 18
Require Import Bool Arith List CpdtTactics.
Set Implicit Arguments.
19

20
(** In general, similar commands will be hidden in the book rendering of each chapter's source code, so you will need to insert them in from-scratch replayings of the code that is presented.  To be more specific, every chapter begins with the above two lines, with the import list tweaked as appropriate, considering which definitions the chapter uses.  The second command above affects the default behavior of definitions regarding type inference. *)
21 22


Adam Chlipala's avatar
Adam Chlipala committed
23
(** * Arithmetic Expressions Over Natural Numbers *)
24

Adam Chlipala's avatar
Adam Chlipala committed
25
(** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *)
26

Adam Chlipala's avatar
Adam Chlipala committed
27
(** ** Source Language *)
28

29
(** We begin with the syntax of the source language.%\index{Vernacular commands!Inductive}% *)
30

Adam Chlipala's avatar
Adam Chlipala committed
31
Inductive binop : Set := Plus | Times.
32

33
(** Our first line of Coq code should be unsurprising to ML and Haskell programmers.  We define an %\index{algebraic datatypes}%algebraic datatype [binop] to stand for the binary operators of our source language.  There are just two wrinkles compared to ML and Haskell.  First, we use the keyword [Inductive], in place of <<data>>, <<datatype>>, or <<type>>.  This is not just a trivial surface syntax difference; inductive types in Coq are much more expressive than garden variety algebraic datatypes, essentially enabling us to encode all of mathematics, though we begin humbly in this chapter.  Second, there is the %\index{Gallina terms!Set}%[: Set] fragment, which declares that we are defining a datatype that should be thought of as a constituent of programs.  Later, we will see other options for defining datatypes in the universe of proofs or in an infinite hierarchy of universes, encompassing both programs and proofs, that is useful in higher-order constructions. *)
34

Adam Chlipala's avatar
Adam Chlipala committed
35 36 37
Inductive exp : Set :=
| Const : nat -> exp
| Binop : binop -> exp -> exp -> exp.
38

39 40
(** Now we define the type of arithmetic expressions.  We write that a constant may be built from one argument, a natural number; and a binary operation may be built from a choice of operator and two operand expressions.

41
A note for readers following along in the PDF version: %\index{coqdoc}%coqdoc supports pretty-printing of tokens in %\LaTeX{}%#LaTeX# or HTML.  Where you see a right arrow character, the source contains the ASCII text <<->>>.  Other examples of this substitution appearing in this chapter are a double right arrow for <<=>>>, the inverted %`%#'#A' symbol for <<forall>>, and the Cartesian product %`%#'#X' for <<*>>.  When in doubt about the ASCII version of a symbol, you can consult the chapter source code.
42 43 44

%\medskip%

45
Now we are ready to say what these programs mean.  We will do this by writing an %\index{interpreters}%interpreter that can be thought of as a trivial operational or denotational semantics.  (If you are not familiar with these semantic techniques, no need to worry; we will stick to "common sense" constructions.)%\index{Vernacular commands!Definition}% *)
46

Adam Chlipala's avatar
Adam Chlipala committed
47 48 49 50 51
Definition binopDenote (b : binop) : nat -> nat -> nat :=
  match b with
    | Plus => plus
    | Times => mult
  end.
52

53
(** The meaning of a binary operator is a binary function over naturals, defined with pattern-matching notation analogous to the <<case>> and <<match>> of ML and Haskell, and referring to the functions [plus] and [mult] from the Coq standard library.  The keyword [Definition] is Coq's all-purpose notation for binding a term of the programming language to a name, with some associated syntactic sugar, like the notation we see here for defining a function.  That sugar could be expanded to yield this definition:
54 55 56 57 58 59
[[
Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
  match b with
    | Plus => plus
    | Times => mult
  end.
60 61
]]

62 63 64 65 66 67 68
In this example, we could also omit all of the type annotations, arriving at:
[[
Definition binopDenote := fun b =>
  match b with
    | Plus => plus
    | Times => mult
  end.
69 70
]]

71
Languages like Haskell and ML have a convenient%\index{principal types}\index{type inference}% _principal types_ property, which gives us strong guarantees about how effective type inference will be.  Unfortunately, Coq's type system is so expressive that any kind of "complete" type inference is impossible, and the task even seems to be hard heuristically in practice.  Nonetheless, Coq includes some very helpful heuristics, many of them copying the workings of Haskell and ML type-checkers for programs that fall in simple fragments of Coq's language.
72

73
This is as good a time as any to mention the preponderance of different languages associated with Coq.  The theoretical foundation of Coq is a formal system called the%\index{Calculus of Inductive Constructions}\index{CIC|see{Calculus of Inductive Constructions}}% _Calculus of Inductive Constructions_ (CIC)%~\cite{CIC}%, which is an extension of the older%\index{Calculus of Constructions}\index{CoC|see{Calculus of Constructions}}% _Calculus of Constructions_ (CoC)%~\cite{CoC}%.  CIC is quite a spartan foundation, which is helpful for proving metatheory but not so helpful for real development.  Still, it is nice to know that it has been proved that CIC enjoys properties like%\index{strong normalization}% _strong normalization_ %\cite{CIC}%, meaning that every program (and, more importantly, every proof term) terminates; and%\index{relative consistency}% _relative consistency_ %\cite{SetsInTypes}% with systems like versions of %\index{Zermelo-Fraenkel set theory}%Zermelo-Fraenkel set theory, which roughly means that you can believe that Coq proofs mean that the corresponding propositions are "really true," if you believe in set theory.
74

75
Coq is actually based on an extension of CIC called%\index{Gallina}% _Gallina_.  The text after the [:=] and before the period in the last code example is a term of Gallina.  Gallina adds many useful features that are not compiled internally to more primitive CIC features.  The important metatheorems about CIC have not been extended to the full breadth of these features, but most Coq users do not seem to lose much sleep over this omission.
76

77
Next, there is%\index{Ltac}% _Ltac_, Coq's domain-specific language for writing proofs and decision procedures. We will see some basic examples of Ltac later in this chapter, and much of this book is devoted to more involved Ltac examples.
78

79
Finally, commands like [Inductive] and [Definition] are part of%\index{Vernacular commands}% _the Vernacular_, which includes all sorts of useful queries and requests to the Coq system.  Every Coq source file is a series of vernacular commands, where many command forms take arguments that are Gallina or Ltac programs.  (Actually, Coq source files are more like _trees_ of vernacular commands, thanks to various nested scoping constructs.)
80 81 82

%\medskip%

83
We can give a simple definition of the meaning of an expression:%\index{Vernacular commands!Fixpoint}% *)
84

Adam Chlipala's avatar
Adam Chlipala committed
85 86 87 88 89
Fixpoint expDenote (e : exp) : nat :=
  match e with
    | Const n => n
    | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
  end.
90

91 92
(** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint].  The rest should be old hat for functional programmers. *)

93
(** It is convenient to be able to test definitions before starting to prove things about them.  We can verify that our semantics is sensible by evaluating some sample uses, using the command %\index{Vernacular commands!Eval}%[Eval].  This command takes an argument expressing a%\index{reduction strategy}% _reduction strategy_, or an "order of evaluation."  Unlike with ML, which hardcodes an _eager_ reduction strategy, or Haskell, which hardcodes a _lazy_ strategy, in Coq we are free to choose between these and many other orders of evaluation, because all Coq programs terminate.  In fact, Coq silently checked %\index{termination checking}%termination of our [Fixpoint] definition above, using a simple heuristic based on monotonically decreasing size of arguments across recursive calls.  Specifically, recursive calls must be made on arguments that were pulled out of the original recursive argument with [match] expressions.  (In Chapter 7, we will see some ways of getting around this restriction, though simply removing the restriction would leave Coq useless as a theorem proving tool, for reasons we will start to learn about in the next chapter.)
94 95

To return to our test evaluations, we run the [Eval] command using the [simpl] evaluation strategy, whose definition is best postponed until we have learned more about Coq's foundations, but which usually gets the job done. *)
Adam Chlipala's avatar
Adam Chlipala committed
96 97

Eval simpl in expDenote (Const 42).
98 99
(** [= 42 : nat] *)

Adam Chlipala's avatar
Adam Chlipala committed
100
Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)).
101 102
(** [= 4 : nat] *)

Adam Chlipala's avatar
Adam Chlipala committed
103
Eval simpl in expDenote (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
104
(** [= 28 : nat] *)
105

106 107 108
(** %\smallskip{}%Nothing too surprising goes on here, so we are ready to move on to the target language of our compiler. *)


Adam Chlipala's avatar
Adam Chlipala committed
109
(** ** Target Language *)
110

111 112
(** We will compile our source programs onto a simple stack machine, whose syntax is: *)

Adam Chlipala's avatar
Adam Chlipala committed
113
Inductive instr : Set :=
114 115
| iConst : nat -> instr
| iBinop : binop -> instr.
Adam Chlipala's avatar
Adam Chlipala committed
116 117 118 119

Definition prog := list instr.
Definition stack := list nat.

120 121
(** An instruction either pushes a constant onto the stack or pops two arguments, applies a binary operator to them, and pushes the result onto the stack.  A program is a list of instructions, and a stack is a list of natural numbers.

122
We can give instructions meanings as functions from stacks to optional stacks, where running an instruction results in [None] in case of a stack underflow and results in [Some s'] when the result of execution is the new stack [s'].  %\index{Gallina operators!::}%The infix operator [::] is "list cons" from the Coq standard library.%\index{Gallina terms!option}% *)
123

Adam Chlipala's avatar
Adam Chlipala committed
124 125
Definition instrDenote (i : instr) (s : stack) : option stack :=
  match i with
126 127
    | iConst n => Some (n :: s)
    | iBinop b =>
Adam Chlipala's avatar
Adam Chlipala committed
128 129 130 131 132 133
      match s with
        | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
        | _ => None
      end
  end.

134
(** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *)
135 136 137 138 139 140 141 142 143 144

Fixpoint progDenote (p : prog) (s : stack) : option stack :=
  match p with
    | nil => Some s
    | i :: p' =>
      match instrDenote i s with
        | None => None
        | Some s' => progDenote p' s'
      end
  end.
145

146 147
(** With the two programming languages defined, we can turn to the compiler definition. *)

148

149
(** ** Translation *)
150

151
(** Our compiler itself is now unsurprising.  The list concatenation operator %\index{Gallina operators!++}%[++] comes from the Coq standard library. *)
152

Adam Chlipala's avatar
Adam Chlipala committed
153 154
Fixpoint compile (e : exp) : prog :=
  match e with
155 156
    | Const n => iConst n :: nil
    | Binop b e1 e2 => compile e2 ++ compile e1 ++ iBinop b :: nil
Adam Chlipala's avatar
Adam Chlipala committed
157
  end.
158 159


Adam Chlipala's avatar
Adam Chlipala committed
160 161 162
(** Before we set about proving that this compiler is correct, we can try a few test runs, using our sample programs from earlier. *)

Eval simpl in compile (Const 42).
163
(** [= iConst 42 :: nil : prog] *)
164

Adam Chlipala's avatar
Adam Chlipala committed
165
Eval simpl in compile (Binop Plus (Const 2) (Const 2)).
166
(** [= iConst 2 :: iConst 2 :: iBinop Plus :: nil : prog] *)
167

Adam Chlipala's avatar
Adam Chlipala committed
168
Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
169
(** [= iConst 7 :: iConst 2 :: iConst 2 :: iBinop Plus :: iBinop Times :: nil : prog] *)
Adam Chlipala's avatar
Adam Chlipala committed
170

171
(** %\smallskip{}%We can also run our compiled programs and check that they give the right results. *)
Adam Chlipala's avatar
Adam Chlipala committed
172 173

Eval simpl in progDenote (compile (Const 42)) nil.
174 175
(** [= Some (42 :: nil) : option stack] *)

Adam Chlipala's avatar
Adam Chlipala committed
176
Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil.
177 178
(** [= Some (4 :: nil) : option stack] *)

179 180
Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2))
  (Const 7))) nil.
181
(** [= Some (28 :: nil) : option stack] *)
Adam Chlipala's avatar
Adam Chlipala committed
182

183
(** %\smallskip{}%So far so good, but how can we be sure the compiler operates correctly for _all_ input programs? *)
Adam Chlipala's avatar
Adam Chlipala committed
184

Adam Chlipala's avatar
Adam Chlipala committed
185
(** ** Translation Correctness *)
186

187
(** 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:%\index{Vernacular commands!Theorem}% *)
Adam Chlipala's avatar
Adam Chlipala committed
188

Adam Chlipala's avatar
Adam Chlipala committed
189
Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
Adam Chlipala's avatar
Adam Chlipala committed
190 191 192
(* begin hide *)
Abort.
(* end hide *)
Adam Chlipala's avatar
Adam Chlipala committed
193
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
194

195
(** 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%\index{strengthening the induction hypothesis}% _strengthening the induction hypothesis_.  We do that by proving an auxiliary lemma, using the command [Lemma] that is a synonym for [Theorem], conventionally used for less important theorems that appear in the proofs of primary theorems.%\index{Vernacular commands!Lemma}% *)
Adam Chlipala's avatar
Adam Chlipala committed
196

197 198
Lemma compile_correct' : forall e p s,
  progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
Adam Chlipala's avatar
Adam Chlipala committed
199

200
(** After the period in the [Lemma] command, we are in%\index{interactive proof-editing mode}% _the interactive proof-editing mode_.  We find ourselves staring at this ominous screen of text:
Adam Chlipala's avatar
Adam Chlipala committed
201 202 203 204 205

[[
1 subgoal
  
 ============================
206 207
  forall (e : exp) (p : list instr) (s : stack),
   progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)  
208
 
Adam Chlipala's avatar
Adam Chlipala committed
209 210
]]

211
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 %\index{subgoals}%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.
Adam Chlipala's avatar
Adam Chlipala committed
212

213
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 %\index{hypotheses}%hypotheses, if we had any.  Below the line is the %\index{conclusion}%conclusion, which, in general, is to be proved from the hypotheses.
Adam Chlipala's avatar
Adam Chlipala committed
214

215
We manipulate the proof state by running commands called%\index{tactics}% _tactics_.  Let us start out by running one of the most important tactics:%\index{tactics!induction}%
Adam Chlipala's avatar
Adam Chlipala committed
216 217
*)

Adam Chlipala's avatar
Adam Chlipala committed
218 219
  induction e.

Adam Chlipala's avatar
Adam Chlipala committed
220 221 222
(** 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:

[[
223 224
2 subgoals

Adam Chlipala's avatar
Adam Chlipala committed
225 226 227 228 229
 n : nat
 ============================
 forall (s : stack) (p : list instr),
   progDenote (compile (Const n) ++ p) s =
   progDenote p (expDenote (Const n) :: s)
230 231 232

subgoal 2 is

Adam Chlipala's avatar
Adam Chlipala committed
233 234 235
  forall (s : stack) (p : list instr),
    progDenote (compile (Binop b e1 e2) ++ p) s =
    progDenote p (expDenote (Binop b e1 e2) :: s)
236
 
Adam Chlipala's avatar
Adam Chlipala committed
237 238
]]

239
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 %\index{free variable}%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 %\index{structural induction}%structural induction.
Adam Chlipala's avatar
Adam Chlipala committed
240

241
We begin the first case with another very common tactic.%\index{tactics!intros}%
Adam Chlipala's avatar
Adam Chlipala committed
242 243
*)

Adam Chlipala's avatar
Adam Chlipala committed
244
  intros.
Adam Chlipala's avatar
Adam Chlipala committed
245 246 247 248 249 250 251 252 253 254

(** The current subgoal changes to:
[[

 n : nat
 s : stack
 p : list instr
 ============================
 progDenote (compile (Const n) ++ p) s =
 progDenote p (expDenote (Const n) :: s)
255
 
Adam Chlipala's avatar
Adam Chlipala committed
256 257 258 259
]]

We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.

260
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.%\index{tactics!unfold}%
Adam Chlipala's avatar
Adam Chlipala committed
261 262
*)

Adam Chlipala's avatar
Adam Chlipala committed
263
  unfold compile.
Adam Chlipala's avatar
Adam Chlipala committed
264 265 266 267 268
(** [[
 n : nat
 s : stack
 p : list instr
 ============================
269
 progDenote ((iConst n :: nil) ++ p) s =
Adam Chlipala's avatar
Adam Chlipala committed
270
 progDenote p (expDenote (Const n) :: s)
271
 
272 273
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
274

Adam Chlipala's avatar
Adam Chlipala committed
275
  unfold expDenote.
Adam Chlipala's avatar
Adam Chlipala committed
276 277 278 279 280
(** [[
 n : nat
 s : stack
 p : list instr
 ============================
281
 progDenote ((iConst n :: nil) ++ p) s = progDenote p (n :: s)
282
 
Adam Chlipala's avatar
Adam Chlipala committed
283 284
]]

285
We only need to unfold the first occurrence of [progDenote] to prove the goal.  An [at] clause used with [unfold] specifies a particular occurrence of an identifier to unfold, where we count occurrences from left to right.%\index{tactics!unfold}% *)
Adam Chlipala's avatar
Adam Chlipala committed
286 287 288 289 290 291 292 293 294 295 296 297 298 299 300 301

  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
302
      end) ((iConst n :: nil) ++ p) s =
Adam Chlipala's avatar
Adam Chlipala committed
303
   progDenote p (n :: s)
304
 
Adam Chlipala's avatar
Adam Chlipala committed
305 306
]]

307
This last [unfold] has left us with an anonymous fixpoint version of [progDenote], which will generally happen when unfolding recursive definitions.  Note that Coq has automatically renamed the [fix] arguments [p] and [s] to [p0] and [s0], to avoid clashes with our local free variables.  There is also a subterm [None (A:=stack)], which has an annotation specifying that the type of the term ought to be [option stack].  This is phrased as an explicit instantiation of a named type parameter [A] from the definition of [option].
308 309

Fortunately, in this case, we can eliminate the complications of anonymous recursion 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, which applies the same reduction strategy that we used earlier with [Eval] (and whose details we still postpone).%\index{tactics!simpl}%
Adam Chlipala's avatar
Adam Chlipala committed
310 311
*)

Adam Chlipala's avatar
Adam Chlipala committed
312
  simpl.
Adam Chlipala's avatar
Adam Chlipala committed
313 314 315 316 317 318 319 320 321 322 323 324 325 326 327
(** [[
 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)
328
 
Adam Chlipala's avatar
Adam Chlipala committed
329 330
]]

331
Now we can unexpand the definition of [progDenote]:%\index{tactics!fold}%
Adam Chlipala's avatar
Adam Chlipala committed
332 333 334 335 336 337 338 339 340
*)

  fold progDenote.
(** [[
 n : nat
 s : stack
 p : list instr
 ============================
 progDenote p (n :: s) = progDenote p (n :: s)
341
 
Adam Chlipala's avatar
Adam Chlipala committed
342 343
]]

344
It looks like we are at the end of this case, since we have a trivial equality.  Indeed, a single tactic finishes us off:%\index{tactics!reflexivity}%
Adam Chlipala's avatar
Adam Chlipala committed
345 346
*)

Adam Chlipala's avatar
Adam Chlipala committed
347 348
  reflexivity.

Adam Chlipala's avatar
Adam Chlipala committed
349 350 351 352 353 354 355 356 357 358 359 360 361 362
(** 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)
363
 
Adam Chlipala's avatar
Adam Chlipala committed
364 365
]]

366
We see our first example of %\index{hypotheses}%hypotheses above the double-dashed line.  They are the inductive hypotheses [IHe1] and [IHe2] corresponding to the subterms [e1] and [e2], respectively.
Adam Chlipala's avatar
Adam Chlipala committed
367

368
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. %\index{tactics!intros}\index{tactics!unfold}\index{tactics!fold}% *)
Adam Chlipala's avatar
Adam Chlipala committed
369

Adam Chlipala's avatar
Adam Chlipala committed
370 371 372 373 374
  intros.
  unfold compile.
  fold compile.
  unfold expDenote.
  fold expDenote.
Adam Chlipala's avatar
Adam Chlipala committed
375

Adam Chlipala's avatar
Adam Chlipala committed
376
(** Now we arrive at a point where the tactics we have seen so far are insufficient.  No further definition unfoldings get us anywhere, so we will need to try something different.
Adam Chlipala's avatar
Adam Chlipala committed
377 378 379 380 381 382 383 384 385 386 387 388

[[
  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
  ============================
389
   progDenote ((compile e2 ++ compile e1 ++ iBinop b :: nil) ++ p) s =
Adam Chlipala's avatar
Adam Chlipala committed
390
   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
391
 
Adam Chlipala's avatar
Adam Chlipala committed
392 393
]]

394 395 396
What we need is the associative law of list concatenation, which is available as a theorem [app_assoc_reverse] in the standard library.%\index{Vernacular commands!Check}% *)

Check app_assoc.
397
(** %\vspace{-.15in}%[[
398
app_assoc_reverse
Adam Chlipala's avatar
Adam Chlipala committed
399
     : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
400
 
Adam Chlipala's avatar
Adam Chlipala committed
401 402
]]

403
If we did not already know the name of the theorem, we could use the %\index{Vernacular commands!SearchRewrite}%[SearchRewrite] command to find it, based on a pattern that we would like to rewrite: *)
404 405

SearchRewrite ((_ ++ _) ++ _).
406
(** %\vspace{-.15in}%[[
407 408 409 410 411 412
app_assoc_reverse:
  forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
]]
%\vspace{-.25in}%
[[
app_assoc: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n
413 414 415
 
]]

416
We use [app_assoc_reverse] to perform a rewrite: %\index{tactics!rewrite}% *)
Adam Chlipala's avatar
Adam Chlipala committed
417

418
  rewrite app_assoc_reverse.
Adam Chlipala's avatar
Adam Chlipala committed
419

420
(** %\noindent{}%changing the conclusion to:
Adam Chlipala's avatar
Adam Chlipala committed
421

422
[[
423
   progDenote (compile e2 ++ (compile e1 ++ iBinop b :: nil) ++ p) s =
Adam Chlipala's avatar
Adam Chlipala committed
424
   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
425
 
Adam Chlipala's avatar
Adam Chlipala committed
426 427
]]

428
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.%\index{tactics!rewrite}% *)
Adam Chlipala's avatar
Adam Chlipala committed
429

Adam Chlipala's avatar
Adam Chlipala committed
430
  rewrite IHe2.
Adam Chlipala's avatar
Adam Chlipala committed
431
(** [[
432
   progDenote ((compile e1 ++ iBinop b :: nil) ++ p) (expDenote e2 :: s) =
Adam Chlipala's avatar
Adam Chlipala committed
433
   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
434
 
Adam Chlipala's avatar
Adam Chlipala committed
435 436
]]

437
The same process lets us apply the remaining hypothesis.%\index{tactics!rewrite}% *)
Adam Chlipala's avatar
Adam Chlipala committed
438

439
  rewrite app_assoc_reverse.
Adam Chlipala's avatar
Adam Chlipala committed
440
  rewrite IHe1.
Adam Chlipala's avatar
Adam Chlipala committed
441
(** [[
442
   progDenote ((iBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
Adam Chlipala's avatar
Adam Chlipala committed
443
   progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
444
 
Adam Chlipala's avatar
Adam Chlipala committed
445 446
]]

447
Now we can apply a similar sequence of tactics to the one that ended the proof of the first case.%\index{tactics!unfold}\index{tactics!simpl}\index{tactics!fold}\index{tactics!reflexivity}%
Adam Chlipala's avatar
Adam Chlipala committed
448 449 450
*)

  unfold progDenote at 1.
Adam Chlipala's avatar
Adam Chlipala committed
451
  simpl.
Adam Chlipala's avatar
Adam Chlipala committed
452
  fold progDenote.
Adam Chlipala's avatar
Adam Chlipala committed
453
  reflexivity.
Adam Chlipala's avatar
Adam Chlipala committed
454

455
(** And the proof is completed, as indicated by the message: *)
Adam Chlipala's avatar
Adam Chlipala committed
456

457 458 459 460 461
(**
<<
  Proof completed.
>>
*)
Adam Chlipala's avatar
Adam Chlipala committed
462

463
(** 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.%\index{Vernacular commands!Abort}%
Adam Chlipala's avatar
Adam Chlipala committed
464 465
*)

Adam Chlipala's avatar
Adam Chlipala committed
466 467
Abort.

468 469
(** %\index{tactics!induction}\index{tactics!crush}% *)

Adam Chlipala's avatar
Adam Chlipala committed
470
Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s =
Adam Chlipala's avatar
Adam Chlipala committed
471 472 473 474
  progDenote p (expDenote e :: s).
  induction e; crush.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
475
(** 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 %\index{tactics!semicolon}%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.
Adam Chlipala's avatar
Adam Chlipala committed
476

477
The [crush] tactic comes from the library associated with this book and is not part of the Coq standard library.  The book's library contains a number of other tactics that are especially helpful in highly automated proofs.
478

479
The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it.  The tactic commands we have written above are an example of a _proof script_, or a series of Ltac programs; while [Qed] uses the result of the script to generate a _proof term_, a well-typed term of Gallina.  To believe that a theorem is true, we only need to trust that the (relatively simple) checker for proof terms is correct; the use of proof scripts is immaterial.  Part I of this book will introduce the principles behind encoding all proofs as terms of Gallina.
480 481

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.%\index{tactics!intros}% *)
Adam Chlipala's avatar
Adam Chlipala committed
482

Adam Chlipala's avatar
Adam Chlipala committed
483
Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
Adam Chlipala's avatar
Adam Chlipala committed
484 485 486 487 488
  intros.
(** [[
  e : exp
  ============================
   progDenote (compile e) nil = Some (expDenote e :: nil)
489
 
Adam Chlipala's avatar
Adam Chlipala committed
490 491
]]

Adam Chlipala's avatar
Adam Chlipala committed
492
At this point, we want to massage the lefthand side to match the statement of [compile_correct'].  A theorem from the standard library is useful: *)
Adam Chlipala's avatar
Adam Chlipala committed
493 494 495 496 497

Check app_nil_end.
(** [[
app_nil_end
     : forall (A : Type) (l : list A), l = l ++ nil
498
]]
499
%\index{tactics!rewrite}% *)
Adam Chlipala's avatar
Adam Chlipala committed
500

Adam Chlipala's avatar
Adam Chlipala committed
501
  rewrite (app_nil_end (compile e)).
Adam Chlipala's avatar
Adam Chlipala committed
502

503
(** 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.  The [rewrite] tactic might choose the wrong place to rewrite if we did not specify which we want.
Adam Chlipala's avatar
Adam Chlipala committed
504 505 506 507 508

[[
  e : exp
  ============================
   progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
509
 
Adam Chlipala's avatar
Adam Chlipala committed
510 511
]]

512
Now we can apply the lemma.%\index{tactics!rewrite}% *)
Adam Chlipala's avatar
Adam Chlipala committed
513

Adam Chlipala's avatar
Adam Chlipala committed
514
  rewrite compile_correct'.
Adam Chlipala's avatar
Adam Chlipala committed
515 516 517 518
(** [[
  e : exp
  ============================
   progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
519
 
Adam Chlipala's avatar
Adam Chlipala committed
520 521
]]

522
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 %\index{tactics!reflexivity}%[reflexivity] does the normalization and checks that the two results are syntactically equal.%\index{tactics!reflexivity}% *)
Adam Chlipala's avatar
Adam Chlipala committed
523

Adam Chlipala's avatar
Adam Chlipala committed
524 525
  reflexivity.
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
526
(* end thide *)
527

528
(** This proof can be shortened and made automated, but we leave that task as an exercise for the reader. *)
529

530

Adam Chlipala's avatar
Adam Chlipala committed
531
(** * Typed Expressions *)
532 533 534

(** In this section, we will build on the initial example by adding additional expression forms that depend on static typing of terms for safety. *)

Adam Chlipala's avatar
Adam Chlipala committed
535
(** ** Source Language *)
536

537 538
(** We define a trivial language of types to classify our expressions: *)

539 540
Inductive type : Set := Nat | Bool.

541 542 543
(** Like most programming languages, Coq uses case-sensitive variable names, so that our user-defined type [type] is distinct from the [Type] keyword that we have already seen appear in the statement of a polymorphic theorem (and that we will meet in more detail later), and our constructor names [Nat] and [Bool] are distinct from the types [nat] and [bool] in the standard library.

   Now we define an expanded set of binary operators. *)
544

545 546 547 548 549 550
Inductive tbinop : type -> type -> type -> Set :=
| TPlus : tbinop Nat Nat Nat
| TTimes : tbinop Nat Nat Nat
| TEq : forall t, tbinop t t Bool
| TLt : tbinop Nat Nat Bool.

551
(** The definition of [tbinop] is different from [binop] in an important way.  Where we declared that [binop] has type [Set], here we declare that [tbinop] has type [type -> type -> type -> Set].  We define [tbinop] as an _indexed type family_.  Indexed inductive types are at the heart of Coq's expressive power; almost everything else of interest is defined in terms of them.
552

Adam Chlipala's avatar
Adam Chlipala committed
553
The intuitive explanation of [tbinop] is that a [tbinop t1 t2 t] is a binary operator whose operands should have types [t1] and [t2], and whose result has type [t].  For instance, constructor [TLt] (for less-than comparison of numbers) is assigned type [tbinop Nat Nat Bool], meaning the operator's arguments are naturals and its result is Boolean.  The type of [TEq] introduces a small bit of additional complication via polymorphism: we want to allow equality comparison of any two values of any type, as long as they have the _same_ type.
554

555 556
ML and Haskell have indexed algebraic datatypes.  For instance, their list types are indexed by the type of data that the list carries.  However, compared to Coq, ML and Haskell 98 place two important restrictions on datatype definitions.

557
First, the indices of the range of each data constructor must be type variables bound at the top level of the datatype definition.  There is no way to do what we did here, where we, for instance, say that [TPlus] is a constructor building a [tbinop] whose indices are all fixed at [Nat].  %\index{generalized algebraic datatypes}\index{GADTs|see{generalized algebraic datatypes}}% _Generalized algebraic datatypes_ (GADTs)%~\cite{GADT}% are a popular feature in %\index{GHC Haskell}%GHC Haskell and other languages that removes this first restriction.
558

559
The second restriction is not lifted by GADTs.  In ML and Haskell, indices of types must be types and may not be _expressions_.  In Coq, types may be indexed by arbitrary Gallina terms.  Type indices can live in the same universe as programs, and we can compute with them just like regular programs.  Haskell supports a hobbled form of computation in type indices based on %\index{Haskell}%multi-parameter type classes, and recent extensions like type functions bring Haskell programming even closer to "real" functional programming with types, but, without dependent typing, there must always be a gap between how one programs with types and how one programs normally.
560 561
*)

562
(** We can define a similar type family for typed expressions, where a term of type [texp t] can be assigned object language type [t].  (It is conventional in the world of interactive theorem proving to call the language of the proof assistant the%\index{meta language}% _meta language_ and a language being formalized the%\index{object language}% _object language_.) *)
563

564 565 566
Inductive texp : type -> Set :=
| TNConst : nat -> texp Nat
| TBConst : bool -> texp Bool
567
| TBinop : forall t1 t2 t, tbinop t1 t2 t -> texp t1 -> texp t2 -> texp t.
568

569
(** Thanks to our use of dependent types, every well-typed [texp] represents a well-typed source expression, by construction.  This turns out to be very convenient for many things we might want to do with expressions.  For instance, it is easy to adapt our interpreter approach to defining semantics.  We start by defining a function mapping the types of our object language into Coq types: *)
570

571 572 573 574 575 576
Definition typeDenote (t : type) : Set :=
  match t with
    | Nat => nat
    | Bool => bool
  end.

577
(** It can take a few moments to come to terms with the fact that [Set], the type of types of programs, is itself a first-class type, and that we can write functions that return [Set]s.  Past that wrinkle, the definition of [typeDenote] is trivial, relying on the [nat] and [bool] types from the Coq standard library.  We can interpret binary operators by relying on standard-library equality test functions [eqb] and [beq_nat] for Booleans and naturals, respectively, along with a less-than test [leb]: *)
578

Adam Chlipala's avatar
Adam Chlipala committed
579 580 581 582 583
Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
  : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
  match b with
    | TPlus => plus
    | TTimes => mult
584 585
    | TEq Nat => beq_nat
    | TEq Bool => eqb
586
    | TLt => leb
Adam Chlipala's avatar
Adam Chlipala committed
587 588
  end.

589
(** This function has just a few differences from the denotation functions we saw earlier.  First, [tbinop] is an indexed type, so its indices become additional arguments to [tbinopDenote].  Second, we need to perform a genuine%\index{dependent pattern matching}% _dependent pattern match_, where the necessary _type_ of each case body depends on the _value_ that has been matched.  At this early stage, we will not go into detail on the many subtle aspects of Gallina that support dependent pattern-matching, but the subject is central to Part II of the book.
590

591 592 593
The same tricks suffice to define an expression denotation function in an unsurprising way:
*)

Adam Chlipala's avatar
Adam Chlipala committed
594 595
Fixpoint texpDenote t (e : texp t) : typeDenote t :=
  match e with
596 597 598 599 600
    | TNConst n => n
    | TBConst b => b
    | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
  end.

601 602 603
(** We can evaluate a few example programs to convince ourselves that this semantics is correct. *)

Eval simpl in texpDenote (TNConst 42).
Adam Chlipala's avatar
Adam Chlipala committed
604 605
(** [= 42 : typeDenote Nat] *)

606 607 608
(* begin hide *)
Eval simpl in texpDenote (TBConst false).
(* end hide *)
609
Eval simpl in texpDenote (TBConst true).
Adam Chlipala's avatar
Adam Chlipala committed
610 611
(** [= true : typeDenote Bool] *)

612 613
Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2))
  (TNConst 7)).
Adam Chlipala's avatar
Adam Chlipala committed
614 615
(** [= 28 : typeDenote Nat] *)

616 617
Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2))
  (TNConst 7)).
618
(** [= false : typeDenote Bool] *)
Adam Chlipala's avatar
Adam Chlipala committed
619

620 621
Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
  (TNConst 7)).
Adam Chlipala's avatar
Adam Chlipala committed
622
(** [= true : typeDenote Bool] *)
623

624 625
(** %\smallskip{}%Now we are ready to define a suitable stack machine target for compilation. *)

626

Adam Chlipala's avatar
Adam Chlipala committed
627
(** ** Target Language *)
628

629
(** In the example of the untyped language, stack machine programs could encounter stack underflows and "get stuck."  This was unfortunate, since we had to deal with this complication even though we proved that our compiler never produced underflowing programs.  We could have used dependent types to force all stack machine programs to be underflow-free.
Adam Chlipala's avatar
Adam Chlipala committed
630 631 632 633 634

For our new languages, besides underflow, we also have the problem of stack slots with naturals instead of bools or vice versa.  This time, we will use indexed typed families to avoid the need to reason about potential failures.

We start by defining stack types, which classify sets of possible stacks. *)

635 636
Definition tstack := list type.

Adam Chlipala's avatar
Adam Chlipala committed
637 638 639 640
(** Any stack classified by a [tstack] must have exactly as many elements, and each stack element must have the type found in the same position of the stack type.

We can define instructions in terms of stack types, where every instruction's type tells us what initial stack type it expects and what final stack type it will produce. *)

641
Inductive tinstr : tstack -> tstack -> Set :=
642 643
| TiNConst : forall s, nat -> tinstr s (Nat :: s)
| TiBConst : forall s, bool -> tinstr s (Bool :: s)
644
| TiBinop : forall arg1 arg2 res s,
645 646 647
  tbinop arg1 arg2 res
  -> tinstr (arg1 :: arg2 :: s) (res :: s).

Adam Chlipala's avatar
Adam Chlipala committed
648 649
(** Stack machine programs must be a similar inductive family, since, if we again used the [list] type family, we would not be able to guarantee that intermediate stack types match within a program. *)

650 651 652 653 654 655 656
Inductive tprog : tstack -> tstack -> Set :=
| TNil : forall s, tprog s s
| TCons : forall s1 s2 s3,
  tinstr s1 s2
  -> tprog s2 s3
  -> tprog s1 s3.

Adam Chlipala's avatar
Adam Chlipala committed
657 658
(** Now, to define the semantics of our new target language, we need a representation for stacks at runtime.  We will again take advantage of type information to define types of value stacks that, by construction, contain the right number and types of elements. *)

659 660 661 662 663 664
Fixpoint vstack (ts : tstack) : Set :=
  match ts with
    | nil => unit
    | t :: ts' => typeDenote t * vstack ts'
  end%type.

665
(** This is another [Set]-valued function.  This time it is recursive, which is perfectly valid, since [Set] is not treated specially in determining which functions may be written.  We say that the value stack of an empty stack type is any value of type [unit], which has just a single value, [tt].  A nonempty stack type leads to a value stack that is a pair, whose first element has the proper type and whose second element follows the representation for the remainder of the stack type.  We write [%]%\index{notation scopes}\coqdocvar{%#<tt>#type#</tt>#%}% as an instruction to Coq's extensible parser.  In particular, this directive applies to the whole [match] expression, which we ask to be parsed as though it were a type, so that the operator [*] is interpreted as Cartesian product instead of, say, multiplication.  (Note that this use of %\coqdocvar{%#<tt>#type#</tt>#%}% has no connection to the inductive type [type] that we have defined.)
Adam Chlipala's avatar
Adam Chlipala committed
666

667
This idea of programming with types can take a while to internalize, but it enables a very simple definition of instruction denotation.  Our definition is like what you might expect from a Lisp-like version of ML that ignored type information.  Nonetheless, the fact that [tinstrDenote] passes the type-checker guarantees that our stack machine programs can never go wrong.  We use a special form of [let] to destructure a multi-level tuple. *)
Adam Chlipala's avatar
Adam Chlipala committed
668

669
Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
Adam Chlipala's avatar
Adam Chlipala committed
670
  match i with
671 672
    | TiNConst _ n => fun s => (n, s)
    | TiBConst _ b => fun s => (b, s)
673
    | TiBinop _ _ _ _ b => fun s =>
674 675
      let '(arg1, (arg2, s')) := s in
        ((tbinopDenote b) arg1 arg2, s')
676 677
  end.

Adam Chlipala's avatar
Adam Chlipala committed
678 679 680
(** Why do we choose to use an anonymous function to bind the initial stack in every case of the [match]?  Consider this well-intentioned but invalid alternative version:
[[
Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
Adam Chlipala's avatar
Adam Chlipala committed
681
  match i with
682 683
    | TiNConst _ n => (n, s)
    | TiBConst _ b => (b, s)
684
    | TiBinop _ _ _ _ b =>
685 686
      let '(arg1, (arg2, s')) := s in
        ((tbinopDenote b) arg1 arg2, s')
Adam Chlipala's avatar
Adam Chlipala committed
687
  end.
688 689
]]

690
The Coq type checker complains that:
Adam Chlipala's avatar
Adam Chlipala committed
691

692
<<
Adam Chlipala's avatar
Adam Chlipala committed
693
The term "(n, s)" has type "(nat * vstack ts)%type"
Adam Chlipala's avatar
Adam Chlipala committed
694
 while it is expected to have type "vstack ?119".
695
>>
Adam Chlipala's avatar
Adam Chlipala committed
696

697
This and other mysteries of Coq dependent typing we postpone until Part II of the book.  The upshot of our later discussion is that it is often useful to push inside of [match] branches those function parameters whose types depend on the type of the value being matched.  Our later, more complete treatement of Gallina's typing rules will explain why this helps.
Adam Chlipala's avatar
Adam Chlipala committed
698 699 700 701
*)

(** We finish the semantics with a straightforward definition of program denotation. *)

Adam Chlipala's avatar
Adam Chlipala committed
702 703
Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' :=
  match p with
704 705 706 707
    | TNil _ => fun s => s
    | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
  end.

708 709
(** The same argument-postponing trick is crucial for this definition. *)

710 711 712

(** ** Translation *)

Adam Chlipala's avatar
Adam Chlipala committed
713 714
(** To define our compilation, it is useful to have an auxiliary function for concatenating two stack machine programs. *)

Adam Chlipala's avatar
Adam Chlipala committed
715 716
Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') : tprog ts' ts'' -> tprog ts ts'' :=
  match p with
717 718 719 720
    | TNil _ => fun p' => p'
    | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
  end.

Adam Chlipala's avatar
Adam Chlipala committed
721 722
(** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)

Adam Chlipala's avatar
Adam Chlipala committed
723 724
Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) :=
  match e with
725 726
    | TNConst n => TCons (TiNConst _ n) (TNil _)
    | TBConst b => TCons (TiBConst _ b) (TNil _)
727
    | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
728
      (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _)))
729 730
  end.

731
(** One interesting feature of the definition is the underscores appearing to the right of [=>] arrows.  Haskell and ML programmers are quite familiar with compilers that infer type parameters to polymorphic values.  In Coq, it is possible to go even further and ask the system to infer arbitrary terms, by writing underscores in place of specific values.  You may have noticed that we have been calling functions without specifying all of their arguments.  For instance, the recursive calls here to [tcompile] omit the [t] argument.  Coq's _implicit argument_ mechanism automatically inserts underscores for arguments that it will probably be able to infer.  Inference of such values is far from complete, though; generally, it only works in cases similar to those encountered with polymorphic type instantiation in Haskell and ML.
Adam Chlipala's avatar
Adam Chlipala committed
732 733 734

The underscores here are being filled in with stack types.  That is, the Coq type inferencer is, in a sense, inferring something about the flow of control in the translated programs.  We can take a look at exactly which values are filled in: *)

735
Print tcompile.
736
(** %\vspace{-.15in}%[[
Adam Chlipala's avatar
Adam Chlipala committed
737 738 739 740
tcompile = 
fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
  tprog ts (t :: ts) :=
  match e in (texp t0) return (tprog ts (t0 :: ts)) with
741 742
  | TNConst n => TCons (TiNConst ts n) (TNil (Nat :: ts))
  | TBConst b => TCons (TiBConst ts b) (TNil (Bool :: ts))
Adam Chlipala's avatar
Adam Chlipala committed
743 744 745
  | TBinop arg1 arg2 res b e1 e2 =>
      tconcat (tcompile arg2 e2 ts)
        (tconcat (tcompile arg1 e1 (arg2 :: ts))
746
           (TCons (TiBinop ts b) (TNil (res :: ts))))
Adam Chlipala's avatar
Adam Chlipala committed
747 748
  end
     : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts)
749 750
]]
*)
Adam Chlipala's avatar
Adam Chlipala committed
751 752 753 754 755


(** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)

Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
756
(** [= (42, tt) : vstack (Nat :: nil)] *)
Adam Chlipala's avatar
Adam Chlipala committed
757

Adam Chlipala's avatar
Adam Chlipala committed
758
Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
759
(** [= (true, tt) : vstack (Bool :: nil)] *)
Adam Chlipala's avatar
Adam Chlipala committed
760

761 762
Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2)
  (TNConst 2)) (TNConst 7)) nil) tt.
763
(** [= (28, tt) : vstack (Nat :: nil)] *)
Adam Chlipala's avatar
Adam Chlipala committed
764

765 766
Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2)
  (TNConst 2)) (TNConst 7)) nil) tt.
767
(** [= (false, tt) : vstack (Bool :: nil)] *)
Adam Chlipala's avatar
Adam Chlipala committed
768

769 770
Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
  (TNConst 7)) nil) tt.
771
(** [= (true, tt) : vstack (Bool :: nil)] *)
Adam Chlipala's avatar
Adam Chlipala committed
772

773 774
(** %\smallskip{}%The compiler seems to be working, so let us turn to proving that it _always_ works. *)

775

Adam Chlipala's avatar
Adam Chlipala committed
776 777 778 779
(** ** Translation Correctness *)

(** We can state a correctness theorem similar to the last one. *)

Adam Chlipala's avatar
Adam Chlipala committed
780 781
Theorem tcompile_correct : forall t (e : texp t),
  tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
Adam Chlipala's avatar
Adam Chlipala committed
782 783 784
(* begin hide *)
Abort.
(* end hide *)
Adam Chlipala's avatar
Adam Chlipala committed
785
(* begin thide *)
Adam Chlipala's avatar
Adam Chlipala committed
786

787
(** Again, we need to strengthen the theorem statement so that the induction will go through.  This time, to provide an excuse to demonstrate different tactics, I will develop an alternative approach to this kind of proof, stating the key lemma as: *)
788

Adam Chlipala's avatar
Adam Chlipala committed
789 790
Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
  tprogDenote (tcompile e ts) s = (texpDenote e, s).
Adam Chlipala's avatar
Adam Chlipala committed
791

792
(** While lemma [compile_correct'] quantified over a program that is the "continuation"%~\cite{continuations}% for the expression we are considering, here we avoid drawing in any extra syntactic elements.  In addition to the source expression and its type, we also quantify over an initial stack type and a stack compatible with it.  Running the compilation of the program starting from that stack, we should arrive at a stack that differs only in having the program's denotation pushed onto it.
Adam Chlipala's avatar
Adam Chlipala committed
793 794 795

   Let us try to prove this theorem in the same way that we settled on in the last section. *)

796
  induction e; crush.
Adam Chlipala's avatar
Adam Chlipala committed
797 798 799 800 801 802 803

(** We are left with this unproved conclusion:

[[
tprogDenote
     (tconcat (tcompile e2 ts)
        (tconcat (tcompile e1 (arg2 :: ts))
804
           (TCons (TiBinop ts t) (TNil (res :: ts))))) s =
Adam Chlipala's avatar
Adam Chlipala committed
805
   (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
Adam Chlipala's avatar
Adam Chlipala committed
806
 
Adam Chlipala's avatar
Adam Chlipala committed
807 808
]]

809
We need an analogue to the [app_assoc_reverse] theorem that we used to rewrite the goal in the last section.  We can abort this proof and prove such a lemma about [tconcat].
Adam Chlipala's avatar
Adam Chlipala committed
810
*)
Adam Chlipala's avatar
Adam Chlipala committed
811

812 813
Abort.

Adam Chlipala's avatar
Adam Chlipala committed
814
Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
815 816 817 818 819 820
  (s : vstack ts),
  tprogDenote (tconcat p p') s
  = tprogDenote p' (tprogDenote p s).
  induction p; crush.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
821 822
(** This one goes through completely automatically.

823
Some code behind the scenes registers [app_assoc_reverse] for use by [crush].  We must register [tconcat_correct] similarly to get the same effect:%\index{Vernacular commands!Hint Rewrite}% *)
Adam Chlipala's avatar
Adam Chlipala committed
824

825
Hint Rewrite tconcat_correct.
826

827
(** Here we meet the pervasive concept of a _hint_.  Many proofs can be found through exhaustive enumerations of combinations of possible proof steps; hints provide the set of steps to consider.  The tactic [crush] is applying such brute force search for us silently, and it will consider more possibilities as we add more hints.  This particular hint asks that the lemma be used for left-to-right rewriting.
828

829
Now we are ready to return to [tcompile_correct'], proving it automatically this time. *)
Adam Chlipala's avatar
Adam Chlipala committed
830

Adam Chlipala's avatar
Adam Chlipala committed
831 832
Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
  tprogDenote (tcompile e ts) s = (texpDenote e, s).
833 834 835
  induction e; crush.
Qed.

Adam Chlipala's avatar
Adam Chlipala committed
836 837
(** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)

838
Hint Rewrite tcompile_correct'.
839

Adam Chlipala's avatar
Adam Chlipala committed
840 841
Theorem tcompile_correct : forall t (e : texp t),
  tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
842 843
  crush.
Qed.
Adam Chlipala's avatar
Adam Chlipala committed
844
(* end thide *)
845

846
(** It is probably worth emphasizing that we are doing more than building mathematical models.  Our compilers are functional programs that can be executed efficiently.  One strategy for doing so is based on%\index{program extraction}% _program extraction_, which generates OCaml code from Coq developments.  For instance, we run a command to output the OCaml version of [tcompile]:%\index{Vernacular commands!Extraction}% *)
847 848 849 850 851 852 853 854 855 856 857 858 859 860 861 862 863 864 865 866 867 868 869

Extraction tcompile.

(** <<
let rec tcompile t e ts =
  match e with
  | TNConst n ->
    TCons (ts, (Cons (Nat, ts)), (Cons (Nat, ts)), (TiNConst (ts, n)), (TNil
      (Cons (Nat, ts))))
  | TBConst b ->
    TCons (ts, (Cons (Bool, ts)), (Cons (Bool, ts)), (TiBConst (ts, b)),
      (TNil (Cons (Bool, ts))))
  | TBinop (t1, t2, t0, b, e1, e2) ->
    tconcat ts (Cons (t2, ts)) (Cons (t0, ts)) (tcompile t2 e2 ts)
      (tconcat (Cons (t2, ts)) (Cons (t1, (Cons (t2, ts)))) (Cons (t0, ts))
        (tcompile t1 e1 (Cons (t2, ts))) (TCons ((Cons (t1, (Cons (t2,
        ts)))), (Cons (t0, ts)), (Cons (t0, ts)), (TiBinop (t1, t2, t0, ts,
        b)), (TNil (Cons (t0, ts))))))
>>

We can compile this code with the usual OCaml compiler and obtain an executable program with halfway decent performance.

This chapter has been a whirlwind tour through two examples of the style of Coq development that I advocate.  Parts II and III of the book focus on the key elements of that style, namely dependent types and scripted proof automation, respectively.  Before we get there, we will spend some time in Part I on more standard foundational material.  Part I may still be of interest to seasoned Coq hackers, since I follow the highly automated proof style even at that early stage. *)