annotate src/StackMachine.v @ 311:4cb3ba8604bc

Pass through first half of StackMachine, along with some reorganization of the build process
author Adam Chlipala <adam@chlipala.net>
date Mon, 29 Aug 2011 15:31:06 -0400
parents d2cb78f54454
children 495153a41819
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@307 24 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 %\index{Vernacular commands!Require}%[Require Import] #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%Arith%}%#</span></span># [Bool] #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%List%}%#</span></span># [Tactics.] and %\index{Vernacular commands!Set Implicit Arguments}%[Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.] at the start of the file, to match some code hidden in this rendering of the chapter source. 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 some imports of other modules, followed by [Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.], where the latter affects the default behavior of definitions regarding type inference.
adam@307 25 *)
adamc@9 26
adamc@9 27
adamc@20 28 (** * Arithmetic Expressions Over Natural Numbers *)
adamc@2 29
adamc@40 30 (** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *)
adamc@9 31
adamc@20 32 (** ** Source Language *)
adamc@9 33
adam@311 34 (** We begin with the syntax of the source language.%\index{Vernacular commands!Inductive}% *)
adamc@2 35
adamc@4 36 Inductive binop : Set := Plus | Times.
adamc@2 37
adam@311 38 (** 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 %\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 %\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. *)
adamc@9 39
adamc@4 40 Inductive exp : Set :=
adamc@4 41 | Const : nat -> exp
adamc@4 42 | Binop : binop -> exp -> exp -> exp.
adamc@2 43
adamc@9 44 (** 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 45
adam@311 46 A note for readers following along in the PDF version: %\index{coqdoc}%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 47
adamc@9 48 %\medskip%
adamc@9 49
adam@311 50 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}% *)
adamc@9 51
adamc@4 52 Definition binopDenote (b : binop) : nat -> nat -> nat :=
adamc@4 53 match b with
adamc@4 54 | Plus => plus
adamc@4 55 | Times => mult
adamc@4 56 end.
adamc@2 57
adamc@9 58 (** 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 59
adamc@9 60 [[
adamc@9 61 Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
adamc@9 62 match b with
adamc@9 63 | Plus => plus
adamc@9 64 | Times => mult
adamc@9 65 end.
adamc@9 66
adamc@205 67 ]]
adamc@205 68
adamc@9 69 In this example, we could also omit all of the type annotations, arriving at:
adamc@9 70
adamc@9 71 [[
adamc@9 72 Definition binopDenote := fun b =>
adamc@9 73 match b with
adamc@9 74 | Plus => plus
adamc@9 75 | Times => mult
adamc@9 76 end.
adamc@9 77
adamc@205 78 ]]
adamc@205 79
adam@311 80 Languages like Haskell and ML have a convenient %\index{principal typing}\index{type inference}\emph{%#<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 81
adam@311 82 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}}\emph{%#<i>#Calculus of Inductive Constructions (CIC)#</i>#%}~\cite{CIC}%, which is an extension of the older %\index{Calculus of Constructions}\index{CoC|see{Calculus of Constructions}}\emph{%#<i>#Calculus of Constructions (CoC)#</i>#%}~\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}\emph{%#<i>#strong normalization#</i>#%}~\cite{CIC}%, meaning that every program (and, more importantly, every proof term) terminates; and %\index{relative consistency}\emph{%#<i>#relative consistency#</i>#%}~\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.
adamc@9 83
adam@311 84 Coq is actually based on an extension of CIC called %\index{Gallina}\emph{%#<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 85
adam@311 86 Next, there is %\index{Ltac}\emph{%#<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 87
adam@311 88 Finally, commands like [Inductive] and [Definition] are part of %\index{Vernacular commands}\emph{%#<i>#the Vernacular#</i>#%}%, 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 %\emph{%#<i>#trees#</i>#%}% of vernacular commands, thanks to various nested scoping constructs.)
adamc@9 89
adamc@9 90 %\medskip%
adamc@9 91
adam@311 92 We can give a simple definition of the meaning of an expression:%\index{Vernacular commands!Fixpoint}% *)
adamc@9 93
adamc@4 94 Fixpoint expDenote (e : exp) : nat :=
adamc@4 95 match e with
adamc@4 96 | Const n => n
adamc@4 97 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
adamc@4 98 end.
adamc@2 99
adamc@9 100 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *)
adamc@2 101
adam@311 102 (** 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}\emph{%#<i>#reduction strategy#</i>#%}%, or an %``%#"#order of evaluation.#"#%''% Unlike with ML, which hardcodes an %\emph{%#<i>#eager#</i>#%}% reduction strategy, or Haskell, which hardcodes a %\emph{%#<i>#lazy#</i>#%}% 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.
adam@311 103
adam@311 104 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. *)
adamc@16 105
adamc@16 106 Eval simpl in expDenote (Const 42).
adamc@205 107 (** [= 42 : nat] *)
adamc@205 108
adamc@16 109 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)).
adamc@205 110 (** [= 4 : nat] *)
adamc@205 111
adamc@16 112 Eval simpl in expDenote (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adamc@205 113 (** [= 28 : nat] *)
adamc@9 114
adamc@20 115 (** ** Target Language *)
adamc@4 116
adamc@10 117 (** We will compile our source programs onto a simple stack machine, whose syntax is: *)
adamc@2 118
adamc@4 119 Inductive instr : Set :=
adam@311 120 | iConst : nat -> instr
adam@311 121 | iBinop : binop -> instr.
adamc@2 122
adamc@4 123 Definition prog := list instr.
adamc@4 124 Definition stack := list nat.
adamc@2 125
adamc@10 126 (** 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 127
adam@311 128 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}% *)
adamc@10 129
adamc@4 130 Definition instrDenote (i : instr) (s : stack) : option stack :=
adamc@4 131 match i with
adam@311 132 | iConst n => Some (n :: s)
adam@311 133 | iBinop b =>
adamc@4 134 match s with
adamc@4 135 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
adamc@4 136 | _ => None
adamc@4 137 end
adamc@4 138 end.
adamc@2 139
adam@311 140 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *)
adamc@206 141
adamc@206 142 Fixpoint progDenote (p : prog) (s : stack) : option stack :=
adamc@206 143 match p with
adamc@206 144 | nil => Some s
adamc@206 145 | i :: p' =>
adamc@206 146 match instrDenote i s with
adamc@206 147 | None => None
adamc@206 148 | Some s' => progDenote p' s'
adamc@206 149 end
adamc@206 150 end.
adamc@2 151
adamc@4 152
adamc@9 153 (** ** Translation *)
adamc@4 154
adam@311 155 (** Our compiler itself is now unsurprising. The list concatenation operator %\index{Gallina operators!++}%[++] comes from the Coq standard library. *)
adamc@2 156
adamc@4 157 Fixpoint compile (e : exp) : prog :=
adamc@4 158 match e with
adam@311 159 | Const n => iConst n :: nil
adam@311 160 | Binop b e1 e2 => compile e2 ++ compile e1 ++ iBinop b :: nil
adamc@4 161 end.
adamc@2 162
adamc@2 163
adamc@16 164 (** 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 165
adamc@16 166 Eval simpl in compile (Const 42).
adam@311 167 (** [= iConst 42 :: nil : prog] *)
adamc@206 168
adamc@16 169 Eval simpl in compile (Binop Plus (Const 2) (Const 2)).
adam@311 170 (** [= iConst 2 :: iConst 2 :: iBinop Plus :: nil : prog] *)
adamc@206 171
adamc@16 172 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adam@311 173 (** [= iConst 7 :: iConst 2 :: iConst 2 :: iBinop Plus :: iBinop Times :: nil : prog] *)
adamc@16 174
adamc@40 175 (** We can also run our compiled programs and check that they give the right results. *)
adamc@16 176
adamc@16 177 Eval simpl in progDenote (compile (Const 42)) nil.
adamc@206 178 (** [= Some (42 :: nil) : option stack] *)
adamc@206 179
adamc@16 180 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil.
adamc@206 181 (** [= Some (4 :: nil) : option stack] *)
adamc@206 182
adam@311 183 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2))
adam@311 184 (Const 7))) nil.
adamc@206 185 (** [= Some (28 :: nil) : option stack] *)
adamc@16 186
adamc@16 187
adamc@20 188 (** ** Translation Correctness *)
adamc@4 189
adam@311 190 (** 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}% *)
adamc@11 191
adamc@26 192 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 193 (* begin hide *)
adamc@11 194 Abort.
adamc@11 195 (* end hide *)
adamc@22 196 (* begin thide *)
adamc@11 197
adam@311 198 (** 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}\emph{%#<i>#strengthening the induction hypothesis#</i>#%}%. 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}% *)
adamc@2 199
adamc@206 200 Lemma compile_correct' : forall e p s,
adamc@206 201 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
adamc@11 202
adam@311 203 (** After the period in the [Lemma] command, we are in %\index{interactive proof-editing mode}\emph{%#<i>#the interactive proof-editing mode#</i>#%}%. We find ourselves staring at this ominous screen of text:
adamc@11 204
adamc@11 205 [[
adamc@11 206 1 subgoal
adamc@11 207
adamc@11 208 ============================
adamc@15 209 forall (e : exp) (p : list instr) (s : stack),
adamc@15 210 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
adamc@206 211
adamc@11 212 ]]
adamc@11 213
adam@311 214 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.
adamc@11 215
adam@311 216 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.
adamc@11 217
adam@311 218 We manipulate the proof state by running commands called %\index{tactics}\emph{%#<i>#tactics#</i>#%}%. Let us start out by running one of the most important tactics:%\index{tactics!induction}%
adamc@11 219 *)
adamc@11 220
adamc@4 221 induction e.
adamc@2 222
adamc@11 223 (** 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 224
adam@311 225 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@311 226
adamc@11 227 [[
adamc@11 228 n : nat
adamc@11 229 ============================
adamc@11 230 forall (s : stack) (p : list instr),
adamc@11 231 progDenote (compile (Const n) ++ p) s =
adamc@11 232 progDenote p (expDenote (Const n) :: s)
adamc@11 233 ]]
adam@311 234 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adamc@11 235 [[
adamc@11 236 forall (s : stack) (p : list instr),
adamc@11 237 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 238 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 239
adamc@11 240 ]]
adamc@11 241
adam@311 242 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.
adamc@11 243
adam@311 244 We begin the first case with another very common tactic.%\index{tactics!intros}%
adamc@11 245 *)
adamc@11 246
adamc@4 247 intros.
adamc@11 248
adamc@11 249 (** The current subgoal changes to:
adamc@11 250 [[
adamc@11 251
adamc@11 252 n : nat
adamc@11 253 s : stack
adamc@11 254 p : list instr
adamc@11 255 ============================
adamc@11 256 progDenote (compile (Const n) ++ p) s =
adamc@11 257 progDenote p (expDenote (Const n) :: s)
adamc@206 258
adamc@11 259 ]]
adamc@11 260
adamc@11 261 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
adamc@11 262
adam@311 263 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}%
adamc@11 264 *)
adamc@11 265
adam@311 266 (* begin hide *)
adamc@4 267 unfold compile.
adam@311 268 (* end hide *)
adam@311 269 (** %\coqdockw{unfold} \coqdocdefinition{compile}.%#<tt>unfold compile.</tt># *)
adamc@11 270 (** [[
adamc@11 271 n : nat
adamc@11 272 s : stack
adamc@11 273 p : list instr
adamc@11 274 ============================
adam@311 275 progDenote ((iConst n :: nil) ++ p) s =
adamc@11 276 progDenote p (expDenote (Const n) :: s)
adamc@206 277
adam@302 278 ]]
adam@302 279 *)
adamc@11 280
adam@311 281 (* begin hide *)
adamc@4 282 unfold expDenote.
adam@311 283 (* end hide *)
adam@311 284 (** %\coqdockw{unfold} \coqdocdefinition{expDenote}.%#<tt>unfold expDenote.</tt># *)
adamc@11 285 (** [[
adamc@11 286 n : nat
adamc@11 287 s : stack
adamc@11 288 p : list instr
adamc@11 289 ============================
adam@311 290 progDenote ((iConst n :: nil) ++ p) s = progDenote p (n :: s)
adamc@206 291
adamc@11 292 ]]
adamc@11 293
adam@311 294 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}% *)
adamc@11 295
adam@311 296 (* begin hide *)
adamc@11 297 unfold progDenote at 1.
adam@311 298 (* end hide *)
adam@311 299 (** %\coqdockw{unfold} \coqdocdefinition{progDenote} \coqdoctac{at} 1.%#<tt>unfold progDenote at 1.</tt># *)
adamc@11 300 (** [[
adamc@11 301 n : nat
adamc@11 302 s : stack
adamc@11 303 p : list instr
adamc@11 304 ============================
adamc@11 305 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 306 option stack :=
adamc@11 307 match p0 with
adamc@11 308 | nil => Some s0
adamc@11 309 | i :: p' =>
adamc@11 310 match instrDenote i s0 with
adamc@11 311 | Some s' => progDenote p' s'
adamc@11 312 | None => None (A:=stack)
adamc@11 313 end
adam@311 314 end) ((iConst n :: nil) ++ p) s =
adamc@11 315 progDenote p (n :: s)
adamc@206 316
adamc@11 317 ]]
adamc@11 318
adam@311 319 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 %\coqdocconstructor{None} (\coqdocvar{A}:=\coqdocdefinition{stack})%#<tt>None (A:=stack)</tt>#, which has an annotation specifying that the type of the term ought to be [option A]. This is phrased as an explicit instantiation of a named type parameter [A] from the definition of [option].
adam@311 320
adam@311 321 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}%
adamc@11 322 *)
adamc@11 323
adamc@4 324 simpl.
adamc@11 325 (** [[
adamc@11 326 n : nat
adamc@11 327 s : stack
adamc@11 328 p : list instr
adamc@11 329 ============================
adamc@11 330 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 331 option stack :=
adamc@11 332 match p0 with
adamc@11 333 | nil => Some s0
adamc@11 334 | i :: p' =>
adamc@11 335 match instrDenote i s0 with
adamc@11 336 | Some s' => progDenote p' s'
adamc@11 337 | None => None (A:=stack)
adamc@11 338 end
adamc@11 339 end) p (n :: s) = progDenote p (n :: s)
adamc@206 340
adamc@11 341 ]]
adamc@11 342
adam@311 343 Now we can unexpand the definition of [progDenote]:%\index{tactics!fold}%
adamc@11 344 *)
adamc@11 345
adam@311 346 (* begin hide *)
adamc@11 347 fold progDenote.
adam@311 348 (* end hide *)
adam@311 349 (** %\coqdockw{fold} \coqdocdefinition{progDenote}.%#<tt>fold progDenote.</tt># *)
adamc@11 350 (** [[
adamc@11 351 n : nat
adamc@11 352 s : stack
adamc@11 353 p : list instr
adamc@11 354 ============================
adamc@11 355 progDenote p (n :: s) = progDenote p (n :: s)
adamc@206 356
adamc@11 357 ]]
adamc@11 358
adam@311 359 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}%
adamc@11 360 *)
adamc@11 361
adamc@4 362 reflexivity.
adamc@2 363
adamc@11 364 (** On to the second inductive case:
adamc@11 365
adamc@11 366 [[
adamc@11 367 b : binop
adamc@11 368 e1 : exp
adamc@11 369 IHe1 : forall (s : stack) (p : list instr),
adamc@11 370 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 371 e2 : exp
adamc@11 372 IHe2 : forall (s : stack) (p : list instr),
adamc@11 373 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 374 ============================
adamc@11 375 forall (s : stack) (p : list instr),
adamc@11 376 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 377 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 378
adamc@11 379 ]]
adamc@11 380
adam@311 381 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.
adamc@11 382
adam@311 383 We start out the same way as before, introducing new free variables and unfolding and folding the appropriate definitions. The seemingly frivolous [unfold]/%\coqdockw{fold}%#<tt>fold</tt># pairs are actually accomplishing useful work, because [unfold] will sometimes perform easy simplifications. %\index{tactics!intros}\index{tactics!unfold}\index{tactics!fold}% *)
adamc@11 384
adamc@4 385 intros.
adam@311 386 (* begin hide *)
adamc@4 387 unfold compile.
adamc@4 388 fold compile.
adamc@4 389 unfold expDenote.
adamc@4 390 fold expDenote.
adam@311 391 (* end hide *)
adam@311 392 (** %\coqdockw{unfold} \coqdocdefinition{compile}.
adam@311 393 \coqdockw{fold} \coqdocdefinition{compile}.
adam@311 394 \coqdockw{unfold} \coqdocdefinition{expDenote}.
adam@311 395 \coqdockw{fold} \coqdocdefinition{expDenote}.%
adam@311 396 #<tt>unfold compile. fold compile. unfold expDenote. fold expDenote.</tt># *)
adamc@11 397
adamc@44 398 (** 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 399
adamc@11 400 [[
adamc@11 401 b : binop
adamc@11 402 e1 : exp
adamc@11 403 IHe1 : forall (s : stack) (p : list instr),
adamc@11 404 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 405 e2 : exp
adamc@11 406 IHe2 : forall (s : stack) (p : list instr),
adamc@11 407 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 408 s : stack
adamc@11 409 p : list instr
adamc@11 410 ============================
adam@311 411 progDenote ((compile e2 ++ compile e1 ++ iBinop b :: nil) ++ p) s =
adamc@11 412 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 413
adamc@11 414 ]]
adamc@11 415
adam@311 416 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}% *)
adamc@11 417
adam@311 418 Check app_assoc.
adam@311 419
adamc@11 420 (** [[
adam@311 421 app_assoc_reverse
adamc@11 422 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adamc@206 423
adamc@11 424 ]]
adamc@11 425
adam@311 426 If we did not already know the name of the theorem, we could use the %\index{Vernacular commands!SearchRewrite}\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% command to find it, based on a pattern that we would like to rewrite: *)
adam@277 427
adam@311 428 (* begin hide *)
adam@277 429 SearchRewrite ((_ ++ _) ++ _).
adam@311 430 (* end hide *)
adam@311 431 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [((_ ++ _) ++ _).] *)
adam@277 432 (** [[
adam@311 433 app_assoc_reverse:
adam@311 434 forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adam@311 435 ]]
adam@311 436 %\vspace{-.25in}%
adam@311 437 [[
adam@311 438 app_assoc: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n
adam@277 439
adam@277 440 ]]
adam@277 441
adam@311 442 We use [app_assoc_reverse] to perform a rewrite: %\index{tactics!rewrite}% *)
adamc@11 443
adam@311 444 rewrite app_assoc_reverse.
adamc@11 445
adamc@206 446 (** changing the conclusion to:
adamc@11 447
adamc@206 448 [[
adam@311 449 progDenote (compile e2 ++ (compile e1 ++ iBinop b :: nil) ++ p) s =
adamc@11 450 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 451
adamc@11 452 ]]
adamc@11 453
adam@311 454 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}% *)
adamc@11 455
adamc@4 456 rewrite IHe2.
adamc@11 457 (** [[
adam@311 458 progDenote ((compile e1 ++ iBinop b :: nil) ++ p) (expDenote e2 :: s) =
adamc@11 459 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 460
adamc@11 461 ]]
adamc@11 462
adam@311 463 The same process lets us apply the remaining hypothesis.%\index{tactics!rewrite}% *)
adamc@11 464
adam@311 465 rewrite app_assoc_reverse.
adamc@4 466 rewrite IHe1.
adamc@11 467 (** [[
adam@311 468 progDenote ((iBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
adamc@11 469 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 470
adamc@11 471 ]]
adamc@11 472
adam@311 473 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}%
adamc@11 474 *)
adamc@11 475
adam@311 476 (* begin hide *)
adamc@11 477 unfold progDenote at 1.
adamc@4 478 simpl.
adamc@11 479 fold progDenote.
adamc@4 480 reflexivity.
adam@311 481 (* end hide *)
adam@311 482 (** %\coqdockw{unfold} \coqdocdefinition{progDenote} \coqdoctac{at} 1. \coqdockw{simpl}. \coqdockw{fold} \coqdocdefinition{progDenote}. \coqdockw{reflexivity}.%#<tt>unfold progDenote at 1. simpl. fold progDenote. reflexivity.</tt># *)
adamc@11 483
adam@311 484 (** And the proof is completed, as indicated by the message: *)
adamc@11 485
adam@311 486 (** %\coqdockw{Proof} \coqdockw{completed}.%#<tt>Proof completed.</tt># *)
adamc@11 487
adam@311 488 (** 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}%
adamc@11 489 *)
adamc@11 490
adamc@4 491 Abort.
adamc@2 492
adam@311 493 (** %\index{tactics!induction}\index{tactics!crush}% *)
adam@311 494
adamc@26 495 Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 496 progDenote p (expDenote e :: s).
adamc@4 497 induction e; crush.
adamc@4 498 Qed.
adamc@2 499
adam@311 500 (** 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{tacticals!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.
adamc@11 501
adamc@210 502 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 503
adam@311 504 The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it.
adam@311 505
adam@311 506 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}% *)
adamc@11 507
adamc@26 508 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 509 intros.
adamc@11 510 (** [[
adamc@11 511 e : exp
adamc@11 512 ============================
adamc@11 513 progDenote (compile e) nil = Some (expDenote e :: nil)
adamc@206 514
adamc@11 515 ]]
adamc@11 516
adamc@26 517 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 518
adamc@11 519 Check app_nil_end.
adamc@11 520 (** [[
adamc@11 521 app_nil_end
adamc@11 522 : forall (A : Type) (l : list A), l = l ++ nil
adam@302 523 ]]
adam@311 524 %\index{tactics!rewrite}% *)
adamc@11 525
adamc@4 526 rewrite (app_nil_end (compile e)).
adamc@11 527
adamc@11 528 (** 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 529
adamc@11 530 [[
adamc@11 531 e : exp
adamc@11 532 ============================
adamc@11 533 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
adamc@206 534
adamc@11 535 ]]
adamc@11 536
adam@311 537 Now we can apply the lemma.%\index{tactics!rewrite}% *)
adamc@11 538
adamc@26 539 rewrite compile_correct'.
adamc@11 540 (** [[
adamc@11 541 e : exp
adamc@11 542 ============================
adamc@11 543 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
adamc@206 544
adamc@11 545 ]]
adamc@11 546
adam@311 547 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}% *)
adamc@11 548
adamc@4 549 reflexivity.
adamc@4 550 Qed.
adamc@22 551 (* end thide *)
adamc@14 552
adam@311 553 (** This proof can be shortened and made automated, but we leave that as an exercise for the reader. *)
adam@311 554
adamc@14 555
adamc@20 556 (** * Typed Expressions *)
adamc@14 557
adamc@14 558 (** 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 559
adamc@20 560 (** ** Source Language *)
adamc@14 561
adamc@15 562 (** We define a trivial language of types to classify our expressions: *)
adamc@15 563
adamc@14 564 Inductive type : Set := Nat | Bool.
adamc@14 565
adam@277 566 (** 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 567
adam@277 568 Now we define an expanded set of binary operators. *)
adamc@15 569
adamc@14 570 Inductive tbinop : type -> type -> type -> Set :=
adamc@14 571 | TPlus : tbinop Nat Nat Nat
adamc@14 572 | TTimes : tbinop Nat Nat Nat
adamc@14 573 | TEq : forall t, tbinop t t Bool
adamc@14 574 | TLt : tbinop Nat Nat Bool.
adamc@14 575
adam@307 576 (** 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 %\emph{%#<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 577
adamc@15 578 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 579
adam@307 580 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]. %\emph{%#<i>#Generalized algebraic datatypes (GADTs)#</i>#%}% are a popular feature in GHC Haskell and other languages that removes this first restriction.
adamc@15 581
adam@307 582 The second restriction is not lifted by GADTs. In ML and Haskell, indices of types must be types and may not be %\emph{%#<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 583 *)
adamc@15 584
adamc@15 585 (** We can define a similar type family for typed expressions. *)
adamc@15 586
adamc@14 587 Inductive texp : type -> Set :=
adamc@14 588 | TNConst : nat -> texp Nat
adamc@14 589 | TBConst : bool -> texp Bool
adamc@14 590 | TBinop : forall arg1 arg2 res, tbinop arg1 arg2 res -> texp arg1 -> texp arg2 -> texp res.
adamc@14 591
adamc@15 592 (** 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 593
adamc@14 594 Definition typeDenote (t : type) : Set :=
adamc@14 595 match t with
adamc@14 596 | Nat => nat
adamc@14 597 | Bool => bool
adamc@14 598 end.
adamc@14 599
adamc@15 600 (** 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 601
adam@292 602 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 603 *)
adamc@15 604
adam@279 605 Fixpoint lessThan (n1 n2 : nat) : bool :=
adamc@14 606 match n1, n2 with
adamc@14 607 | O, S _ => true
adam@279 608 | S n1', S n2' => lessThan n1' n2'
adamc@14 609 | _, _ => false
adamc@14 610 end.
adamc@14 611
adam@277 612 (** Now we can interpret binary operators, relying on standard-library equality test functions [eqb] and [beq_nat] for booleans and naturals, respectively: *)
adamc@15 613
adamc@14 614 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@14 615 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@207 616 match b in (tbinop arg1 arg2 res)
adamc@207 617 return (typeDenote arg1 -> typeDenote arg2 -> typeDenote res) with
adamc@14 618 | TPlus => plus
adamc@14 619 | TTimes => mult
adam@277 620 | TEq Nat => beq_nat
adam@277 621 | TEq Bool => eqb
adam@279 622 | TLt => lessThan
adamc@14 623 end.
adamc@14 624
adam@307 625 (** 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 %\emph{%#<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 626
adam@307 627 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 %\emph{%#<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 628
adamc@207 629 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 630
adamc@207 631 (* begin hide *)
adamc@207 632 Reset tbinopDenote.
adamc@207 633 (* end hide *)
adamc@207 634 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@207 635 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@207 636 match b with
adamc@207 637 | TPlus => plus
adamc@207 638 | TTimes => mult
adam@277 639 | TEq Nat => beq_nat
adam@277 640 | TEq Bool => eqb
adam@279 641 | TLt => lessThan
adamc@207 642 end.
adamc@207 643
adamc@207 644 (**
adamc@15 645 The same tricks suffice to define an expression denotation function in an unsurprising way:
adamc@15 646 *)
adamc@15 647
adamc@207 648 Fixpoint texpDenote t (e : texp t) : typeDenote t :=
adamc@207 649 match e with
adamc@14 650 | TNConst n => n
adamc@14 651 | TBConst b => b
adamc@14 652 | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
adamc@14 653 end.
adamc@14 654
adamc@17 655 (** We can evaluate a few example programs to convince ourselves that this semantics is correct. *)
adamc@17 656
adamc@17 657 Eval simpl in texpDenote (TNConst 42).
adamc@207 658 (** [= 42 : typeDenote Nat] *)
adamc@207 659
adamc@17 660 Eval simpl in texpDenote (TBConst true).
adamc@207 661 (** [= true : typeDenote Bool] *)
adamc@207 662
adamc@17 663 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
adamc@207 664 (** [= 28 : typeDenote Nat] *)
adamc@207 665
adamc@17 666 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
adamc@207 667 (** [= false : typeDenote Bool] *)
adamc@207 668
adamc@17 669 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
adamc@207 670 (** [= true : typeDenote Bool] *)
adamc@17 671
adamc@14 672
adamc@20 673 (** ** Target Language *)
adamc@14 674
adam@292 675 (** 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 676
adamc@18 677 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 678
adamc@18 679 We start by defining stack types, which classify sets of possible stacks. *)
adamc@18 680
adamc@14 681 Definition tstack := list type.
adamc@14 682
adamc@18 683 (** 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 684
adamc@18 685 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 686
adamc@14 687 Inductive tinstr : tstack -> tstack -> Set :=
adamc@14 688 | TINConst : forall s, nat -> tinstr s (Nat :: s)
adamc@14 689 | TIBConst : forall s, bool -> tinstr s (Bool :: s)
adam@311 690 | TiBinop : forall arg1 arg2 res s,
adamc@14 691 tbinop arg1 arg2 res
adamc@14 692 -> tinstr (arg1 :: arg2 :: s) (res :: s).
adamc@14 693
adamc@18 694 (** 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 695
adamc@14 696 Inductive tprog : tstack -> tstack -> Set :=
adamc@14 697 | TNil : forall s, tprog s s
adamc@14 698 | TCons : forall s1 s2 s3,
adamc@14 699 tinstr s1 s2
adamc@14 700 -> tprog s2 s3
adamc@14 701 -> tprog s1 s3.
adamc@14 702
adamc@18 703 (** 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 704
adamc@14 705 Fixpoint vstack (ts : tstack) : Set :=
adamc@14 706 match ts with
adamc@14 707 | nil => unit
adamc@14 708 | t :: ts' => typeDenote t * vstack ts'
adamc@14 709 end%type.
adamc@14 710
adamc@210 711 (** 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 712
adamc@207 713 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 714
adamc@14 715 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
adamc@207 716 match i with
adamc@14 717 | TINConst _ n => fun s => (n, s)
adamc@14 718 | TIBConst _ b => fun s => (b, s)
adam@311 719 | TiBinop _ _ _ _ b => fun s =>
adamc@14 720 match s with
adamc@14 721 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@14 722 end
adamc@14 723 end.
adamc@14 724
adamc@18 725 (** 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 726
adamc@18 727 [[
adamc@18 728 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
adamc@207 729 match i with
adamc@18 730 | TINConst _ n => (n, s)
adamc@18 731 | TIBConst _ b => (b, s)
adam@311 732 | TiBinop _ _ _ _ b =>
adamc@18 733 match s with
adamc@18 734 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@18 735 end
adamc@18 736 end.
adamc@18 737
adamc@205 738 ]]
adamc@205 739
adamc@18 740 The Coq type-checker complains that:
adamc@18 741
adamc@18 742 [[
adamc@18 743 The term "(n, s)" has type "(nat * vstack ts)%type"
adamc@207 744 while it is expected to have type "vstack ?119".
adamc@207 745
adamc@207 746 ]]
adamc@207 747
adamc@207 748 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 749
adamc@207 750 [[
adamc@207 751 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
adamc@207 752 match i in tinstr ts ts' return vstack ts' with
adamc@207 753 | TINConst _ n => (n, s)
adamc@207 754 | TIBConst _ b => (b, s)
adam@311 755 | TiBinop _ _ _ _ b =>
adamc@207 756 match s with
adamc@207 757 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@207 758 end
adamc@207 759 end.
adamc@207 760
adamc@207 761 ]]
adamc@207 762
adamc@207 763 Now the error message changes.
adamc@207 764
adamc@207 765 [[
adamc@207 766 The term "(n, s)" has type "(nat * vstack ts)%type"
adamc@207 767 while it is expected to have type "vstack (Nat :: t)".
adamc@207 768
adamc@18 769 ]]
adamc@18 770
adamc@18 771 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 772
adam@307 773 There %\emph{%#<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 774
adamc@18 775 *)
adamc@18 776
adamc@18 777 (** We finish the semantics with a straightforward definition of program denotation. *)
adamc@18 778
adamc@207 779 Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' :=
adamc@207 780 match p with
adamc@14 781 | TNil _ => fun s => s
adamc@14 782 | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
adamc@14 783 end.
adamc@14 784
adamc@14 785
adamc@14 786 (** ** Translation *)
adamc@14 787
adamc@19 788 (** To define our compilation, it is useful to have an auxiliary function for concatenating two stack machine programs. *)
adamc@19 789
adamc@207 790 Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') : tprog ts' ts'' -> tprog ts ts'' :=
adamc@207 791 match p with
adamc@14 792 | TNil _ => fun p' => p'
adamc@14 793 | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
adamc@14 794 end.
adamc@14 795
adamc@19 796 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)
adamc@19 797
adamc@207 798 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) :=
adamc@207 799 match e with
adamc@14 800 | TNConst n => TCons (TINConst _ n) (TNil _)
adamc@14 801 | TBConst b => TCons (TIBConst _ b) (TNil _)
adamc@14 802 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
adam@311 803 (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _)))
adamc@14 804 end.
adamc@14 805
adam@307 806 (** 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 %\emph{%#<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 807
adamc@19 808 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 809
adamc@14 810 Print tcompile.
adamc@19 811 (** [[
adamc@19 812 tcompile =
adamc@19 813 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
adamc@19 814 tprog ts (t :: ts) :=
adamc@19 815 match e in (texp t0) return (tprog ts (t0 :: ts)) with
adamc@19 816 | TNConst n => TCons (TINConst ts n) (TNil (Nat :: ts))
adamc@19 817 | TBConst b => TCons (TIBConst ts b) (TNil (Bool :: ts))
adamc@19 818 | TBinop arg1 arg2 res b e1 e2 =>
adamc@19 819 tconcat (tcompile arg2 e2 ts)
adamc@19 820 (tconcat (tcompile arg1 e1 (arg2 :: ts))
adam@311 821 (TCons (TiBinop ts b) (TNil (res :: ts))))
adamc@19 822 end
adamc@19 823 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts)
adam@302 824 ]]
adam@302 825 *)
adamc@19 826
adamc@19 827
adamc@19 828 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)
adamc@19 829
adamc@19 830 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
adamc@207 831 (** [= (42, tt) : vstack (Nat :: nil)] *)
adamc@207 832
adamc@19 833 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
adamc@207 834 (** [= (true, tt) : vstack (Bool :: nil)] *)
adamc@207 835
adamc@19 836 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
adamc@207 837 (** [= (28, tt) : vstack (Nat :: nil)] *)
adamc@207 838
adamc@19 839 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
adamc@207 840 (** [= (false, tt) : vstack (Bool :: nil)] *)
adamc@207 841
adamc@19 842 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
adamc@207 843 (** [= (true, tt) : vstack (Bool :: nil)] *)
adamc@19 844
adamc@14 845
adamc@20 846 (** ** Translation Correctness *)
adamc@20 847
adamc@20 848 (** We can state a correctness theorem similar to the last one. *)
adamc@20 849
adamc@207 850 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 851 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@20 852 (* begin hide *)
adamc@20 853 Abort.
adamc@20 854 (* end hide *)
adamc@22 855 (* begin thide *)
adamc@20 856
adamc@20 857 (** 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 858
adamc@207 859 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 860 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@20 861
adam@292 862 (** 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 863
adamc@20 864 Let us try to prove this theorem in the same way that we settled on in the last section. *)
adamc@20 865
adamc@14 866 induction e; crush.
adamc@20 867
adamc@20 868 (** We are left with this unproved conclusion:
adamc@20 869
adamc@20 870 [[
adamc@20 871 tprogDenote
adamc@20 872 (tconcat (tcompile e2 ts)
adamc@20 873 (tconcat (tcompile e1 (arg2 :: ts))
adam@311 874 (TCons (TiBinop ts t) (TNil (res :: ts))))) s =
adamc@20 875 (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
adamc@207 876
adamc@20 877 ]]
adamc@20 878
adamc@20 879 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 880 *)
adamc@207 881
adamc@14 882 Abort.
adamc@14 883
adamc@26 884 Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
adamc@14 885 (s : vstack ts),
adamc@14 886 tprogDenote (tconcat p p') s
adamc@14 887 = tprogDenote p' (tprogDenote p s).
adamc@14 888 induction p; crush.
adamc@14 889 Qed.
adamc@14 890
adamc@20 891 (** This one goes through completely automatically.
adamc@20 892
adamc@26 893 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 894
adamc@26 895 Hint Rewrite tconcat_correct : cpdt.
adamc@14 896
adamc@26 897 (** 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 898
adamc@207 899 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 900 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@14 901 induction e; crush.
adamc@14 902 Qed.
adamc@14 903
adamc@20 904 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
adamc@20 905
adamc@26 906 Hint Rewrite tcompile_correct' : cpdt.
adamc@14 907
adamc@207 908 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 909 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@14 910 crush.
adamc@14 911 Qed.
adamc@22 912 (* end thide *)