annotate src/StackMachine.v @ 302:7b38729be069

Tweak mark-up to support coqdoc 8.3
author Adam Chlipala <adam@chlipala.net>
date Mon, 17 Jan 2011 15:12:30 -0500
parents 123f466faedc
children f7ac1e0c0d60
rev   line source
adam@298 1 (* Copyright (c) 2008-2011, Adam Chlipala
adamc@2 2 *
adamc@2 3 * This work is licensed under a
adamc@2 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@2 5 * Unported License.
adamc@2 6 * The license text is available at:
adamc@2 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@2 8 *)
adamc@2 9
adamc@3 10 (* begin hide *)
adam@277 11 Require Import Arith Bool List.
adamc@2 12
adamc@2 13 Require Import Tactics.
adamc@14 14
adamc@14 15 Set Implicit Arguments.
adamc@3 16 (* end hide *)
adamc@2 17
adamc@2 18
adamc@25 19 (** %\chapter{Some Quick Examples}% *)
adamc@25 20
adamc@25 21
adam@279 22 (** 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.
adam@279 23
adam@298 24 I assume that you have installed Coq and Proof General. The code in this book is tested with Coq versions 8.2pl2 and 8.3pl1, though parts may work with other versions.
adamc@9 25
adam@278 26 To set up your Proof General environment to process the source to this chapter, a few simple steps are required.
adamc@25 27
adam@278 28 %\begin{enumerate}%#<ol>#
adam@278 29
adam@278 30 %\item %#<li>#Get the book source from
adam@278 31 %\begin{center}\url{http://adam.chlipala.net/cpdt/cpdt.tgz}\end{center}%#<blockquote><tt><a href="http://adam.chlipala.net/cpdt/cpdt.tgz">http://adam.chlipala.net/cpdt/cpdt.tgz</a></tt></blockquote></li>#
adam@278 32
adam@278 33 %\item %#<li>#Unpack the tarball to some directory %\texttt{%#<tt>#DIR#</tt>#%}%.#</li>#
adam@278 34
adam@278 35 %\item %#<li>#Run %\texttt{%#<tt>#make#</tt>#%}% in %\texttt{%#<tt>#DIR#</tt>#%}%.#</li>#
adam@278 36
adam@278 37 %\item %#<li>#There are some minor headaches associated with getting Proof General to pass the proper command line arguments to the %\texttt{%#<tt>#coqtop#</tt>#%}% program, which provides the interactive Coq toplevel. The best way to add settings that will be shared by many source files is to add a custom variable setting to your %\texttt{%#<tt>#.emacs#</tt>#%}% file, like this:
adamc@25 38 %\begin{verbatim}%#<pre>#(custom-set-variables
adamc@25 39 ...
adam@278 40 '(coq-prog-args '("-I" "DIR/src"))
adamc@25 41 ...
adamc@25 42 )#</pre>#%\end{verbatim}%
adam@278 43 The extra arguments demonstrated here are the proper choices for working with the code for this book. The ellipses stand for other Emacs customization settings you may already have. It can be helpful to save several alternate sets of flags in your %\texttt{%#<tt>#.emacs#</tt>#%}% file, with all but one commented out within the %\texttt{%#<tt>#custom-set-variables#</tt>#%}% block at any given time.#</li>#
adam@278 44
adam@278 45 #</ol>#%\end{enumerate}%
adam@278 46
adam@278 47 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 two lines [Require Import Arith Bool List Tactics.] and [Set Implicit Arguments.] at the start of the file, to match some code hidden in this rendering of the chapter source, and be sure to run the Coq binary %\texttt{%#<tt>#coqtop#</tt>#%}% with the command-line argument %\texttt{%#<tt>#-I DIR/src#</tt>#%}%. If you have installed Proof General properly, it should start automatically when you visit a %\texttt{%#<tt>#.v#</tt>#%}% buffer in Emacs.
adamc@11 48
adamc@11 49 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. *)
adamc@9 50
adamc@9 51
adamc@20 52 (** * Arithmetic Expressions Over Natural Numbers *)
adamc@2 53
adamc@40 54 (** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *)
adamc@9 55
adamc@20 56 (** ** Source Language *)
adamc@9 57
adamc@9 58 (** We begin with the syntax of the source language. *)
adamc@2 59
adamc@4 60 Inductive binop : Set := Plus | Times.
adamc@2 61
adamc@9 62 (** Our first line of Coq code should be unsurprising to ML and Haskell programmers. We define an 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 %\texttt{%#<tt>#data#</tt>#%}%, %\texttt{%#<tt>#datatype#</tt>#%}%, or %\texttt{%#<tt>#type#</tt>#%}%. 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 [: 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. *)
adamc@9 63
adamc@4 64 Inductive exp : Set :=
adamc@4 65 | Const : nat -> exp
adamc@4 66 | Binop : binop -> exp -> exp -> exp.
adamc@2 67
adamc@9 68 (** 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.
adamc@9 69
adam@277 70 A note for readers following along in the PDF version: coqdoc supports pretty-printing of tokens in LaTeX or HTML. Where you see a right arrow character, the source contains the ASCII text %\texttt{%#<tt>#->#</tt>#%}%. Other examples of this substitution appearing in this chapter are a double right arrow for %\texttt{%#<tt>#=>#</tt>#%}%, the inverted 'A' symbol for %\texttt{%#<tt>#forall#</tt>#%}%, and the Cartesian product 'X' for %\texttt{%#<tt>#*#</tt>#%}%. When in doubt about the ASCII version of a symbol, you can consult the chapter source code.
adamc@9 71
adamc@9 72 %\medskip%
adamc@9 73
adam@292 74 Now we are ready to say what these programs mean. We will do this by writing an 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.) *)
adamc@9 75
adamc@4 76 Definition binopDenote (b : binop) : nat -> nat -> nat :=
adamc@4 77 match b with
adamc@4 78 | Plus => plus
adamc@4 79 | Times => mult
adamc@4 80 end.
adamc@2 81
adamc@9 82 (** The meaning of a binary operator is a binary function over naturals, defined with pattern-matching notation analogous to the %\texttt{%#<tt>#case#</tt>#%}% and %\texttt{%#<tt>#match#</tt>#%}% 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:
adamc@9 83
adamc@9 84 [[
adamc@9 85 Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
adamc@9 86 match b with
adamc@9 87 | Plus => plus
adamc@9 88 | Times => mult
adamc@9 89 end.
adamc@9 90
adamc@205 91 ]]
adamc@205 92
adamc@9 93 In this example, we could also omit all of the type annotations, arriving at:
adamc@9 94
adamc@9 95 [[
adamc@9 96 Definition binopDenote := fun b =>
adamc@9 97 match b with
adamc@9 98 | Plus => plus
adamc@9 99 | Times => mult
adamc@9 100 end.
adamc@9 101
adamc@205 102 ]]
adamc@205 103
adam@292 104 Languages like Haskell and ML have a convenient %\textit{%#<i>#principal typing#</i>#%}% 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.
adamc@9 105
adam@294 106 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 %\textit{%#<i>#Calculus of Inductive Constructions (CIC)#</i>#%}%, which is an extension of the older %\textit{%#<i>#Calculus of Constructions (CoC)#</i>#%}%. 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 %\textit{%#<i>#strong normalization#</i>#%}%, meaning that every program (and, more importantly, every proof term) terminates; and %\textit{%#<i>#relative consistency#</i>#%}% with systems like versions of 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.
adamc@9 107
adamc@40 108 Coq is actually based on an extension of CIC called %\textit{%#<i>#Gallina#</i>#%}%. 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.
adamc@9 109
adamc@9 110 Commands like [Inductive] and [Definition] are part of %\textit{%#<i>#the vernacular#</i>#%}%, which includes all sorts of useful queries and requests to the Coq system.
adamc@9 111
adamc@9 112 Finally, there is %\textit{%#<i>#Ltac#</i>#%}%, 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.
adamc@9 113
adamc@9 114 %\medskip%
adamc@9 115
adamc@9 116 We can give a simple definition of the meaning of an expression: *)
adamc@9 117
adamc@4 118 Fixpoint expDenote (e : exp) : nat :=
adamc@4 119 match e with
adamc@4 120 | Const n => n
adamc@4 121 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
adamc@4 122 end.
adamc@2 123
adamc@9 124 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *)
adamc@2 125
adamc@16 126 (** 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. *)
adamc@16 127
adamc@16 128 Eval simpl in expDenote (Const 42).
adamc@205 129 (** [= 42 : nat] *)
adamc@205 130
adamc@16 131 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)).
adamc@205 132 (** [= 4 : nat] *)
adamc@205 133
adamc@16 134 Eval simpl in expDenote (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adamc@205 135 (** [= 28 : nat] *)
adamc@9 136
adamc@20 137 (** ** Target Language *)
adamc@4 138
adamc@10 139 (** We will compile our source programs onto a simple stack machine, whose syntax is: *)
adamc@2 140
adamc@4 141 Inductive instr : Set :=
adamc@4 142 | IConst : nat -> instr
adamc@4 143 | IBinop : binop -> instr.
adamc@2 144
adamc@4 145 Definition prog := list instr.
adamc@4 146 Definition stack := list nat.
adamc@2 147
adamc@10 148 (** 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.
adamc@10 149
adam@292 150 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']. [::] is the %``%#"#list cons#"#%''% operator from the Coq standard library. *)
adamc@10 151
adamc@4 152 Definition instrDenote (i : instr) (s : stack) : option stack :=
adamc@4 153 match i with
adamc@4 154 | IConst n => Some (n :: s)
adamc@4 155 | IBinop b =>
adamc@4 156 match s with
adamc@4 157 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
adamc@4 158 | _ => None
adamc@4 159 end
adamc@4 160 end.
adamc@2 161
adamc@206 162 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program.
adamc@10 163
adamc@206 164 [[
adamc@4 165 Fixpoint progDenote (p : prog) (s : stack) {struct p} : option stack :=
adamc@4 166 match p with
adamc@4 167 | nil => Some s
adamc@4 168 | i :: p' =>
adamc@4 169 match instrDenote i s with
adamc@4 170 | None => None
adamc@4 171 | Some s' => progDenote p' s'
adamc@4 172 end
adamc@4 173 end.
adamc@2 174
adamc@206 175 ]]
adamc@206 176
adamc@206 177 There is one interesting difference compared to our previous example of a [Fixpoint]. This recursive function takes two arguments, [p] and [s]. It is critical for the soundness of Coq that every program terminate, so a shallow syntactic termination check is imposed on every recursive function definition. One of the function parameters must be designated to decrease monotonically across recursive calls. That is, every recursive call must use a version of that argument that has been pulled out of the current argument by some number of [match] expressions. [expDenote] has only one argument, so we did not need to specify which of its arguments decreases. For [progDenote], we resolve the ambiguity by writing [{struct p}] to indicate that argument [p] decreases structurally.
adamc@206 178
adamc@206 179 Recent versions of Coq will also infer a termination argument, so that we may write simply: *)
adamc@206 180
adamc@206 181 Fixpoint progDenote (p : prog) (s : stack) : option stack :=
adamc@206 182 match p with
adamc@206 183 | nil => Some s
adamc@206 184 | i :: p' =>
adamc@206 185 match instrDenote i s with
adamc@206 186 | None => None
adamc@206 187 | Some s' => progDenote p' s'
adamc@206 188 end
adamc@206 189 end.
adamc@2 190
adamc@4 191
adamc@9 192 (** ** Translation *)
adamc@4 193
adamc@10 194 (** Our compiler itself is now unsurprising. [++] is the list concatenation operator from the Coq standard library. *)
adamc@2 195
adamc@4 196 Fixpoint compile (e : exp) : prog :=
adamc@4 197 match e with
adamc@4 198 | Const n => IConst n :: nil
adamc@4 199 | Binop b e1 e2 => compile e2 ++ compile e1 ++ IBinop b :: nil
adamc@4 200 end.
adamc@2 201
adamc@2 202
adamc@16 203 (** Before we set about proving that this compiler is correct, we can try a few test runs, using our sample programs from earlier. *)
adamc@16 204
adamc@16 205 Eval simpl in compile (Const 42).
adamc@206 206 (** [= IConst 42 :: nil : prog] *)
adamc@206 207
adamc@16 208 Eval simpl in compile (Binop Plus (Const 2) (Const 2)).
adamc@206 209 (** [= IConst 2 :: IConst 2 :: IBinop Plus :: nil : prog] *)
adamc@206 210
adamc@16 211 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adamc@206 212 (** [= IConst 7 :: IConst 2 :: IConst 2 :: IBinop Plus :: IBinop Times :: nil : prog] *)
adamc@16 213
adamc@40 214 (** We can also run our compiled programs and check that they give the right results. *)
adamc@16 215
adamc@16 216 Eval simpl in progDenote (compile (Const 42)) nil.
adamc@206 217 (** [= Some (42 :: nil) : option stack] *)
adamc@206 218
adamc@16 219 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil.
adamc@206 220 (** [= Some (4 :: nil) : option stack] *)
adamc@206 221
adamc@16 222 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7))) nil.
adamc@206 223 (** [= Some (28 :: nil) : option stack] *)
adamc@16 224
adamc@16 225
adamc@20 226 (** ** Translation Correctness *)
adamc@4 227
adamc@11 228 (** 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: *)
adamc@11 229
adamc@26 230 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 231 (* begin hide *)
adamc@11 232 Abort.
adamc@11 233 (* end hide *)
adamc@22 234 (* begin thide *)
adamc@11 235
adam@292 236 (** 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:
adamc@11 237 *)
adamc@2 238
adamc@206 239 Lemma compile_correct' : forall e p s,
adamc@206 240 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
adamc@11 241
adamc@11 242 (** 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:
adamc@11 243
adamc@11 244 [[
adamc@11 245 1 subgoal
adamc@11 246
adamc@11 247 ============================
adamc@15 248 forall (e : exp) (p : list instr) (s : stack),
adamc@15 249 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
adamc@206 250
adamc@11 251 ]]
adamc@11 252
adamc@11 253 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.
adamc@11 254
adamc@11 255 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.
adamc@11 256
adamc@11 257 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:
adamc@11 258 *)
adamc@11 259
adamc@4 260 induction e.
adamc@2 261
adamc@11 262 (** 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:
adamc@11 263
adamc@11 264 [[
adamc@11 265 2 subgoals
adamc@11 266
adamc@11 267 n : nat
adamc@11 268 ============================
adamc@11 269 forall (s : stack) (p : list instr),
adamc@11 270 progDenote (compile (Const n) ++ p) s =
adamc@11 271 progDenote p (expDenote (Const n) :: s)
adamc@11 272 ]]
adamc@11 273 [[
adamc@11 274 subgoal 2 is:
adamc@11 275 forall (s : stack) (p : list instr),
adamc@11 276 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 277 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 278
adamc@11 279 ]]
adamc@11 280
adamc@11 281 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.
adamc@11 282
adamc@11 283 We begin the first case with another very common tactic.
adamc@11 284 *)
adamc@11 285
adamc@4 286 intros.
adamc@11 287
adamc@11 288 (** The current subgoal changes to:
adamc@11 289 [[
adamc@11 290
adamc@11 291 n : nat
adamc@11 292 s : stack
adamc@11 293 p : list instr
adamc@11 294 ============================
adamc@11 295 progDenote (compile (Const n) ++ p) s =
adamc@11 296 progDenote p (expDenote (Const n) :: s)
adamc@206 297
adamc@11 298 ]]
adamc@11 299
adamc@11 300 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
adamc@11 301
adamc@11 302 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.
adamc@11 303 *)
adamc@11 304
adamc@4 305 unfold compile.
adamc@11 306 (** [[
adamc@11 307 n : nat
adamc@11 308 s : stack
adamc@11 309 p : list instr
adamc@11 310 ============================
adamc@11 311 progDenote ((IConst n :: nil) ++ p) s =
adamc@11 312 progDenote p (expDenote (Const n) :: s)
adamc@206 313
adam@302 314 ]]
adam@302 315 *)
adamc@11 316
adamc@4 317 unfold expDenote.
adam@302 318
adamc@11 319 (** [[
adamc@11 320 n : nat
adamc@11 321 s : stack
adamc@11 322 p : list instr
adamc@11 323 ============================
adamc@11 324 progDenote ((IConst n :: nil) ++ p) s = progDenote p (n :: s)
adamc@206 325
adamc@11 326 ]]
adamc@11 327
adamc@11 328 We only need to unfold the first occurrence of [progDenote] to prove the goal: *)
adamc@11 329
adamc@11 330 unfold progDenote at 1.
adamc@11 331
adamc@11 332 (** [[
adamc@11 333
adamc@11 334 n : nat
adamc@11 335 s : stack
adamc@11 336 p : list instr
adamc@11 337 ============================
adamc@11 338 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 339 option stack :=
adamc@11 340 match p0 with
adamc@11 341 | nil => Some s0
adamc@11 342 | i :: p' =>
adamc@11 343 match instrDenote i s0 with
adamc@11 344 | Some s' => progDenote p' s'
adamc@11 345 | None => None (A:=stack)
adamc@11 346 end
adamc@11 347 end) ((IConst n :: nil) ++ p) s =
adamc@11 348 progDenote p (n :: s)
adamc@206 349
adamc@11 350 ]]
adamc@11 351
adamc@11 352 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:
adamc@11 353 *)
adamc@11 354
adamc@4 355 simpl.
adamc@11 356 (** [[
adamc@11 357 n : nat
adamc@11 358 s : stack
adamc@11 359 p : list instr
adamc@11 360 ============================
adamc@11 361 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 362 option stack :=
adamc@11 363 match p0 with
adamc@11 364 | nil => Some s0
adamc@11 365 | i :: p' =>
adamc@11 366 match instrDenote i s0 with
adamc@11 367 | Some s' => progDenote p' s'
adamc@11 368 | None => None (A:=stack)
adamc@11 369 end
adamc@11 370 end) p (n :: s) = progDenote p (n :: s)
adamc@206 371
adamc@11 372 ]]
adamc@11 373
adamc@11 374 Now we can unexpand the definition of [progDenote]:
adamc@11 375 *)
adamc@11 376
adamc@11 377 fold progDenote.
adamc@11 378
adamc@11 379 (** [[
adamc@11 380
adamc@11 381 n : nat
adamc@11 382 s : stack
adamc@11 383 p : list instr
adamc@11 384 ============================
adamc@11 385 progDenote p (n :: s) = progDenote p (n :: s)
adamc@206 386
adamc@11 387 ]]
adamc@11 388
adamc@11 389 It looks like we are at the end of this case, since we have a trivial equality. Indeed, a single tactic finishes us off:
adamc@11 390 *)
adamc@11 391
adamc@4 392 reflexivity.
adamc@2 393
adamc@11 394 (** On to the second inductive case:
adamc@11 395
adamc@11 396 [[
adamc@11 397 b : binop
adamc@11 398 e1 : exp
adamc@11 399 IHe1 : forall (s : stack) (p : list instr),
adamc@11 400 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 401 e2 : exp
adamc@11 402 IHe2 : forall (s : stack) (p : list instr),
adamc@11 403 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 404 ============================
adamc@11 405 forall (s : stack) (p : list instr),
adamc@11 406 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 407 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 408
adamc@11 409 ]]
adamc@11 410
adamc@11 411 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.
adamc@11 412
adamc@11 413 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. *)
adamc@11 414
adamc@4 415 intros.
adamc@4 416 unfold compile.
adamc@4 417 fold compile.
adamc@4 418 unfold expDenote.
adamc@4 419 fold expDenote.
adamc@11 420
adamc@44 421 (** 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.
adamc@11 422
adamc@11 423 [[
adamc@11 424 b : binop
adamc@11 425 e1 : exp
adamc@11 426 IHe1 : forall (s : stack) (p : list instr),
adamc@11 427 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 428 e2 : exp
adamc@11 429 IHe2 : forall (s : stack) (p : list instr),
adamc@11 430 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 431 s : stack
adamc@11 432 p : list instr
adamc@11 433 ============================
adamc@11 434 progDenote ((compile e2 ++ compile e1 ++ IBinop b :: nil) ++ p) s =
adamc@11 435 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 436
adamc@11 437 ]]
adamc@11 438
adam@277 439 What we need is the associative law of list concatenation, which is available as a theorem [app_ass] in the standard library. *)
adamc@11 440
adamc@11 441 Check app_ass.
adamc@11 442 (** [[
adamc@11 443 app_ass
adamc@11 444 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adamc@206 445
adamc@11 446 ]]
adamc@11 447
adam@277 448 If we did not already know the name of the theorem, we could use the [SearchRewrite] command to find it, based on a pattern that we would like to rewrite: *)
adam@277 449
adam@277 450 SearchRewrite ((_ ++ _) ++ _).
adam@277 451 (** [[
adam@277 452 app_ass: forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adam@277 453 ass_app: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n
adam@277 454
adam@277 455 ]]
adam@277 456
adamc@11 457 We use it to perform a rewrite: *)
adamc@11 458
adamc@4 459 rewrite app_ass.
adamc@11 460
adamc@206 461 (** changing the conclusion to:
adamc@11 462
adamc@206 463 [[
adamc@11 464 progDenote (compile e2 ++ (compile e1 ++ IBinop b :: nil) ++ p) s =
adamc@11 465 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 466
adamc@11 467 ]]
adamc@11 468
adamc@11 469 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: *)
adamc@11 470
adamc@4 471 rewrite IHe2.
adamc@11 472 (** [[
adamc@11 473 progDenote ((compile e1 ++ IBinop b :: nil) ++ p) (expDenote e2 :: s) =
adamc@11 474 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 475
adamc@11 476 ]]
adamc@11 477
adamc@11 478 The same process lets us apply the remaining hypothesis. *)
adamc@11 479
adamc@4 480 rewrite app_ass.
adamc@4 481 rewrite IHe1.
adamc@11 482 (** [[
adamc@11 483 progDenote ((IBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
adamc@11 484 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 485
adamc@11 486 ]]
adamc@11 487
adamc@11 488 Now we can apply a similar sequence of tactics to that that ended the proof of the first case.
adamc@11 489 *)
adamc@11 490
adamc@11 491 unfold progDenote at 1.
adamc@4 492 simpl.
adamc@11 493 fold progDenote.
adamc@4 494 reflexivity.
adamc@11 495
adamc@11 496 (** And the proof is completed, as indicated by the message:
adamc@11 497
adamc@11 498 [[
adamc@11 499 Proof completed.
adamc@11 500
adamc@205 501 ]]
adamc@205 502
adamc@11 503 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.
adamc@11 504 *)
adamc@11 505
adamc@4 506 Abort.
adamc@2 507
adamc@26 508 Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 509 progDenote p (expDenote e :: s).
adamc@4 510 induction e; crush.
adamc@4 511 Qed.
adamc@2 512
adamc@11 513 (** 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.
adamc@11 514
adamc@210 515 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.
adamc@210 516
adamc@11 517 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. *)
adamc@11 518
adamc@26 519 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 520 intros.
adamc@11 521 (** [[
adamc@11 522 e : exp
adamc@11 523 ============================
adamc@11 524 progDenote (compile e) nil = Some (expDenote e :: nil)
adamc@206 525
adamc@11 526 ]]
adamc@11 527
adamc@26 528 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: *)
adamc@11 529
adamc@11 530 Check app_nil_end.
adamc@11 531 (** [[
adamc@11 532 app_nil_end
adamc@11 533 : forall (A : Type) (l : list A), l = l ++ nil
adam@302 534 ]]
adam@302 535 *)
adamc@11 536
adamc@4 537 rewrite (app_nil_end (compile e)).
adamc@11 538
adamc@11 539 (** 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.
adamc@11 540
adamc@11 541 [[
adamc@11 542 e : exp
adamc@11 543 ============================
adamc@11 544 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
adamc@206 545
adamc@11 546 ]]
adamc@11 547
adamc@11 548 Now we can apply the lemma. *)
adamc@11 549
adamc@26 550 rewrite compile_correct'.
adamc@11 551 (** [[
adamc@11 552 e : exp
adamc@11 553 ============================
adamc@11 554 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
adamc@206 555
adamc@11 556 ]]
adamc@11 557
adamc@11 558 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. *)
adamc@11 559
adamc@4 560 reflexivity.
adamc@4 561 Qed.
adamc@22 562 (* end thide *)
adamc@14 563
adamc@14 564
adamc@20 565 (** * Typed Expressions *)
adamc@14 566
adamc@14 567 (** In this section, we will build on the initial example by adding additional expression forms that depend on static typing of terms for safety. *)
adamc@14 568
adamc@20 569 (** ** Source Language *)
adamc@14 570
adamc@15 571 (** We define a trivial language of types to classify our expressions: *)
adamc@15 572
adamc@14 573 Inductive type : Set := Nat | Bool.
adamc@14 574
adam@277 575 (** 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.
adam@277 576
adam@277 577 Now we define an expanded set of binary operators. *)
adamc@15 578
adamc@14 579 Inductive tbinop : type -> type -> type -> Set :=
adamc@14 580 | TPlus : tbinop Nat Nat Nat
adamc@14 581 | TTimes : tbinop Nat Nat Nat
adamc@14 582 | TEq : forall t, tbinop t t Bool
adamc@14 583 | TLt : tbinop Nat Nat Bool.
adamc@14 584
adamc@15 585 (** 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 %\textit{%#<i>#indexed type family#</i>#%}%. Indexed inductive types are at the heart of Coq's expressive power; almost everything else of interest is defined in terms of them.
adamc@15 586
adamc@15 587 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.
adamc@15 588
adamc@15 589 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]. %\textit{%#<i>#Generalized algebraic datatypes (GADTs)#</i>#%}% are a popular feature in GHC Haskell and other languages that removes this first restriction.
adamc@15 590
adam@292 591 The second restriction is not lifted by GADTs. In ML and Haskell, indices of types must be types and may not be %\textit{%#<i>#expressions#</i>#%}%. 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 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.
adamc@15 592 *)
adamc@15 593
adamc@15 594 (** We can define a similar type family for typed expressions. *)
adamc@15 595
adamc@14 596 Inductive texp : type -> Set :=
adamc@14 597 | TNConst : nat -> texp Nat
adamc@14 598 | TBConst : bool -> texp Bool
adamc@14 599 | TBinop : forall arg1 arg2 res, tbinop arg1 arg2 res -> texp arg1 -> texp arg2 -> texp res.
adamc@14 600
adamc@15 601 (** 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 languages into Coq types: *)
adamc@15 602
adamc@14 603 Definition typeDenote (t : type) : Set :=
adamc@14 604 match t with
adamc@14 605 | Nat => nat
adamc@14 606 | Bool => bool
adamc@14 607 end.
adamc@14 608
adamc@15 609 (** 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.
adamc@15 610
adam@292 611 We need to define one auxiliary function, implementing a boolean binary %``%#"#less-than#"#%''% operator, which only appears in the standard library with a type fancier than what we are prepared to deal with here. The code is entirely standard and ML-like, with the one caveat being that the Coq [nat] type uses a unary representation, where [O] is zero and [S n] is the successor of [n].
adamc@15 612 *)
adamc@15 613
adam@279 614 Fixpoint lessThan (n1 n2 : nat) : bool :=
adamc@14 615 match n1, n2 with
adamc@14 616 | O, S _ => true
adam@279 617 | S n1', S n2' => lessThan n1' n2'
adamc@14 618 | _, _ => false
adamc@14 619 end.
adamc@14 620
adam@277 621 (** Now we can interpret binary operators, relying on standard-library equality test functions [eqb] and [beq_nat] for booleans and naturals, respectively: *)
adamc@15 622
adamc@14 623 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@14 624 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@207 625 match b in (tbinop arg1 arg2 res)
adamc@207 626 return (typeDenote arg1 -> typeDenote arg2 -> typeDenote res) with
adamc@14 627 | TPlus => plus
adamc@14 628 | TTimes => mult
adam@277 629 | TEq Nat => beq_nat
adam@277 630 | TEq Bool => eqb
adam@279 631 | TLt => lessThan
adamc@14 632 end.
adamc@14 633
adamc@207 634 (** 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 %\textit{%#<i>#dependent pattern match#</i>#%}% to come up with a definition of this function that type-checks. In each branch of the [match], we need to use branch-specific information about the indices to [tbinop]. General type inference that takes such information into account is undecidable, so it is often necessary to write annotations, like we see above on the line with [match].
adamc@15 635
adamc@273 636 The [in] annotation restates the type of the term being case-analyzed. Though we use the same names for the indices as we use in the type of the original argument binder, these are actually fresh variables, and they are %\textit{%#<i>#binding occurrences#</i>#%}%. Their scope is the [return] clause. That is, [arg1], [arg2], and [res] are new bound variables bound only within the return clause [typeDenote arg1 -> typeDenote arg2 -> typeDenote res]. By being explicit about the functional relationship between the type indices and the match result, we regain decidable type inference.
adamc@15 637
adamc@207 638 In fact, recent Coq versions use some heuristics that can save us the trouble of writing [match] annotations, and those heuristics get the job done in this case. We can get away with writing just: *)
adamc@207 639
adamc@207 640 (* begin hide *)
adamc@207 641 Reset tbinopDenote.
adamc@207 642 (* end hide *)
adamc@207 643 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@207 644 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@207 645 match b with
adamc@207 646 | TPlus => plus
adamc@207 647 | TTimes => mult
adam@277 648 | TEq Nat => beq_nat
adam@277 649 | TEq Bool => eqb
adam@279 650 | TLt => lessThan
adamc@207 651 end.
adamc@207 652
adamc@207 653 (**
adamc@15 654 The same tricks suffice to define an expression denotation function in an unsurprising way:
adamc@15 655 *)
adamc@15 656
adamc@207 657 Fixpoint texpDenote t (e : texp t) : typeDenote t :=
adamc@207 658 match e with
adamc@14 659 | TNConst n => n
adamc@14 660 | TBConst b => b
adamc@14 661 | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
adamc@14 662 end.
adamc@14 663
adamc@17 664 (** We can evaluate a few example programs to convince ourselves that this semantics is correct. *)
adamc@17 665
adamc@17 666 Eval simpl in texpDenote (TNConst 42).
adamc@207 667 (** [= 42 : typeDenote Nat] *)
adamc@207 668
adamc@17 669 Eval simpl in texpDenote (TBConst true).
adamc@207 670 (** [= true : typeDenote Bool] *)
adamc@207 671
adamc@17 672 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
adamc@207 673 (** [= 28 : typeDenote Nat] *)
adamc@207 674
adamc@17 675 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
adamc@207 676 (** [= false : typeDenote Bool] *)
adamc@207 677
adamc@17 678 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
adamc@207 679 (** [= true : typeDenote Bool] *)
adamc@17 680
adamc@14 681
adamc@20 682 (** ** Target Language *)
adamc@14 683
adam@292 684 (** Now we want to define a suitable stack machine target for compilation. 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.
adamc@18 685
adamc@18 686 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.
adamc@18 687
adamc@18 688 We start by defining stack types, which classify sets of possible stacks. *)
adamc@18 689
adamc@14 690 Definition tstack := list type.
adamc@14 691
adamc@18 692 (** 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.
adamc@18 693
adamc@18 694 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. *)
adamc@18 695
adamc@14 696 Inductive tinstr : tstack -> tstack -> Set :=
adamc@14 697 | TINConst : forall s, nat -> tinstr s (Nat :: s)
adamc@14 698 | TIBConst : forall s, bool -> tinstr s (Bool :: s)
adamc@14 699 | TIBinop : forall arg1 arg2 res s,
adamc@14 700 tbinop arg1 arg2 res
adamc@14 701 -> tinstr (arg1 :: arg2 :: s) (res :: s).
adamc@14 702
adamc@18 703 (** 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. *)
adamc@18 704
adamc@14 705 Inductive tprog : tstack -> tstack -> Set :=
adamc@14 706 | TNil : forall s, tprog s s
adamc@14 707 | TCons : forall s1 s2 s3,
adamc@14 708 tinstr s1 s2
adamc@14 709 -> tprog s2 s3
adamc@14 710 -> tprog s1 s3.
adamc@14 711
adamc@18 712 (** 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. *)
adamc@18 713
adamc@14 714 Fixpoint vstack (ts : tstack) : Set :=
adamc@14 715 match ts with
adamc@14 716 | nil => unit
adamc@14 717 | t :: ts' => typeDenote t * vstack ts'
adamc@14 718 end%type.
adamc@14 719
adamc@210 720 (** 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 [%type] so that Coq knows to interpret [*] as Cartesian product rather than multiplication.
adamc@18 721
adamc@207 722 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. *)
adamc@18 723
adamc@14 724 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
adamc@207 725 match i with
adamc@14 726 | TINConst _ n => fun s => (n, s)
adamc@14 727 | TIBConst _ b => fun s => (b, s)
adamc@14 728 | TIBinop _ _ _ _ b => fun s =>
adamc@14 729 match s with
adamc@14 730 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@14 731 end
adamc@14 732 end.
adamc@14 733
adamc@18 734 (** 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:
adamc@18 735
adamc@18 736 [[
adamc@18 737 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
adamc@207 738 match i with
adamc@18 739 | TINConst _ n => (n, s)
adamc@18 740 | TIBConst _ b => (b, s)
adamc@18 741 | TIBinop _ _ _ _ b =>
adamc@18 742 match s with
adamc@18 743 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@18 744 end
adamc@18 745 end.
adamc@18 746
adamc@205 747 ]]
adamc@205 748
adamc@18 749 The Coq type-checker complains that:
adamc@18 750
adamc@18 751 [[
adamc@18 752 The term "(n, s)" has type "(nat * vstack ts)%type"
adamc@207 753 while it is expected to have type "vstack ?119".
adamc@207 754
adamc@207 755 ]]
adamc@207 756
adamc@207 757 The text [?119] stands for a unification variable. We can try to help Coq figure out the value of this variable with an explicit annotation on our [match] expression.
adamc@207 758
adamc@207 759 [[
adamc@207 760 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
adamc@207 761 match i in tinstr ts ts' return vstack ts' with
adamc@207 762 | TINConst _ n => (n, s)
adamc@207 763 | TIBConst _ b => (b, s)
adamc@207 764 | TIBinop _ _ _ _ b =>
adamc@207 765 match s with
adamc@207 766 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@207 767 end
adamc@207 768 end.
adamc@207 769
adamc@207 770 ]]
adamc@207 771
adamc@207 772 Now the error message changes.
adamc@207 773
adamc@207 774 [[
adamc@207 775 The term "(n, s)" has type "(nat * vstack ts)%type"
adamc@207 776 while it is expected to have type "vstack (Nat :: t)".
adamc@207 777
adamc@18 778 ]]
adamc@18 779
adamc@18 780 Recall from our earlier discussion of [match] annotations that we write the annotations to express to the type-checker the relationship between the type indices of the case object and the result type of the [match]. Coq chooses to assign to the wildcard [_] after [TINConst] the name [t], and the type error is telling us that the type checker cannot prove that [t] is the same as [ts]. By moving [s] out of the [match], we lose the ability to express, with [in] and [return] clauses, the relationship between the shared index [ts] of [s] and [i].
adamc@18 781
adamc@18 782 There %\textit{%#<i>#are#</i>#%}% reasonably general ways of getting around this problem without pushing binders inside [match]es. However, the alternatives are significantly more involved, and the technique we use here is almost certainly the best choice, whenever it applies.
adamc@18 783
adamc@18 784 *)
adamc@18 785
adamc@18 786 (** We finish the semantics with a straightforward definition of program denotation. *)
adamc@18 787
adamc@207 788 Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' :=
adamc@207 789 match p with
adamc@14 790 | TNil _ => fun s => s
adamc@14 791 | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
adamc@14 792 end.
adamc@14 793
adamc@14 794
adamc@14 795 (** ** Translation *)
adamc@14 796
adamc@19 797 (** To define our compilation, it is useful to have an auxiliary function for concatenating two stack machine programs. *)
adamc@19 798
adamc@207 799 Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') : tprog ts' ts'' -> tprog ts ts'' :=
adamc@207 800 match p with
adamc@14 801 | TNil _ => fun p' => p'
adamc@14 802 | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
adamc@14 803 end.
adamc@14 804
adamc@19 805 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)
adamc@19 806
adamc@207 807 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) :=
adamc@207 808 match e with
adamc@14 809 | TNConst n => TCons (TINConst _ n) (TNil _)
adamc@14 810 | TBConst b => TCons (TIBConst _ b) (TNil _)
adamc@14 811 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
adamc@14 812 (tconcat (tcompile e1 _) (TCons (TIBinop _ b) (TNil _)))
adamc@14 813 end.
adamc@14 814
adamc@40 815 (** 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 %\textit{%#<i>#implicit argument#</i>#%}% 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.
adamc@19 816
adamc@19 817 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: *)
adamc@19 818
adamc@14 819 Print tcompile.
adamc@19 820 (** [[
adamc@19 821 tcompile =
adamc@19 822 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
adamc@19 823 tprog ts (t :: ts) :=
adamc@19 824 match e in (texp t0) return (tprog ts (t0 :: ts)) with
adamc@19 825 | TNConst n => TCons (TINConst ts n) (TNil (Nat :: ts))
adamc@19 826 | TBConst b => TCons (TIBConst ts b) (TNil (Bool :: ts))
adamc@19 827 | TBinop arg1 arg2 res b e1 e2 =>
adamc@19 828 tconcat (tcompile arg2 e2 ts)
adamc@19 829 (tconcat (tcompile arg1 e1 (arg2 :: ts))
adamc@19 830 (TCons (TIBinop ts b) (TNil (res :: ts))))
adamc@19 831 end
adamc@19 832 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts)
adam@302 833 ]]
adam@302 834 *)
adamc@19 835
adamc@19 836
adamc@19 837 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)
adamc@19 838
adamc@19 839 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
adamc@207 840 (** [= (42, tt) : vstack (Nat :: nil)] *)
adamc@207 841
adamc@19 842 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
adamc@207 843 (** [= (true, tt) : vstack (Bool :: nil)] *)
adamc@207 844
adamc@19 845 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
adamc@207 846 (** [= (28, tt) : vstack (Nat :: nil)] *)
adamc@207 847
adamc@19 848 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
adamc@207 849 (** [= (false, tt) : vstack (Bool :: nil)] *)
adamc@207 850
adamc@19 851 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
adamc@207 852 (** [= (true, tt) : vstack (Bool :: nil)] *)
adamc@19 853
adamc@14 854
adamc@20 855 (** ** Translation Correctness *)
adamc@20 856
adamc@20 857 (** We can state a correctness theorem similar to the last one. *)
adamc@20 858
adamc@207 859 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 860 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@20 861 (* begin hide *)
adamc@20 862 Abort.
adamc@20 863 (* end hide *)
adamc@22 864 (* begin thide *)
adamc@20 865
adamc@20 866 (** Again, we need to strengthen the theorem statement so that the induction will go through. This time, I will develop an alternative approach to this kind of proof, stating the key lemma as: *)
adamc@14 867
adamc@207 868 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 869 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@20 870
adam@292 871 (** While lemma [compile_correct'] quantified over a program that is the %``%#"#continuation#"#%''% 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.
adamc@20 872
adamc@20 873 Let us try to prove this theorem in the same way that we settled on in the last section. *)
adamc@20 874
adamc@14 875 induction e; crush.
adamc@20 876
adamc@20 877 (** We are left with this unproved conclusion:
adamc@20 878
adamc@20 879 [[
adamc@20 880 tprogDenote
adamc@20 881 (tconcat (tcompile e2 ts)
adamc@20 882 (tconcat (tcompile e1 (arg2 :: ts))
adamc@20 883 (TCons (TIBinop ts t) (TNil (res :: ts))))) s =
adamc@20 884 (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
adamc@207 885
adamc@20 886 ]]
adamc@20 887
adamc@20 888 We need an analogue to the [app_ass] theorem that we used to rewrite the goal in the last section. We can abort this proof and prove such a lemma about [tconcat].
adamc@20 889 *)
adamc@207 890
adamc@14 891 Abort.
adamc@14 892
adamc@26 893 Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
adamc@14 894 (s : vstack ts),
adamc@14 895 tprogDenote (tconcat p p') s
adamc@14 896 = tprogDenote p' (tprogDenote p s).
adamc@14 897 induction p; crush.
adamc@14 898 Qed.
adamc@14 899
adamc@20 900 (** This one goes through completely automatically.
adamc@20 901
adamc@26 902 Some code behind the scenes registers [app_ass] for use by [crush]. We must register [tconcat_correct] similarly to get the same effect: *)
adamc@20 903
adamc@26 904 Hint Rewrite tconcat_correct : cpdt.
adamc@14 905
adamc@26 906 (** We ask that the lemma be used for left-to-right rewriting, and we ask for the hint to be added to the hint database called [cpdt], which is the database used by [crush]. Now we are ready to return to [tcompile_correct'], proving it automatically this time. *)
adamc@20 907
adamc@207 908 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 909 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@14 910 induction e; crush.
adamc@14 911 Qed.
adamc@14 912
adamc@20 913 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
adamc@20 914
adamc@26 915 Hint Rewrite tcompile_correct' : cpdt.
adamc@14 916
adamc@207 917 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 918 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@14 919 crush.
adamc@14 920 Qed.
adamc@22 921 (* end thide *)