annotate src/StackMachine.v @ 316:2aaff91f5258

Pass through InductiveTypes, through end of reflexive types
author Adam Chlipala <adam@chlipala.net>
date Sun, 11 Sep 2011 17:22:36 -0400
parents d5787b70cf48
children cbeccef45f4e
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@312 11 Require Import Bool Arith List.
adamc@2 12
adam@314 13 Require Import CpdtTactics.
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@314 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
adam@314 25
adam@314 26 %\index{Vernacular commands!Require}%[Require Import Bool] #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%Arith%}%#</span></span># #<span class="inlinecode"><span class="id" type="var">#%\coqdocconstructor{%List%}%#</span></span># [CpdtTactics.]
adam@314 27
adam@314 28 %\noindent{}%and
adam@314 29
adam@314 30 %\index{Vernacular commands!Set Implicit Arguments}%[Set Implicit] #<span class="inlinecode"><span class="id" type="keyword">#%\coqdockw{%Arguments%}%#</span></span>#[.]
adam@314 31
adam@314 32 %\noindent{}%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 33 *)
adamc@9 34
adamc@9 35
adamc@20 36 (** * Arithmetic Expressions Over Natural Numbers *)
adamc@2 37
adamc@40 38 (** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *)
adamc@9 39
adamc@20 40 (** ** Source Language *)
adamc@9 41
adam@311 42 (** We begin with the syntax of the source language.%\index{Vernacular commands!Inductive}% *)
adamc@2 43
adamc@4 44 Inductive binop : Set := Plus | Times.
adamc@2 45
adam@311 46 (** 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 47
adamc@4 48 Inductive exp : Set :=
adamc@4 49 | Const : nat -> exp
adamc@4 50 | Binop : binop -> exp -> exp -> exp.
adamc@2 51
adamc@9 52 (** 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 53
adam@311 54 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 55
adamc@9 56 %\medskip%
adamc@9 57
adam@311 58 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 59
adamc@4 60 Definition binopDenote (b : binop) : nat -> nat -> nat :=
adamc@4 61 match b with
adamc@4 62 | Plus => plus
adamc@4 63 | Times => mult
adamc@4 64 end.
adamc@2 65
adamc@9 66 (** 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 67
adamc@9 68 [[
adamc@9 69 Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
adamc@9 70 match b with
adamc@9 71 | Plus => plus
adamc@9 72 | Times => mult
adamc@9 73 end.
adamc@9 74
adamc@205 75 ]]
adamc@205 76
adamc@9 77 In this example, we could also omit all of the type annotations, arriving at:
adamc@9 78
adamc@9 79 [[
adamc@9 80 Definition binopDenote := fun b =>
adamc@9 81 match b with
adamc@9 82 | Plus => plus
adamc@9 83 | Times => mult
adamc@9 84 end.
adamc@9 85
adamc@205 86 ]]
adamc@205 87
adam@314 88 Languages like Haskell and ML have a convenient %\index{principal types}\index{type inference}\emph{%#<i>#principal types#</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 89
adam@311 90 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 91
adam@311 92 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 93
adam@311 94 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 95
adam@311 96 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 97
adamc@9 98 %\medskip%
adamc@9 99
adam@311 100 We can give a simple definition of the meaning of an expression:%\index{Vernacular commands!Fixpoint}% *)
adamc@9 101
adamc@4 102 Fixpoint expDenote (e : exp) : nat :=
adamc@4 103 match e with
adamc@4 104 | Const n => n
adamc@4 105 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
adamc@4 106 end.
adamc@2 107
adamc@9 108 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *)
adamc@2 109
adam@311 110 (** 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 111
adam@311 112 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 113
adamc@16 114 Eval simpl in expDenote (Const 42).
adamc@205 115 (** [= 42 : nat] *)
adamc@205 116
adamc@16 117 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)).
adamc@205 118 (** [= 4 : nat] *)
adamc@205 119
adamc@16 120 Eval simpl in expDenote (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adamc@205 121 (** [= 28 : nat] *)
adamc@9 122
adamc@20 123 (** ** Target Language *)
adamc@4 124
adamc@10 125 (** We will compile our source programs onto a simple stack machine, whose syntax is: *)
adamc@2 126
adamc@4 127 Inductive instr : Set :=
adam@311 128 | iConst : nat -> instr
adam@311 129 | iBinop : binop -> instr.
adamc@2 130
adamc@4 131 Definition prog := list instr.
adamc@4 132 Definition stack := list nat.
adamc@2 133
adamc@10 134 (** 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 135
adam@311 136 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 137
adamc@4 138 Definition instrDenote (i : instr) (s : stack) : option stack :=
adamc@4 139 match i with
adam@311 140 | iConst n => Some (n :: s)
adam@311 141 | iBinop b =>
adamc@4 142 match s with
adamc@4 143 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
adamc@4 144 | _ => None
adamc@4 145 end
adamc@4 146 end.
adamc@2 147
adam@311 148 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *)
adamc@206 149
adamc@206 150 Fixpoint progDenote (p : prog) (s : stack) : option stack :=
adamc@206 151 match p with
adamc@206 152 | nil => Some s
adamc@206 153 | i :: p' =>
adamc@206 154 match instrDenote i s with
adamc@206 155 | None => None
adamc@206 156 | Some s' => progDenote p' s'
adamc@206 157 end
adamc@206 158 end.
adamc@2 159
adamc@4 160
adamc@9 161 (** ** Translation *)
adamc@4 162
adam@311 163 (** Our compiler itself is now unsurprising. The list concatenation operator %\index{Gallina operators!++}%[++] comes from the Coq standard library. *)
adamc@2 164
adamc@4 165 Fixpoint compile (e : exp) : prog :=
adamc@4 166 match e with
adam@311 167 | Const n => iConst n :: nil
adam@311 168 | Binop b e1 e2 => compile e2 ++ compile e1 ++ iBinop b :: nil
adamc@4 169 end.
adamc@2 170
adamc@2 171
adamc@16 172 (** 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 173
adamc@16 174 Eval simpl in compile (Const 42).
adam@311 175 (** [= iConst 42 :: nil : prog] *)
adamc@206 176
adamc@16 177 Eval simpl in compile (Binop Plus (Const 2) (Const 2)).
adam@311 178 (** [= iConst 2 :: iConst 2 :: iBinop Plus :: nil : prog] *)
adamc@206 179
adamc@16 180 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adam@311 181 (** [= iConst 7 :: iConst 2 :: iConst 2 :: iBinop Plus :: iBinop Times :: nil : prog] *)
adamc@16 182
adamc@40 183 (** We can also run our compiled programs and check that they give the right results. *)
adamc@16 184
adamc@16 185 Eval simpl in progDenote (compile (Const 42)) nil.
adamc@206 186 (** [= Some (42 :: nil) : option stack] *)
adamc@206 187
adamc@16 188 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil.
adamc@206 189 (** [= Some (4 :: nil) : option stack] *)
adamc@206 190
adam@311 191 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2))
adam@311 192 (Const 7))) nil.
adamc@206 193 (** [= Some (28 :: nil) : option stack] *)
adamc@16 194
adamc@16 195
adamc@20 196 (** ** Translation Correctness *)
adamc@4 197
adam@311 198 (** 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 199
adamc@26 200 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 201 (* begin hide *)
adamc@11 202 Abort.
adamc@11 203 (* end hide *)
adamc@22 204 (* begin thide *)
adamc@11 205
adam@311 206 (** 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 207
adamc@206 208 Lemma compile_correct' : forall e p s,
adamc@206 209 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
adamc@11 210
adam@311 211 (** 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 212
adamc@11 213 [[
adamc@11 214 1 subgoal
adamc@11 215
adamc@11 216 ============================
adamc@15 217 forall (e : exp) (p : list instr) (s : stack),
adamc@15 218 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
adamc@206 219
adamc@11 220 ]]
adamc@11 221
adam@311 222 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 223
adam@311 224 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 225
adam@311 226 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 227 *)
adamc@11 228
adamc@4 229 induction e.
adamc@2 230
adamc@11 231 (** 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 232
adam@311 233 %\vspace{.1in} \noindent 2 \coqdockw{subgoals}\vspace{-.1in}%#<tt>2 subgoals</tt>#
adam@311 234
adamc@11 235 [[
adamc@11 236 n : nat
adamc@11 237 ============================
adamc@11 238 forall (s : stack) (p : list instr),
adamc@11 239 progDenote (compile (Const n) ++ p) s =
adamc@11 240 progDenote p (expDenote (Const n) :: s)
adamc@11 241 ]]
adam@311 242 %\noindent \coqdockw{subgoal} 2 \coqdockw{is}:%#<tt>subgoal 2 is</tt>#
adamc@11 243 [[
adamc@11 244 forall (s : stack) (p : list instr),
adamc@11 245 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 246 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 247
adamc@11 248 ]]
adamc@11 249
adam@311 250 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 251
adam@311 252 We begin the first case with another very common tactic.%\index{tactics!intros}%
adamc@11 253 *)
adamc@11 254
adamc@4 255 intros.
adamc@11 256
adamc@11 257 (** The current subgoal changes to:
adamc@11 258 [[
adamc@11 259
adamc@11 260 n : nat
adamc@11 261 s : stack
adamc@11 262 p : list instr
adamc@11 263 ============================
adamc@11 264 progDenote (compile (Const n) ++ p) s =
adamc@11 265 progDenote p (expDenote (Const n) :: s)
adamc@206 266
adamc@11 267 ]]
adamc@11 268
adamc@11 269 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
adamc@11 270
adam@311 271 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 272 *)
adamc@11 273
adam@311 274 (* begin hide *)
adamc@4 275 unfold compile.
adam@311 276 (* end hide *)
adam@311 277 (** %\coqdockw{unfold} \coqdocdefinition{compile}.%#<tt>unfold compile.</tt># *)
adamc@11 278 (** [[
adamc@11 279 n : nat
adamc@11 280 s : stack
adamc@11 281 p : list instr
adamc@11 282 ============================
adam@311 283 progDenote ((iConst n :: nil) ++ p) s =
adamc@11 284 progDenote p (expDenote (Const n) :: s)
adamc@206 285
adam@302 286 ]]
adam@302 287 *)
adamc@11 288
adam@311 289 (* begin hide *)
adamc@4 290 unfold expDenote.
adam@311 291 (* end hide *)
adam@311 292 (** %\coqdockw{unfold} \coqdocdefinition{expDenote}.%#<tt>unfold expDenote.</tt># *)
adamc@11 293 (** [[
adamc@11 294 n : nat
adamc@11 295 s : stack
adamc@11 296 p : list instr
adamc@11 297 ============================
adam@311 298 progDenote ((iConst n :: nil) ++ p) s = progDenote p (n :: s)
adamc@206 299
adamc@11 300 ]]
adamc@11 301
adam@311 302 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 303
adam@311 304 (* begin hide *)
adamc@11 305 unfold progDenote at 1.
adam@311 306 (* end hide *)
adam@311 307 (** %\coqdockw{unfold} \coqdocdefinition{progDenote} \coqdoctac{at} 1.%#<tt>unfold progDenote at 1.</tt># *)
adamc@11 308 (** [[
adamc@11 309 n : nat
adamc@11 310 s : stack
adamc@11 311 p : list instr
adamc@11 312 ============================
adamc@11 313 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 314 option stack :=
adamc@11 315 match p0 with
adamc@11 316 | nil => Some s0
adamc@11 317 | i :: p' =>
adamc@11 318 match instrDenote i s0 with
adamc@11 319 | Some s' => progDenote p' s'
adamc@11 320 | None => None (A:=stack)
adamc@11 321 end
adam@311 322 end) ((iConst n :: nil) ++ p) s =
adamc@11 323 progDenote p (n :: s)
adamc@206 324
adamc@11 325 ]]
adamc@11 326
adam@311 327 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 328
adam@311 329 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 330 *)
adamc@11 331
adamc@4 332 simpl.
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) p (n :: s) = progDenote p (n :: s)
adamc@206 348
adamc@11 349 ]]
adamc@11 350
adam@311 351 Now we can unexpand the definition of [progDenote]:%\index{tactics!fold}%
adamc@11 352 *)
adamc@11 353
adam@311 354 (* begin hide *)
adamc@11 355 fold progDenote.
adam@311 356 (* end hide *)
adam@311 357 (** %\coqdockw{fold} \coqdocdefinition{progDenote}.%#<tt>fold progDenote.</tt># *)
adamc@11 358 (** [[
adamc@11 359 n : nat
adamc@11 360 s : stack
adamc@11 361 p : list instr
adamc@11 362 ============================
adamc@11 363 progDenote p (n :: s) = progDenote p (n :: s)
adamc@206 364
adamc@11 365 ]]
adamc@11 366
adam@311 367 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 368 *)
adamc@11 369
adamc@4 370 reflexivity.
adamc@2 371
adamc@11 372 (** On to the second inductive case:
adamc@11 373
adamc@11 374 [[
adamc@11 375 b : binop
adamc@11 376 e1 : exp
adamc@11 377 IHe1 : forall (s : stack) (p : list instr),
adamc@11 378 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 379 e2 : exp
adamc@11 380 IHe2 : forall (s : stack) (p : list instr),
adamc@11 381 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 382 ============================
adamc@11 383 forall (s : stack) (p : list instr),
adamc@11 384 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 385 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 386
adamc@11 387 ]]
adamc@11 388
adam@311 389 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 390
adam@311 391 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 392
adamc@4 393 intros.
adam@311 394 (* begin hide *)
adamc@4 395 unfold compile.
adamc@4 396 fold compile.
adamc@4 397 unfold expDenote.
adamc@4 398 fold expDenote.
adam@311 399 (* end hide *)
adam@311 400 (** %\coqdockw{unfold} \coqdocdefinition{compile}.
adam@311 401 \coqdockw{fold} \coqdocdefinition{compile}.
adam@311 402 \coqdockw{unfold} \coqdocdefinition{expDenote}.
adam@311 403 \coqdockw{fold} \coqdocdefinition{expDenote}.%
adam@311 404 #<tt>unfold compile. fold compile. unfold expDenote. fold expDenote.</tt># *)
adamc@11 405
adamc@44 406 (** 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 407
adamc@11 408 [[
adamc@11 409 b : binop
adamc@11 410 e1 : exp
adamc@11 411 IHe1 : forall (s : stack) (p : list instr),
adamc@11 412 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 413 e2 : exp
adamc@11 414 IHe2 : forall (s : stack) (p : list instr),
adamc@11 415 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 416 s : stack
adamc@11 417 p : list instr
adamc@11 418 ============================
adam@311 419 progDenote ((compile e2 ++ compile e1 ++ iBinop b :: nil) ++ p) s =
adamc@11 420 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 421
adamc@11 422 ]]
adamc@11 423
adam@311 424 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 425
adam@311 426 Check app_assoc.
adam@311 427
adamc@11 428 (** [[
adam@311 429 app_assoc_reverse
adamc@11 430 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adamc@206 431
adamc@11 432 ]]
adamc@11 433
adam@311 434 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 435
adam@311 436 (* begin hide *)
adam@277 437 SearchRewrite ((_ ++ _) ++ _).
adam@311 438 (* end hide *)
adam@311 439 (** %\coqdockw{%#<tt>#SearchRewrite#</tt>#%}% [((_ ++ _) ++ _).] *)
adam@277 440 (** [[
adam@311 441 app_assoc_reverse:
adam@311 442 forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adam@311 443 ]]
adam@311 444 %\vspace{-.25in}%
adam@311 445 [[
adam@311 446 app_assoc: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n
adam@277 447
adam@277 448 ]]
adam@277 449
adam@311 450 We use [app_assoc_reverse] to perform a rewrite: %\index{tactics!rewrite}% *)
adamc@11 451
adam@311 452 rewrite app_assoc_reverse.
adamc@11 453
adamc@206 454 (** changing the conclusion to:
adamc@11 455
adamc@206 456 [[
adam@311 457 progDenote (compile e2 ++ (compile e1 ++ iBinop b :: nil) ++ p) s =
adamc@11 458 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 459
adamc@11 460 ]]
adamc@11 461
adam@311 462 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 463
adamc@4 464 rewrite IHe2.
adamc@11 465 (** [[
adam@311 466 progDenote ((compile e1 ++ iBinop b :: nil) ++ p) (expDenote e2 :: s) =
adamc@11 467 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 468
adamc@11 469 ]]
adamc@11 470
adam@311 471 The same process lets us apply the remaining hypothesis.%\index{tactics!rewrite}% *)
adamc@11 472
adam@311 473 rewrite app_assoc_reverse.
adamc@4 474 rewrite IHe1.
adamc@11 475 (** [[
adam@311 476 progDenote ((iBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
adamc@11 477 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 478
adamc@11 479 ]]
adamc@11 480
adam@311 481 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 482 *)
adamc@11 483
adam@311 484 (* begin hide *)
adamc@11 485 unfold progDenote at 1.
adamc@4 486 simpl.
adamc@11 487 fold progDenote.
adamc@4 488 reflexivity.
adam@311 489 (* end hide *)
adam@311 490 (** %\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 491
adam@311 492 (** And the proof is completed, as indicated by the message: *)
adamc@11 493
adam@311 494 (** %\coqdockw{Proof} \coqdockw{completed}.%#<tt>Proof completed.</tt># *)
adamc@11 495
adam@311 496 (** 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 497 *)
adamc@11 498
adamc@4 499 Abort.
adamc@2 500
adam@311 501 (** %\index{tactics!induction}\index{tactics!crush}% *)
adam@311 502
adamc@26 503 Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 504 progDenote p (expDenote e :: s).
adamc@4 505 induction e; crush.
adamc@4 506 Qed.
adamc@2 507
adam@311 508 (** 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 509
adamc@210 510 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 511
adam@312 512 The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it. The tactic commands we have written above are an example of a %\emph{%#<i>#proof script#</i>#%}%, or a series of Ltac programs; while [Qed] uses the result of the script to generate a %\emph{%#<i>#proof term#</i>#%}%, a well-typed term of Gallina. To believe that a theorem is true, we only need to trust that the (relatively simple) checker for proof terms is correct; the use of proof scripts is immaterial. Part I of this book will introduce the principles behind encoding all proofs as terms of Gallina.
adam@311 513
adam@311 514 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 515
adamc@26 516 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 517 intros.
adamc@11 518 (** [[
adamc@11 519 e : exp
adamc@11 520 ============================
adamc@11 521 progDenote (compile e) nil = Some (expDenote e :: nil)
adamc@206 522
adamc@11 523 ]]
adamc@11 524
adamc@26 525 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 526
adamc@11 527 Check app_nil_end.
adamc@11 528 (** [[
adamc@11 529 app_nil_end
adamc@11 530 : forall (A : Type) (l : list A), l = l ++ nil
adam@302 531 ]]
adam@311 532 %\index{tactics!rewrite}% *)
adamc@11 533
adamc@4 534 rewrite (app_nil_end (compile e)).
adamc@11 535
adamc@11 536 (** 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 537
adamc@11 538 [[
adamc@11 539 e : exp
adamc@11 540 ============================
adamc@11 541 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
adamc@206 542
adamc@11 543 ]]
adamc@11 544
adam@311 545 Now we can apply the lemma.%\index{tactics!rewrite}% *)
adamc@11 546
adamc@26 547 rewrite compile_correct'.
adamc@11 548 (** [[
adamc@11 549 e : exp
adamc@11 550 ============================
adamc@11 551 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
adamc@206 552
adamc@11 553 ]]
adamc@11 554
adam@311 555 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 556
adamc@4 557 reflexivity.
adamc@4 558 Qed.
adamc@22 559 (* end thide *)
adamc@14 560
adam@311 561 (** This proof can be shortened and made automated, but we leave that as an exercise for the reader. *)
adam@311 562
adamc@14 563
adamc@20 564 (** * Typed Expressions *)
adamc@14 565
adamc@14 566 (** 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 567
adamc@20 568 (** ** Source Language *)
adamc@14 569
adamc@15 570 (** We define a trivial language of types to classify our expressions: *)
adamc@15 571
adamc@14 572 Inductive type : Set := Nat | Bool.
adamc@14 573
adam@277 574 (** 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 575
adam@277 576 Now we define an expanded set of binary operators. *)
adamc@15 577
adamc@14 578 Inductive tbinop : type -> type -> type -> Set :=
adamc@14 579 | TPlus : tbinop Nat Nat Nat
adamc@14 580 | TTimes : tbinop Nat Nat Nat
adamc@14 581 | TEq : forall t, tbinop t t Bool
adamc@14 582 | TLt : tbinop Nat Nat Bool.
adamc@14 583
adam@307 584 (** 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 585
adam@312 586 The inuitive explanation of [tbinop] is that a [tbinop t1 t2 t] is a binary operator whose operands should have types [t1] and [t2], and whose result has type [t]. For instance, constructor [TLt] (for less-than comparison of numbers) is assigned type [tbinop Nat Nat Bool], meaning the operator's arguments are naturals and its result is boolean. The type of [TEq] introduces a small bit of additional complication via polymorphism: we want to allow equality comparison of any two values of any type, as long as they have the %\emph{%#<i>#same#</i>#%}% type.
adam@312 587
adamc@15 588 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 589
adam@312 590 First, the indices of the range of each data constructor must be type variables bound at the top level of the datatype definition. There is no way to do what we did here, where we, for instance, say that [TPlus] is a constructor building a [tbinop] whose indices are all fixed at [Nat]. %\index{generalized algebraic datatypes}\index{GADTs|see{generalized algebraic datatypes}}\emph{%#<i>#Generalized algebraic datatypes (GADTs)#</i>#%}~\cite{GADT}% are a popular feature in %\index{GHC Haskell}%GHC Haskell and other languages that removes this first restriction.
adamc@15 591
adam@312 592 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 %\index{Haskell}%multi-parameter type classes, and recent extensions like type functions bring Haskell programming even closer to %``%#"#real#"#%''% functional programming with types, but, without dependent typing, there must always be a gap between how one programs with types and how one programs normally.
adamc@15 593 *)
adamc@15 594
adam@312 595 (** We can define a similar type family for typed expressions, where a term of type [texp t] can be assigned object language type [t]. (It is conventional in the world of interactive theorem proving to call the language of the proof assistant the %\index{meta language}\emph{%#<i>#meta language#</i>#%}% and a language being formalized the %\index{object language}\emph{%#<i>#object language#</i>#%}%.) *)
adamc@15 596
adamc@14 597 Inductive texp : type -> Set :=
adamc@14 598 | TNConst : nat -> texp Nat
adamc@14 599 | TBConst : bool -> texp Bool
adam@312 600 | TBinop : forall t1 t2 t, tbinop t1 t2 t -> texp t1 -> texp t2 -> texp t.
adamc@14 601
adamc@15 602 (** 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 603
adamc@14 604 Definition typeDenote (t : type) : Set :=
adamc@14 605 match t with
adamc@14 606 | Nat => nat
adamc@14 607 | Bool => bool
adamc@14 608 end.
adamc@14 609
adam@312 610 (** It can take a few moments to come to terms with the fact that [Set], the type of types of programs, is itself a first-class type, and that we can write functions that return [Set]s. Past that wrinkle, the definition of [typeDenote] is trivial, relying on the [nat] and [bool] types from the Coq standard library. We can interpret binary operators by relying on standard-library equality test functions [eqb] and [beq_nat] for booleans and naturals, respectively, along with a less-than test [leb]: *)
adamc@15 611
adamc@207 612 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@207 613 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@207 614 match b with
adamc@207 615 | TPlus => plus
adamc@207 616 | TTimes => mult
adam@277 617 | TEq Nat => beq_nat
adam@277 618 | TEq Bool => eqb
adam@312 619 | TLt => leb
adamc@207 620 end.
adamc@207 621
adam@312 622 (** This function has just a few differences from the denotation functions we saw earlier. First, [tbinop] is an indexed type, so its indices become additional arguments to [tbinopDenote]. Second, we need to perform a genuine %\index{dependent pattern matching}\emph{%#<i>#dependent pattern match#</i>#%}%, where the necessary %\emph{%#<i>#type#</i>#%}% of each case body depends on the %\emph{%#<i>#value#</i>#%}% that has been matched. At this early stage, we will not go into detail on the many subtle aspects of Gallina that support dependent pattern-matching, but the subject is central to Part II of the book.
adam@312 623
adamc@15 624 The same tricks suffice to define an expression denotation function in an unsurprising way:
adamc@15 625 *)
adamc@15 626
adamc@207 627 Fixpoint texpDenote t (e : texp t) : typeDenote t :=
adamc@207 628 match e with
adamc@14 629 | TNConst n => n
adamc@14 630 | TBConst b => b
adamc@14 631 | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
adamc@14 632 end.
adamc@14 633
adamc@17 634 (** We can evaluate a few example programs to convince ourselves that this semantics is correct. *)
adamc@17 635
adamc@17 636 Eval simpl in texpDenote (TNConst 42).
adamc@207 637 (** [= 42 : typeDenote Nat] *)
adamc@207 638
adamc@17 639 Eval simpl in texpDenote (TBConst true).
adamc@207 640 (** [= true : typeDenote Bool] *)
adamc@207 641
adam@312 642 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 643 (TNConst 7)).
adamc@207 644 (** [= 28 : typeDenote Nat] *)
adamc@207 645
adam@312 646 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 647 (TNConst 7)).
adam@312 648 (** [= ] %\coqdocconstructor{%#<tt>#false#</tt>#%}% [ : typeDenote Bool] *)
adamc@207 649
adam@312 650 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 651 (TNConst 7)).
adamc@207 652 (** [= true : typeDenote Bool] *)
adamc@17 653
adamc@14 654
adamc@20 655 (** ** Target Language *)
adamc@14 656
adam@292 657 (** 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 658
adamc@18 659 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 660
adamc@18 661 We start by defining stack types, which classify sets of possible stacks. *)
adamc@18 662
adamc@14 663 Definition tstack := list type.
adamc@14 664
adamc@18 665 (** 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 666
adamc@18 667 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 668
adamc@14 669 Inductive tinstr : tstack -> tstack -> Set :=
adam@312 670 | TiNConst : forall s, nat -> tinstr s (Nat :: s)
adam@312 671 | TiBConst : forall s, bool -> tinstr s (Bool :: s)
adam@311 672 | TiBinop : forall arg1 arg2 res s,
adamc@14 673 tbinop arg1 arg2 res
adamc@14 674 -> tinstr (arg1 :: arg2 :: s) (res :: s).
adamc@14 675
adamc@18 676 (** 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 677
adamc@14 678 Inductive tprog : tstack -> tstack -> Set :=
adamc@14 679 | TNil : forall s, tprog s s
adamc@14 680 | TCons : forall s1 s2 s3,
adamc@14 681 tinstr s1 s2
adamc@14 682 -> tprog s2 s3
adamc@14 683 -> tprog s1 s3.
adamc@14 684
adamc@18 685 (** 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 686
adamc@14 687 Fixpoint vstack (ts : tstack) : Set :=
adamc@14 688 match ts with
adamc@14 689 | nil => unit
adamc@14 690 | t :: ts' => typeDenote t * vstack ts'
adamc@14 691 end%type.
adamc@14 692
adam@312 693 (** This is another [Set]-valued function. This time it is recursive, which is perfectly valid, since [Set] is not treated specially in determining which functions may be written. We say that the value stack of an empty stack type is any value of type [unit], which has just a single value, [tt]. A nonempty stack type leads to a value stack that is a pair, whose first element has the proper type and whose second element follows the representation for the remainder of the stack type. We write [%]%\index{notation scopes}\coqdocvar{%#<tt>#type#</tt>#%}% as an instruction to Coq's extensible parser. In particular, this directive applies to the whole [match] expression, which we ask to be parsed as though it were a type, so that the operator [*] is interpreted as Cartesian product instead of, say, multiplication. (Note that this use of %\coqdocvar{%#<tt>#type#</tt>#%}% has no connection to the inductive type [type] that we have defined.)
adamc@18 694
adam@312 695 This idea of programming with types can take a while to internalize, but it enables a very simple definition of instruction denotation. Our definition is like what you might expect from a Lisp-like version of ML that ignored type information. Nonetheless, the fact that [tinstrDenote] passes the type-checker guarantees that our stack machine programs can never go wrong. We use a special form of [let] to destructure a multi-level tuple. *)
adamc@18 696
adamc@14 697 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
adamc@207 698 match i with
adam@312 699 | TiNConst _ n => fun s => (n, s)
adam@312 700 | TiBConst _ b => fun s => (b, s)
adam@311 701 | TiBinop _ _ _ _ b => fun s =>
adam@312 702 let '(arg1, (arg2, s')) := s in
adam@312 703 ((tbinopDenote b) arg1 arg2, s')
adamc@14 704 end.
adamc@14 705
adamc@18 706 (** 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 707 [[
adamc@18 708 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
adamc@207 709 match i with
adam@312 710 | TiNConst _ n => (n, s)
adam@312 711 | TiBConst _ b => (b, s)
adam@311 712 | TiBinop _ _ _ _ b =>
adam@312 713 let '(arg1, (arg2, s')) := s in
adam@312 714 ((tbinopDenote b) arg1 arg2, s')
adamc@18 715 end.
adamc@18 716
adamc@205 717 ]]
adamc@205 718
adamc@18 719 The Coq type-checker complains that:
adamc@18 720
adam@312 721 <<
adamc@18 722 The term "(n, s)" has type "(nat * vstack ts)%type"
adamc@207 723 while it is expected to have type "vstack ?119".
adam@312 724 >>
adamc@207 725
adam@312 726 This and other mysteries of Coq dependent typing we postpone until Part II of the book. The upshot of our later discussion is that it is often useful to push inside of [match] branches those function parameters whose types depend on the type of the value being matched. Our later, more complete treatement of Gallina's typing rules will explain why this helps.
adamc@18 727 *)
adamc@18 728
adamc@18 729 (** We finish the semantics with a straightforward definition of program denotation. *)
adamc@18 730
adamc@207 731 Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' :=
adamc@207 732 match p with
adamc@14 733 | TNil _ => fun s => s
adamc@14 734 | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
adamc@14 735 end.
adamc@14 736
adamc@14 737
adamc@14 738 (** ** Translation *)
adamc@14 739
adamc@19 740 (** To define our compilation, it is useful to have an auxiliary function for concatenating two stack machine programs. *)
adamc@19 741
adamc@207 742 Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') : tprog ts' ts'' -> tprog ts ts'' :=
adamc@207 743 match p with
adamc@14 744 | TNil _ => fun p' => p'
adamc@14 745 | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
adamc@14 746 end.
adamc@14 747
adamc@19 748 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)
adamc@19 749
adamc@207 750 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) :=
adamc@207 751 match e with
adam@312 752 | TNConst n => TCons (TiNConst _ n) (TNil _)
adam@312 753 | TBConst b => TCons (TiBConst _ b) (TNil _)
adamc@14 754 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
adam@311 755 (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _)))
adamc@14 756 end.
adamc@14 757
adam@307 758 (** 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 759
adamc@19 760 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 761
adamc@14 762 Print tcompile.
adamc@19 763 (** [[
adamc@19 764 tcompile =
adamc@19 765 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
adamc@19 766 tprog ts (t :: ts) :=
adamc@19 767 match e in (texp t0) return (tprog ts (t0 :: ts)) with
adam@312 768 | TNConst n => TCons (TiNConst ts n) (TNil (Nat :: ts))
adam@312 769 | TBConst b => TCons (TiBConst ts b) (TNil (Bool :: ts))
adamc@19 770 | TBinop arg1 arg2 res b e1 e2 =>
adamc@19 771 tconcat (tcompile arg2 e2 ts)
adamc@19 772 (tconcat (tcompile arg1 e1 (arg2 :: ts))
adam@311 773 (TCons (TiBinop ts b) (TNil (res :: ts))))
adamc@19 774 end
adamc@19 775 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts)
adam@302 776 ]]
adam@302 777 *)
adamc@19 778
adamc@19 779
adamc@19 780 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)
adamc@19 781
adamc@19 782 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
adam@312 783 (** [= (42, tt) : vstack (][Nat :: nil)] *)
adamc@207 784
adamc@19 785 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
adam@312 786 (** [= (][true][, tt) : vstack (][Bool :: nil)] *)
adamc@207 787
adam@312 788 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2)
adam@312 789 (TNConst 2)) (TNConst 7)) nil) tt.
adam@312 790 (** [= (28, tt) : vstack (][Nat :: nil)] *)
adamc@207 791
adam@312 792 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2)
adam@312 793 (TNConst 2)) (TNConst 7)) nil) tt.
adam@312 794 (** [= (]%\coqdocconstructor{%#<tt>#false#</tt>#%}%[, tt) : vstack (][Bool :: nil)] *)
adamc@207 795
adam@312 796 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 797 (TNConst 7)) nil) tt.
adam@312 798 (** [= (][true][, tt) : vstack (][Bool :: nil)] *)
adamc@19 799
adamc@14 800
adamc@20 801 (** ** Translation Correctness *)
adamc@20 802
adamc@20 803 (** We can state a correctness theorem similar to the last one. *)
adamc@20 804
adamc@207 805 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 806 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@20 807 (* begin hide *)
adamc@20 808 Abort.
adamc@20 809 (* end hide *)
adamc@22 810 (* begin thide *)
adamc@20 811
adam@312 812 (** Again, we need to strengthen the theorem statement so that the induction will go through. This time, to provide an excuse to demonstrate different tactics, I will develop an alternative approach to this kind of proof, stating the key lemma as: *)
adamc@14 813
adamc@207 814 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 815 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@20 816
adam@292 817 (** 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 818
adamc@20 819 Let us try to prove this theorem in the same way that we settled on in the last section. *)
adamc@20 820
adamc@14 821 induction e; crush.
adamc@20 822
adamc@20 823 (** We are left with this unproved conclusion:
adamc@20 824
adamc@20 825 [[
adamc@20 826 tprogDenote
adamc@20 827 (tconcat (tcompile e2 ts)
adamc@20 828 (tconcat (tcompile e1 (arg2 :: ts))
adam@311 829 (TCons (TiBinop ts t) (TNil (res :: ts))))) s =
adamc@20 830 (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
adamc@207 831
adamc@20 832 ]]
adamc@20 833
adam@312 834 We need an analogue to the [app_assoc_reverse] theorem that we used to rewrite the goal in the last section. We can abort this proof and prove such a lemma about [tconcat].
adamc@20 835 *)
adamc@207 836
adamc@14 837 Abort.
adamc@14 838
adamc@26 839 Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
adamc@14 840 (s : vstack ts),
adamc@14 841 tprogDenote (tconcat p p') s
adamc@14 842 = tprogDenote p' (tprogDenote p s).
adamc@14 843 induction p; crush.
adamc@14 844 Qed.
adamc@14 845
adamc@20 846 (** This one goes through completely automatically.
adamc@20 847
adam@316 848 Some code behind the scenes registers [app_assoc_reverse] for use by [crush]. We must register [tconcat_correct] similarly to get the same effect:%\index{Vernacular commands!Hint Rewrite}% *)
adamc@20 849
adam@312 850 (* begin hide *)
adamc@26 851 Hint Rewrite tconcat_correct : cpdt.
adam@312 852 (* end hide *)
adam@312 853 (** %\noindent%[Hint] %\coqdockw{%#<tt>#Rewrite#</tt>#%}% [tconcat_correct : cpdt.] *)
adamc@14 854
adam@312 855 (** Here we meet the pervasive concept of a %\emph{%#<i>#hint#</i>#%}%. Many proofs can be found through exhaustive enumerations of combinations of possible proof steps; hints provide the set of steps to consider. The tactic [crush] is applying such brute force search for us silently, and it will consider more possibilities as we add more hints. This particular hint asks that the lemma be used for left-to-right rewriting, and we ask for the hint to be added to the hint database called [cpdt], which is the database used by [crush]. In general, segragating hints into different databases is helpful to control the performance of proof search, in cases where domain knowledge allows us to narrow the set of proof steps to be considered in brute force search. Part III of this book considers such pragmatic aspects of proof search in much more detail.
adam@312 856
adam@312 857 Now we are ready to return to [tcompile_correct'], proving it automatically this time. *)
adamc@20 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@14 861 induction e; crush.
adamc@14 862 Qed.
adamc@14 863
adamc@20 864 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
adamc@20 865
adam@312 866 (* begin hide *)
adamc@26 867 Hint Rewrite tcompile_correct' : cpdt.
adam@312 868 (* end hide *)
adam@312 869 (** %\noindent%[Hint ]%\coqdockw{%#<tt>#Rewrite#</tt>#%}%[ tcompile_correct' : cpdt.] *)
adamc@14 870
adamc@207 871 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 872 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@14 873 crush.
adamc@14 874 Qed.
adamc@22 875 (* end thide *)
adam@312 876
adam@312 877 (** It is probably worth emphasizing that we are doing more than building mathematical models. Our compilers are functional programs that can be executed efficiently. One strategy for doing so is based on %\index{program extraction}\emph{%#<i>#program extraction#</i>#%}%, which generates OCaml code from Coq developments. For instance, we run a command to output the OCaml version of [tcompile]:%\index{Vernacular commands!Extraction}% *)
adam@312 878
adam@312 879 (* begin hide *)
adam@312 880 Extraction tcompile.
adam@312 881 (* end hide *)
adam@312 882 (** %\noindent\coqdockw{%#<tt>#Extraction#</tt>#%}%[ tcompile.] *)
adam@312 883
adam@312 884 (** <<
adam@312 885 let rec tcompile t e ts =
adam@312 886 match e with
adam@312 887 | TNConst n ->
adam@312 888 TCons (ts, (Cons (Nat, ts)), (Cons (Nat, ts)), (TiNConst (ts, n)), (TNil
adam@312 889 (Cons (Nat, ts))))
adam@312 890 | TBConst b ->
adam@312 891 TCons (ts, (Cons (Bool, ts)), (Cons (Bool, ts)), (TiBConst (ts, b)),
adam@312 892 (TNil (Cons (Bool, ts))))
adam@312 893 | TBinop (t1, t2, t0, b, e1, e2) ->
adam@312 894 tconcat ts (Cons (t2, ts)) (Cons (t0, ts)) (tcompile t2 e2 ts)
adam@312 895 (tconcat (Cons (t2, ts)) (Cons (t1, (Cons (t2, ts)))) (Cons (t0, ts))
adam@312 896 (tcompile t1 e1 (Cons (t2, ts))) (TCons ((Cons (t1, (Cons (t2,
adam@312 897 ts)))), (Cons (t0, ts)), (Cons (t0, ts)), (TiBinop (t1, t2, t0, ts,
adam@312 898 b)), (TNil (Cons (t0, ts))))))
adam@312 899 >>
adam@312 900
adam@312 901 We can compile this code with the usual OCaml compiler and obtain an executable program with halfway decent performance.
adam@312 902
adam@312 903 This chapter has been a whirlwind tour through two examples of the style of Coq development that I advocate. Parts II and III of the book focus on the key elements of that style, namely dependent types and scripted proof automation, respectively. Before we get there, we will spend some time in Part I on more standard foundational material. Part I may still be of interest to seasoned Coq hackers, since I follow the highly automated proof style even at that early stage. *)