annotate src/StackMachine.v @ 541:429e95d23b26

Typo fix
author Adam Chlipala <adam@chlipala.net>
date Sat, 19 Sep 2015 07:44:42 -0400
parents d65e9c1c9041
children af97676583f3
rev   line source
adam@534 1 (* Copyright (c) 2008-2012, 2015, 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@2 10
adamc@25 11 (** %\chapter{Some Quick Examples}% *)
adamc@25 12
adam@447 13 (** I will start off by jumping right in to a fully worked set of examples, building certified compilers from increasingly complicated source languages to stack machines. We will meet a few useful tactics and see how they can be used in manual proofs, and we will also see how easily these proofs can be automated instead. This chapter is not meant to give full explanations of the features that are employed. Rather, it is meant more as an advertisement of what is possible. Later chapters will introduce all of the concepts in bottom-up fashion. In other words, it is expected that most readers will not understand what exactly is going on here, but I hope this demo will whet your appetite for the remaining chapters!
adam@279 14
adam@536 15 As always, you can step through the source file <<StackMachine.v>> for this chapter interactively in Proof General. Alternatively, to get a feel for the whole lifecycle of creating a Coq development, you can enter the pieces of source code in this chapter in a new <<.v>> file in an Emacs buffer. If you do the latter, include these three lines at the start of the file. *)
adam@314 16
adam@536 17 Require Import Bool Arith List Cpdt.CpdtTactics.
adam@419 18 Set Implicit Arguments.
adam@534 19 Set Asymmetric Patterns.
adam@314 20
adam@513 21 (* begin hide *)
adam@513 22 (* begin thide *)
adam@513 23 Definition bleh := app_assoc.
adam@513 24 (* end thide *)
adam@513 25 (* end hide *)
adam@513 26
adam@534 27 (** In general, similar commands will be hidden in the book rendering of each chapter's source code, so you will need to insert them in from-scratch replayings of the code that is presented. To be more specific, every chapter begins with the above three lines, with the import list tweaked as appropriate, considering which definitions the chapter uses. The second command above affects the default behavior of definitions regarding type inference, and the third allows for more concise pattern-matching syntax in Coq versions 8.5 and higher (having no effect in earlier versions). *)
adamc@9 28
adamc@9 29
adamc@20 30 (** * Arithmetic Expressions Over Natural Numbers *)
adamc@2 31
adamc@40 32 (** We will begin with that staple of compiler textbooks, arithmetic expressions over a single type of numbers. *)
adamc@9 33
adamc@20 34 (** ** Source Language *)
adamc@9 35
adam@311 36 (** We begin with the syntax of the source language.%\index{Vernacular commands!Inductive}% *)
adamc@2 37
adamc@4 38 Inductive binop : Set := Plus | Times.
adamc@2 39
adam@447 40 (** Our first line of Coq code should be unsurprising to ML and Haskell programmers. We define an %\index{algebraic datatypes}%algebraic datatype [binop] to stand for the binary operators of our source language. There are just two wrinkles compared to ML and Haskell. First, we use the keyword [Inductive], in place of <<data>>, <<datatype>>, or <<type>>. This is not just a trivial surface syntax difference; inductive types in Coq are much more expressive than garden variety algebraic datatypes, essentially enabling us to encode all of mathematics, though we begin humbly in this chapter. Second, there is the %\index{Gallina terms!Set}%[: Set] fragment, which declares that we are defining a datatype that should be thought of as a constituent of programs. Later, we will see other options for defining datatypes in the universe of proofs or in an infinite hierarchy of universes, encompassing both programs and proofs, that is useful in higher-order constructions. *)
adamc@9 41
adamc@4 42 Inductive exp : Set :=
adamc@4 43 | Const : nat -> exp
adamc@4 44 | Binop : binop -> exp -> exp -> exp.
adamc@2 45
adamc@9 46 (** 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 47
adam@419 48 A note for readers following along in the PDF version: %\index{coqdoc}%coqdoc supports pretty-printing of tokens in %\LaTeX{}%#LaTeX# or HTML. Where you see a right arrow character, the source contains the ASCII text <<->>>. Other examples of this substitution appearing in this chapter are a double right arrow for <<=>>>, the inverted %`%#'#A' symbol for <<forall>>, and the Cartesian product %`%#'#X' for <<*>>. When in doubt about the ASCII version of a symbol, you can consult the chapter source code.
adamc@9 49
adamc@9 50 %\medskip%
adamc@9 51
adam@475 52 Now we are ready to say what programs in our expression language 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 53
adamc@4 54 Definition binopDenote (b : binop) : nat -> nat -> nat :=
adamc@4 55 match b with
adamc@4 56 | Plus => plus
adamc@4 57 | Times => mult
adamc@4 58 end.
adamc@2 59
adam@419 60 (** The meaning of a binary operator is a binary function over naturals, defined with pattern-matching notation analogous to the <<case>> and <<match>> of ML and Haskell, and referring to the functions [plus] and [mult] from the Coq standard library. The keyword [Definition] is Coq's all-purpose notation for binding a term of the programming language to a name, with some associated syntactic sugar, like the notation we see here for defining a function. That sugar could be expanded to yield this definition:
adamc@9 61 [[
adamc@9 62 Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
adamc@9 63 match b with
adamc@9 64 | Plus => plus
adamc@9 65 | Times => mult
adamc@9 66 end.
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 Definition binopDenote := fun b =>
adamc@9 72 match b with
adamc@9 73 | Plus => plus
adamc@9 74 | Times => mult
adamc@9 75 end.
adamc@205 76 ]]
adamc@205 77
adam@475 78 Languages like Haskell and ML have a convenient%\index{principal types}\index{type inference}% _principal types_ property, which gives us strong guarantees about how effective type inference will be. Unfortunately, Coq's type system is so expressive that any kind of "complete" type inference is impossible, and the task even seems to be hard 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 79
adam@475 80 This is as good a time as any to mention the profusion of different languages associated with Coq. The theoretical foundation of Coq is a formal system called the%\index{Calculus of Inductive Constructions}\index{CIC|see{Calculus of Inductive Constructions}}% _Calculus of Inductive Constructions_ (CIC)%~\cite{CIC}%, which is an extension of the older%\index{Calculus of Constructions}\index{CoC|see{Calculus of Constructions}}% _Calculus of Constructions_ (CoC)%~\cite{CoC}%. CIC is quite a spartan foundation, which is helpful for proving metatheory but not so helpful for real development. Still, it is nice to know that it has been proved that CIC enjoys properties like%\index{strong normalization}% _strong normalization_ %\cite{CIC}%, meaning that every program (and, more importantly, every proof term) terminates; and%\index{relative consistency}% _relative consistency_ %\cite{SetsInTypes}% with systems like versions of %\index{Zermelo-Fraenkel set theory}%Zermelo-Fraenkel set theory, which roughly means that you can believe that Coq proofs mean that the corresponding propositions are "really true," if you believe in set theory.
adamc@9 81
adam@475 82 Coq is actually based on an extension of CIC called %\index{Gallina}%Gallina. The text after the [:=] and before the period in the last code example is a term of Gallina. Gallina includes several useful features that must be considered as extensions to CIC. The important metatheorems about CIC have not been extended to the full breadth of the features that go beyond the formalized language, but most Coq users do not seem to lose much sleep over this omission.
adamc@9 83
adam@475 84 Next, there is %\index{Ltac}%Ltac, Coq's domain-specific language for writing proofs and decision procedures. We will see some basic examples of Ltac later in this chapter, and much of this book is devoted to more involved Ltac examples.
adamc@9 85
adam@475 86 Finally, commands like [Inductive] and [Definition] are part of %\index{Vernacular commands}%the Vernacular, which includes all sorts of useful queries and requests to the Coq system. Every Coq source file is a series of vernacular commands, where many command forms take arguments that are Gallina or Ltac programs. (Actually, Coq source files are more like _trees_ of vernacular commands, thanks to various nested scoping constructs.)
adamc@9 87
adamc@9 88 %\medskip%
adamc@9 89
adam@311 90 We can give a simple definition of the meaning of an expression:%\index{Vernacular commands!Fixpoint}% *)
adamc@9 91
adamc@4 92 Fixpoint expDenote (e : exp) : nat :=
adamc@4 93 match e with
adamc@4 94 | Const n => n
adamc@4 95 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
adamc@4 96 end.
adamc@2 97
adamc@9 98 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *)
adamc@2 99
adam@419 100 (** It is convenient to be able to test definitions before starting to prove things about them. We can verify that our semantics is sensible by evaluating some sample uses, using the command %\index{Vernacular commands!Eval}%[Eval]. This command takes an argument expressing a%\index{reduction strategy}% _reduction strategy_, or an "order of evaluation." Unlike with ML, which hardcodes an _eager_ reduction strategy, or Haskell, which hardcodes a _lazy_ strategy, in Coq we are free to choose between these and many other orders of evaluation, because all Coq programs terminate. In fact, Coq silently checked %\index{termination checking}%termination of our [Fixpoint] definition above, using a simple heuristic based on monotonically decreasing size of arguments across recursive calls. Specifically, recursive calls must be made on arguments that were pulled out of the original recursive argument with [match] expressions. (In Chapter 7, we will see some ways of getting around this restriction, though simply removing the restriction would leave Coq useless as a theorem proving tool, for reasons we will start to learn about in the next chapter.)
adam@311 101
adam@311 102 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 103
adamc@16 104 Eval simpl in expDenote (Const 42).
adamc@205 105 (** [= 42 : nat] *)
adamc@205 106
adamc@16 107 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)).
adamc@205 108 (** [= 4 : nat] *)
adamc@205 109
adamc@16 110 Eval simpl in expDenote (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adamc@205 111 (** [= 28 : nat] *)
adamc@9 112
adam@442 113 (** %\smallskip{}%Nothing too surprising goes on here, so we are ready to move on to the target language of our compiler. *)
adam@442 114
adam@442 115
adamc@20 116 (** ** Target Language *)
adamc@4 117
adamc@10 118 (** We will compile our source programs onto a simple stack machine, whose syntax is: *)
adamc@2 119
adamc@4 120 Inductive instr : Set :=
adam@311 121 | iConst : nat -> instr
adam@311 122 | iBinop : binop -> instr.
adamc@2 123
adamc@4 124 Definition prog := list instr.
adamc@4 125 Definition stack := list nat.
adamc@2 126
adamc@10 127 (** 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 128
adam@419 129 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 130
adamc@4 131 Definition instrDenote (i : instr) (s : stack) : option stack :=
adamc@4 132 match i with
adam@311 133 | iConst n => Some (n :: s)
adam@311 134 | iBinop b =>
adamc@4 135 match s with
adamc@4 136 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
adamc@4 137 | _ => None
adamc@4 138 end
adamc@4 139 end.
adamc@2 140
adam@311 141 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *)
adamc@206 142
adamc@206 143 Fixpoint progDenote (p : prog) (s : stack) : option stack :=
adamc@206 144 match p with
adamc@206 145 | nil => Some s
adamc@206 146 | i :: p' =>
adamc@206 147 match instrDenote i s with
adamc@206 148 | None => None
adamc@206 149 | Some s' => progDenote p' s'
adamc@206 150 end
adamc@206 151 end.
adamc@2 152
adam@442 153 (** With the two programming languages defined, we can turn to the compiler definition. *)
adam@442 154
adamc@4 155
adamc@9 156 (** ** Translation *)
adamc@4 157
adam@471 158 (** Our compiler itself is now unsurprising. The list concatenation operator %\index{Gallina operators!++}\coqdocnotation{%#<tt>#++#</tt>#%}% comes from the Coq standard library. *)
adamc@2 159
adamc@4 160 Fixpoint compile (e : exp) : prog :=
adamc@4 161 match e with
adam@311 162 | Const n => iConst n :: nil
adam@311 163 | Binop b e1 e2 => compile e2 ++ compile e1 ++ iBinop b :: nil
adamc@4 164 end.
adamc@2 165
adamc@2 166
adamc@16 167 (** 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 168
adamc@16 169 Eval simpl in compile (Const 42).
adam@311 170 (** [= iConst 42 :: nil : prog] *)
adamc@206 171
adamc@16 172 Eval simpl in compile (Binop Plus (Const 2) (Const 2)).
adam@311 173 (** [= iConst 2 :: iConst 2 :: iBinop Plus :: nil : prog] *)
adamc@206 174
adamc@16 175 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
adam@311 176 (** [= iConst 7 :: iConst 2 :: iConst 2 :: iBinop Plus :: iBinop Times :: nil : prog] *)
adamc@16 177
adam@442 178 (** %\smallskip{}%We can also run our compiled programs and check that they give the right results. *)
adamc@16 179
adamc@16 180 Eval simpl in progDenote (compile (Const 42)) nil.
adamc@206 181 (** [= Some (42 :: nil) : option stack] *)
adamc@206 182
adamc@16 183 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil.
adamc@206 184 (** [= Some (4 :: nil) : option stack] *)
adamc@206 185
adam@311 186 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2))
adam@311 187 (Const 7))) nil.
adamc@206 188 (** [= Some (28 :: nil) : option stack] *)
adamc@16 189
adam@447 190 (** %\smallskip{}%So far so good, but how can we be sure the compiler operates correctly for _all_ input programs? *)
adamc@16 191
adamc@20 192 (** ** Translation Correctness *)
adamc@4 193
adam@311 194 (** 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 195
adamc@26 196 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@22 197 (* begin thide *)
adamc@11 198
adam@419 199 (** Though a pencil-and-paper proof might clock out at this point, writing "by a routine induction on [e]," it turns out not to make sense to attack this proof directly. We need to use the standard trick of%\index{strengthening the induction hypothesis}% _strengthening the induction hypothesis_. We do that by proving an auxiliary lemma, using the command [Lemma] that is a synonym for [Theorem], conventionally used for less important theorems that appear in the proofs of primary theorems.%\index{Vernacular commands!Lemma}% *)
adamc@2 200
adam@469 201 Abort.
adam@469 202
adamc@206 203 Lemma compile_correct' : forall e p s,
adamc@206 204 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
adamc@11 205
adam@399 206 (** After the period in the [Lemma] command, we are in%\index{interactive proof-editing mode}% _the interactive proof-editing mode_. We find ourselves staring at this ominous screen of text:
adamc@11 207
adamc@11 208 [[
adamc@11 209 1 subgoal
adamc@11 210
adamc@11 211 ============================
adamc@15 212 forall (e : exp) (p : list instr) (s : stack),
adamc@15 213 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
adamc@206 214
adamc@11 215 ]]
adamc@11 216
adam@311 217 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 218
adam@311 219 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 220
adam@399 221 We manipulate the proof state by running commands called%\index{tactics}% _tactics_. Let us start out by running one of the most important tactics:%\index{tactics!induction}%
adamc@11 222 *)
adamc@11 223
adamc@4 224 induction e.
adamc@2 225
adamc@11 226 (** 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 227
adam@439 228 [[
adam@439 229 2 subgoals
adam@311 230
adamc@11 231 n : nat
adamc@11 232 ============================
adamc@11 233 forall (s : stack) (p : list instr),
adamc@11 234 progDenote (compile (Const n) ++ p) s =
adamc@11 235 progDenote p (expDenote (Const n) :: s)
adam@439 236
adam@439 237 subgoal 2 is
adam@439 238
adamc@11 239 forall (s : stack) (p : list instr),
adamc@11 240 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 241 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 242
adamc@11 243 ]]
adamc@11 244
adam@311 245 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 246
adam@311 247 We begin the first case with another very common tactic.%\index{tactics!intros}%
adamc@11 248 *)
adamc@11 249
adamc@4 250 intros.
adamc@11 251
adamc@11 252 (** The current subgoal changes to:
adamc@11 253 [[
adamc@11 254
adamc@11 255 n : nat
adamc@11 256 s : stack
adamc@11 257 p : list instr
adamc@11 258 ============================
adamc@11 259 progDenote (compile (Const n) ++ p) s =
adamc@11 260 progDenote p (expDenote (Const n) :: s)
adamc@206 261
adamc@11 262 ]]
adamc@11 263
adamc@11 264 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
adamc@11 265
adam@311 266 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 267 *)
adamc@11 268
adamc@4 269 unfold compile.
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
adamc@4 281 unfold expDenote.
adamc@11 282 (** [[
adamc@11 283 n : nat
adamc@11 284 s : stack
adamc@11 285 p : list instr
adamc@11 286 ============================
adam@311 287 progDenote ((iConst n :: nil) ++ p) s = progDenote p (n :: s)
adamc@206 288
adamc@11 289 ]]
adamc@11 290
adam@311 291 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 292
adamc@11 293 unfold progDenote at 1.
adamc@11 294 (** [[
adamc@11 295 n : nat
adamc@11 296 s : stack
adamc@11 297 p : list instr
adamc@11 298 ============================
adamc@11 299 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 300 option stack :=
adamc@11 301 match p0 with
adamc@11 302 | nil => Some s0
adamc@11 303 | i :: p' =>
adamc@11 304 match instrDenote i s0 with
adamc@11 305 | Some s' => progDenote p' s'
adamc@11 306 | None => None (A:=stack)
adamc@11 307 end
adam@311 308 end) ((iConst n :: nil) ++ p) s =
adamc@11 309 progDenote p (n :: s)
adamc@206 310
adamc@11 311 ]]
adamc@11 312
adam@471 313 This last [unfold] has left us with an anonymous recursive definition of [progDenote] (similarly to how [fun] or "lambda" constructs in general allow anonymous non-recursive functions), which will generally happen when unfolding recursive definitions. Note that Coq has automatically renamed the [fix] arguments [p] and [s] to [p0] and [s0], to avoid clashes with our local free variables. There is also a subterm [None (A:=stack)], which has an annotation specifying that the type of the term ought to be [option stack]. This is phrased as an explicit instantiation of a named type parameter [A] from the definition of [option].
adam@311 314
adam@311 315 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 316 *)
adamc@11 317
adamc@4 318 simpl.
adamc@11 319 (** [[
adamc@11 320 n : nat
adamc@11 321 s : stack
adamc@11 322 p : list instr
adamc@11 323 ============================
adamc@11 324 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 325 option stack :=
adamc@11 326 match p0 with
adamc@11 327 | nil => Some s0
adamc@11 328 | i :: p' =>
adamc@11 329 match instrDenote i s0 with
adamc@11 330 | Some s' => progDenote p' s'
adamc@11 331 | None => None (A:=stack)
adamc@11 332 end
adamc@11 333 end) p (n :: s) = progDenote p (n :: s)
adamc@206 334
adamc@11 335 ]]
adamc@11 336
adam@311 337 Now we can unexpand the definition of [progDenote]:%\index{tactics!fold}%
adamc@11 338 *)
adamc@11 339
adamc@11 340 fold progDenote.
adamc@11 341 (** [[
adamc@11 342 n : nat
adamc@11 343 s : stack
adamc@11 344 p : list instr
adamc@11 345 ============================
adamc@11 346 progDenote p (n :: s) = progDenote p (n :: s)
adamc@206 347
adamc@11 348 ]]
adamc@11 349
adam@311 350 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 351 *)
adamc@11 352
adamc@4 353 reflexivity.
adamc@2 354
adamc@11 355 (** On to the second inductive case:
adamc@11 356
adamc@11 357 [[
adamc@11 358 b : binop
adamc@11 359 e1 : exp
adamc@11 360 IHe1 : forall (s : stack) (p : list instr),
adamc@11 361 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 362 e2 : exp
adamc@11 363 IHe2 : forall (s : stack) (p : list instr),
adamc@11 364 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 365 ============================
adamc@11 366 forall (s : stack) (p : list instr),
adamc@11 367 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 368 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@206 369
adamc@11 370 ]]
adamc@11 371
adam@311 372 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 373
adam@399 374 We start out the same way as before, introducing new free variables and unfolding and folding the appropriate definitions. The seemingly frivolous [unfold]/[fold] pairs are actually accomplishing useful work, because [unfold] will sometimes perform easy simplifications. %\index{tactics!intros}\index{tactics!unfold}\index{tactics!fold}% *)
adamc@11 375
adamc@4 376 intros.
adamc@4 377 unfold compile.
adamc@4 378 fold compile.
adamc@4 379 unfold expDenote.
adamc@4 380 fold expDenote.
adamc@11 381
adamc@44 382 (** 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 383
adamc@11 384 [[
adamc@11 385 b : binop
adamc@11 386 e1 : exp
adamc@11 387 IHe1 : forall (s : stack) (p : list instr),
adamc@11 388 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 389 e2 : exp
adamc@11 390 IHe2 : forall (s : stack) (p : list instr),
adamc@11 391 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 392 s : stack
adamc@11 393 p : list instr
adamc@11 394 ============================
adam@311 395 progDenote ((compile e2 ++ compile e1 ++ iBinop b :: nil) ++ p) s =
adamc@11 396 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 397
adamc@11 398 ]]
adamc@11 399
adam@471 400 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}% (Here and elsewhere, it is possible to tell the difference between inputs and outputs to Coq by periods at the ends of the inputs.) *)
adamc@11 401
adam@469 402 Check app_assoc_reverse.
adam@439 403 (** %\vspace{-.15in}%[[
adam@311 404 app_assoc_reverse
adamc@11 405 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adamc@206 406
adamc@11 407 ]]
adamc@11 408
adam@399 409 If we did not already know the name of the theorem, we could use the %\index{Vernacular commands!SearchRewrite}%[SearchRewrite] command to find it, based on a pattern that we would like to rewrite: *)
adam@277 410
adam@277 411 SearchRewrite ((_ ++ _) ++ _).
adam@439 412 (** %\vspace{-.15in}%[[
adam@311 413 app_assoc_reverse:
adam@311 414 forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adam@311 415 ]]
adam@311 416 %\vspace{-.25in}%
adam@311 417 [[
adam@311 418 app_assoc: forall (A : Type) (l m n : list A), l ++ m ++ n = (l ++ m) ++ n
adam@277 419
adam@277 420 ]]
adam@277 421
adam@311 422 We use [app_assoc_reverse] to perform a rewrite: %\index{tactics!rewrite}% *)
adamc@11 423
adam@311 424 rewrite app_assoc_reverse.
adamc@11 425
adam@439 426 (** %\noindent{}%changing the conclusion to:
adamc@11 427
adamc@206 428 [[
adam@311 429 progDenote (compile e2 ++ (compile e1 ++ iBinop b :: nil) ++ p) s =
adamc@11 430 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 431
adamc@11 432 ]]
adamc@11 433
adam@311 434 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 435
adamc@4 436 rewrite IHe2.
adamc@11 437 (** [[
adam@311 438 progDenote ((compile e1 ++ iBinop b :: nil) ++ p) (expDenote e2 :: s) =
adamc@11 439 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 440
adamc@11 441 ]]
adamc@11 442
adam@311 443 The same process lets us apply the remaining hypothesis.%\index{tactics!rewrite}% *)
adamc@11 444
adam@311 445 rewrite app_assoc_reverse.
adamc@4 446 rewrite IHe1.
adamc@11 447 (** [[
adam@311 448 progDenote ((iBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
adamc@11 449 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@206 450
adamc@11 451 ]]
adamc@11 452
adam@311 453 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 454 *)
adamc@11 455
adamc@11 456 unfold progDenote at 1.
adamc@4 457 simpl.
adamc@11 458 fold progDenote.
adamc@4 459 reflexivity.
adamc@11 460
adam@311 461 (** And the proof is completed, as indicated by the message: *)
adamc@11 462
adam@399 463 (**
adam@399 464 <<
adam@399 465 Proof completed.
adam@399 466 >>
adam@399 467 *)
adamc@11 468
adam@311 469 (** 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 470 *)
adamc@11 471
adamc@4 472 Abort.
adamc@2 473
adam@311 474 (** %\index{tactics!induction}\index{tactics!crush}% *)
adam@311 475
adamc@26 476 Lemma compile_correct' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 477 progDenote p (expDenote e :: s).
adamc@4 478 induction e; crush.
adamc@4 479 Qed.
adamc@2 480
adam@328 481 (** We need only to state the basic inductive proof scheme and call a tactic that automates the tedious reasoning in between. In contrast to the period tactic terminator from our last proof, the %\index{tactics!semicolon}%semicolon tactic separator supports structured, compositional proofs. The tactic [t1; t2] has the effect of running [t1] and then running [t2] on each remaining subgoal. The semicolon is one of the most fundamental building blocks of effective proof automation. The period terminator is very useful for exploratory proving, where you need to see intermediate proof states, but final proofs of any serious complexity should have just one period, terminating a single compound tactic that probably uses semicolons.
adamc@11 482
adam@399 483 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 484
adam@398 485 The %\index{Vernacular commands!Qed}%[Qed] command checks that the proof is finished and, if so, saves it. The tactic commands we have written above are an example of a _proof script_, or a series of Ltac programs; while [Qed] uses the result of the script to generate a _proof term_, a well-typed term of Gallina. To believe that a theorem is true, we only need to trust that the (relatively simple) checker for proof terms is correct; the use of proof scripts is immaterial. Part I of this book will introduce the principles behind encoding all proofs as terms of Gallina.
adam@311 486
adam@311 487 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 488
adamc@26 489 Theorem compile_correct : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 490 intros.
adamc@11 491 (** [[
adamc@11 492 e : exp
adamc@11 493 ============================
adamc@11 494 progDenote (compile e) nil = Some (expDenote e :: nil)
adamc@206 495
adamc@11 496 ]]
adamc@11 497
adamc@26 498 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 499
adamc@11 500 Check app_nil_end.
adamc@11 501 (** [[
adamc@11 502 app_nil_end
adamc@11 503 : forall (A : Type) (l : list A), l = l ++ nil
adam@302 504 ]]
adam@311 505 %\index{tactics!rewrite}% *)
adamc@11 506
adamc@4 507 rewrite (app_nil_end (compile e)).
adamc@11 508
adam@417 509 (** This time, we explicitly specify the value of the variable [l] from the theorem statement, since multiple expressions of list type appear in the conclusion. The [rewrite] tactic might choose the wrong place to rewrite if we did not specify which we want.
adamc@11 510
adamc@11 511 [[
adamc@11 512 e : exp
adamc@11 513 ============================
adamc@11 514 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
adamc@206 515
adamc@11 516 ]]
adamc@11 517
adam@311 518 Now we can apply the lemma.%\index{tactics!rewrite}% *)
adamc@11 519
adamc@26 520 rewrite compile_correct'.
adamc@11 521 (** [[
adamc@11 522 e : exp
adamc@11 523 ============================
adamc@11 524 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
adamc@206 525
adamc@11 526 ]]
adamc@11 527
adam@311 528 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 529
adamc@4 530 reflexivity.
adamc@4 531 Qed.
adamc@22 532 (* end thide *)
adamc@14 533
adam@475 534 (** This proof can be shortened and automated, but we leave that task as an exercise for the reader. *)
adam@311 535
adamc@14 536
adamc@20 537 (** * Typed Expressions *)
adamc@14 538
adamc@14 539 (** 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 540
adamc@20 541 (** ** Source Language *)
adamc@14 542
adamc@15 543 (** We define a trivial language of types to classify our expressions: *)
adamc@15 544
adamc@14 545 Inductive type : Set := Nat | Bool.
adamc@14 546
adam@277 547 (** 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 548
adam@277 549 Now we define an expanded set of binary operators. *)
adamc@15 550
adamc@14 551 Inductive tbinop : type -> type -> type -> Set :=
adamc@14 552 | TPlus : tbinop Nat Nat Nat
adamc@14 553 | TTimes : tbinop Nat Nat Nat
adamc@14 554 | TEq : forall t, tbinop t t Bool
adamc@14 555 | TLt : tbinop Nat Nat Bool.
adamc@14 556
adam@398 557 (** The definition of [tbinop] is different from [binop] in an important way. Where we declared that [binop] has type [Set], here we declare that [tbinop] has type [type -> type -> type -> Set]. We define [tbinop] as an _indexed type family_. Indexed inductive types are at the heart of Coq's expressive power; almost everything else of interest is defined in terms of them.
adamc@15 558
adam@452 559 The intuitive explanation of [tbinop] is that a [tbinop t1 t2 t] is a binary operator whose operands should have types [t1] and [t2], and whose result has type [t]. For instance, constructor [TLt] (for less-than comparison of numbers) is assigned type [tbinop Nat Nat Bool], meaning the operator's arguments are naturals and its result is Boolean. The type of [TEq] introduces a small bit of additional complication via polymorphism: we want to allow equality comparison of any two values of any type, as long as they have the _same_ type.
adam@312 560
adamc@15 561 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 562
adam@469 563 First, the indices of the range of each data constructor must be type variables bound at the top level of the datatype definition. There is no way to do what we did here, where we, for instance, say that [TPlus] is a constructor building a [tbinop] whose indices are all fixed at [Nat]. %\index{generalized algebraic datatypes}\index{GADTs|see{generalized algebraic datatypes}}% _Generalized algebraic datatypes_ (GADTs)%~\cite{GADT}% are a popular feature in %\index{GHC Haskell}%GHC Haskell, OCaml 4, and other languages that removes this first restriction.
adamc@15 564
adam@419 565 The second restriction is not lifted by GADTs. In ML and Haskell, indices of types must be types and may not be _expressions_. In Coq, types may be indexed by arbitrary Gallina terms. Type indices can live in the same universe as programs, and we can compute with them just like regular programs. Haskell supports a hobbled form of computation in type indices based on %\index{Haskell}%multi-parameter type classes, and recent extensions like type functions bring Haskell programming even closer to "real" functional programming with types, but, without dependent typing, there must always be a gap between how one programs with types and how one programs normally.
adamc@15 566 *)
adamc@15 567
adam@399 568 (** We can define a similar type family for typed expressions, where a term of type [texp t] can be assigned object language type [t]. (It is conventional in the world of interactive theorem proving to call the language of the proof assistant the%\index{meta language}% _meta language_ and a language being formalized the%\index{object language}% _object language_.) *)
adamc@15 569
adamc@14 570 Inductive texp : type -> Set :=
adamc@14 571 | TNConst : nat -> texp Nat
adamc@14 572 | TBConst : bool -> texp Bool
adam@312 573 | TBinop : forall t1 t2 t, tbinop t1 t2 t -> texp t1 -> texp t2 -> texp t.
adamc@14 574
adam@447 575 (** Thanks to our use of dependent types, every well-typed [texp] represents a well-typed source expression, by construction. This turns out to be very convenient for many things we might want to do with expressions. For instance, it is easy to adapt our interpreter approach to defining semantics. We start by defining a function mapping the types of our object language into Coq types: *)
adamc@15 576
adamc@14 577 Definition typeDenote (t : type) : Set :=
adamc@14 578 match t with
adamc@14 579 | Nat => nat
adamc@14 580 | Bool => bool
adamc@14 581 end.
adamc@14 582
adam@448 583 (** 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 584
adamc@207 585 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@207 586 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@207 587 match b with
adamc@207 588 | TPlus => plus
adamc@207 589 | TTimes => mult
adam@277 590 | TEq Nat => beq_nat
adam@277 591 | TEq Bool => eqb
adam@312 592 | TLt => leb
adamc@207 593 end.
adamc@207 594
adam@399 595 (** This function has just a few differences from the denotation functions we saw earlier. First, [tbinop] is an indexed type, so its indices become additional arguments to [tbinopDenote]. Second, we need to perform a genuine%\index{dependent pattern matching}% _dependent pattern match_, where the necessary _type_ of each case body depends on the _value_ that has been matched. At this early stage, we will not go into detail on the many subtle aspects of Gallina that support dependent pattern-matching, but the subject is central to Part II of the book.
adam@312 596
adam@471 597 The same tricks suffice to define an expression denotation function in an unsurprising way. Note that the [type] arguments to the [TBinop] constructor must be included explicitly in pattern-matching, but here we write underscores because we do not need to refer to those arguments directly. *)
adamc@15 598
adamc@207 599 Fixpoint texpDenote t (e : texp t) : typeDenote t :=
adamc@207 600 match e with
adamc@14 601 | TNConst n => n
adamc@14 602 | TBConst b => b
adamc@14 603 | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
adamc@14 604 end.
adamc@14 605
adamc@17 606 (** We can evaluate a few example programs to convince ourselves that this semantics is correct. *)
adamc@17 607
adamc@17 608 Eval simpl in texpDenote (TNConst 42).
adamc@207 609 (** [= 42 : typeDenote Nat] *)
adamc@207 610
adam@419 611 (* begin hide *)
adam@419 612 Eval simpl in texpDenote (TBConst false).
adam@419 613 (* end hide *)
adamc@17 614 Eval simpl in texpDenote (TBConst true).
adamc@207 615 (** [= true : typeDenote Bool] *)
adamc@207 616
adam@312 617 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 618 (TNConst 7)).
adamc@207 619 (** [= 28 : typeDenote Nat] *)
adamc@207 620
adam@312 621 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 622 (TNConst 7)).
adam@399 623 (** [= false : typeDenote Bool] *)
adamc@207 624
adam@312 625 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 626 (TNConst 7)).
adamc@207 627 (** [= true : typeDenote Bool] *)
adamc@17 628
adam@442 629 (** %\smallskip{}%Now we are ready to define a suitable stack machine target for compilation. *)
adam@442 630
adamc@14 631
adamc@20 632 (** ** Target Language *)
adamc@14 633
adam@442 634 (** 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 635
adamc@18 636 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 637
adamc@18 638 We start by defining stack types, which classify sets of possible stacks. *)
adamc@18 639
adamc@14 640 Definition tstack := list type.
adamc@14 641
adamc@18 642 (** 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 643
adamc@18 644 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 645
adamc@14 646 Inductive tinstr : tstack -> tstack -> Set :=
adam@312 647 | TiNConst : forall s, nat -> tinstr s (Nat :: s)
adam@312 648 | TiBConst : forall s, bool -> tinstr s (Bool :: s)
adam@311 649 | TiBinop : forall arg1 arg2 res s,
adamc@14 650 tbinop arg1 arg2 res
adamc@14 651 -> tinstr (arg1 :: arg2 :: s) (res :: s).
adamc@14 652
adamc@18 653 (** 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 654
adamc@14 655 Inductive tprog : tstack -> tstack -> Set :=
adamc@14 656 | TNil : forall s, tprog s s
adamc@14 657 | TCons : forall s1 s2 s3,
adamc@14 658 tinstr s1 s2
adamc@14 659 -> tprog s2 s3
adamc@14 660 -> tprog s1 s3.
adamc@14 661
adamc@18 662 (** 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 663
adamc@14 664 Fixpoint vstack (ts : tstack) : Set :=
adamc@14 665 match ts with
adamc@14 666 | nil => unit
adamc@14 667 | t :: ts' => typeDenote t * vstack ts'
adamc@14 668 end%type.
adamc@14 669
adam@312 670 (** 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 671
adam@312 672 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 673
adamc@14 674 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
adamc@207 675 match i with
adam@312 676 | TiNConst _ n => fun s => (n, s)
adam@312 677 | TiBConst _ b => fun s => (b, s)
adam@311 678 | TiBinop _ _ _ _ b => fun s =>
adam@312 679 let '(arg1, (arg2, s')) := s in
adam@312 680 ((tbinopDenote b) arg1 arg2, s')
adamc@14 681 end.
adamc@14 682
adamc@18 683 (** 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 684 [[
adamc@18 685 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
adamc@207 686 match i with
adam@312 687 | TiNConst _ n => (n, s)
adam@312 688 | TiBConst _ b => (b, s)
adam@311 689 | TiBinop _ _ _ _ b =>
adam@312 690 let '(arg1, (arg2, s')) := s in
adam@312 691 ((tbinopDenote b) arg1 arg2, s')
adamc@18 692 end.
adamc@205 693 ]]
adamc@205 694
adam@447 695 The Coq type checker complains that:
adamc@18 696
adam@312 697 <<
adamc@18 698 The term "(n, s)" has type "(nat * vstack ts)%type"
adamc@207 699 while it is expected to have type "vstack ?119".
adam@312 700 >>
adamc@207 701
adam@465 702 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 treatment of Gallina's typing rules will explain why this helps.
adamc@18 703 *)
adamc@18 704
adamc@18 705 (** We finish the semantics with a straightforward definition of program denotation. *)
adamc@18 706
adamc@207 707 Fixpoint tprogDenote ts ts' (p : tprog ts ts') : vstack ts -> vstack ts' :=
adamc@207 708 match p with
adamc@14 709 | TNil _ => fun s => s
adamc@14 710 | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
adamc@14 711 end.
adamc@14 712
adam@447 713 (** The same argument-postponing trick is crucial for this definition. *)
adam@447 714
adamc@14 715
adamc@14 716 (** ** Translation *)
adamc@14 717
adamc@19 718 (** To define our compilation, it is useful to have an auxiliary function for concatenating two stack machine programs. *)
adamc@19 719
adamc@207 720 Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') : tprog ts' ts'' -> tprog ts ts'' :=
adamc@207 721 match p with
adamc@14 722 | TNil _ => fun p' => p'
adamc@14 723 | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
adamc@14 724 end.
adamc@14 725
adamc@19 726 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)
adamc@19 727
adamc@207 728 Fixpoint tcompile t (e : texp t) (ts : tstack) : tprog ts (t :: ts) :=
adamc@207 729 match e with
adam@312 730 | TNConst n => TCons (TiNConst _ n) (TNil _)
adam@312 731 | TBConst b => TCons (TiBConst _ b) (TNil _)
adamc@14 732 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
adam@311 733 (tconcat (tcompile e1 _) (TCons (TiBinop _ b) (TNil _)))
adamc@14 734 end.
adamc@14 735
adam@398 736 (** One interesting feature of the definition is the underscores appearing to the right of [=>] arrows. Haskell and ML programmers are quite familiar with compilers that infer type parameters to polymorphic values. In Coq, it is possible to go even further and ask the system to infer arbitrary terms, by writing underscores in place of specific values. You may have noticed that we have been calling functions without specifying all of their arguments. For instance, the recursive calls here to [tcompile] omit the [t] argument. Coq's _implicit argument_ mechanism automatically inserts underscores for arguments that it will probably be able to infer. Inference of such values is far from complete, though; generally, it only works in cases similar to those encountered with polymorphic type instantiation in Haskell and ML.
adamc@19 737
adamc@19 738 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 739
adamc@14 740 Print tcompile.
adam@439 741 (** %\vspace{-.15in}%[[
adamc@19 742 tcompile =
adamc@19 743 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
adamc@19 744 tprog ts (t :: ts) :=
adamc@19 745 match e in (texp t0) return (tprog ts (t0 :: ts)) with
adam@312 746 | TNConst n => TCons (TiNConst ts n) (TNil (Nat :: ts))
adam@312 747 | TBConst b => TCons (TiBConst ts b) (TNil (Bool :: ts))
adamc@19 748 | TBinop arg1 arg2 res b e1 e2 =>
adamc@19 749 tconcat (tcompile arg2 e2 ts)
adamc@19 750 (tconcat (tcompile arg1 e1 (arg2 :: ts))
adam@311 751 (TCons (TiBinop ts b) (TNil (res :: ts))))
adamc@19 752 end
adamc@19 753 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts)
adam@302 754 ]]
adam@302 755 *)
adamc@19 756
adamc@19 757
adamc@19 758 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)
adamc@19 759
adamc@19 760 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
adam@399 761 (** [= (42, tt) : vstack (Nat :: nil)] *)
adamc@207 762
adamc@19 763 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
adam@399 764 (** [= (true, tt) : vstack (Bool :: nil)] *)
adamc@207 765
adam@312 766 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2)
adam@312 767 (TNConst 2)) (TNConst 7)) nil) tt.
adam@399 768 (** [= (28, tt) : vstack (Nat :: nil)] *)
adamc@207 769
adam@312 770 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2)
adam@312 771 (TNConst 2)) (TNConst 7)) nil) tt.
adam@399 772 (** [= (false, tt) : vstack (Bool :: nil)] *)
adamc@207 773
adam@312 774 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2))
adam@312 775 (TNConst 7)) nil) tt.
adam@399 776 (** [= (true, tt) : vstack (Bool :: nil)] *)
adamc@19 777
adam@442 778 (** %\smallskip{}%The compiler seems to be working, so let us turn to proving that it _always_ works. *)
adam@442 779
adamc@14 780
adamc@20 781 (** ** Translation Correctness *)
adamc@20 782
adamc@20 783 (** We can state a correctness theorem similar to the last one. *)
adamc@20 784
adamc@207 785 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 786 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@20 787 (* begin hide *)
adamc@20 788 Abort.
adamc@20 789 (* end hide *)
adamc@22 790 (* begin thide *)
adamc@20 791
adam@312 792 (** 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 793
adamc@207 794 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 795 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@20 796
adam@419 797 (** While lemma [compile_correct'] quantified over a program that is the "continuation"%~\cite{continuations}% for the expression we are considering, here we avoid drawing in any extra syntactic elements. In addition to the source expression and its type, we also quantify over an initial stack type and a stack compatible with it. Running the compilation of the program starting from that stack, we should arrive at a stack that differs only in having the program's denotation pushed onto it.
adamc@20 798
adamc@20 799 Let us try to prove this theorem in the same way that we settled on in the last section. *)
adamc@20 800
adamc@14 801 induction e; crush.
adamc@20 802
adamc@20 803 (** We are left with this unproved conclusion:
adamc@20 804
adamc@20 805 [[
adamc@20 806 tprogDenote
adamc@20 807 (tconcat (tcompile e2 ts)
adamc@20 808 (tconcat (tcompile e1 (arg2 :: ts))
adam@311 809 (TCons (TiBinop ts t) (TNil (res :: ts))))) s =
adamc@20 810 (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
adamc@207 811
adamc@20 812 ]]
adamc@20 813
adam@312 814 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 815 *)
adamc@207 816
adamc@14 817 Abort.
adamc@14 818
adamc@26 819 Lemma tconcat_correct : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
adamc@14 820 (s : vstack ts),
adamc@14 821 tprogDenote (tconcat p p') s
adamc@14 822 = tprogDenote p' (tprogDenote p s).
adamc@14 823 induction p; crush.
adamc@14 824 Qed.
adamc@14 825
adamc@20 826 (** This one goes through completely automatically.
adamc@20 827
adam@316 828 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 829
adam@375 830 Hint Rewrite tconcat_correct.
adamc@14 831
adam@419 832 (** Here we meet the pervasive concept of a _hint_. Many proofs can be found through exhaustive enumerations of combinations of possible proof steps; hints provide the set of steps to consider. The tactic [crush] is applying such brute force search for us silently, and it will consider more possibilities as we add more hints. This particular hint asks that the lemma be used for left-to-right rewriting.
adam@312 833
adam@312 834 Now we are ready to return to [tcompile_correct'], proving it automatically this time. *)
adamc@20 835
adamc@207 836 Lemma tcompile_correct' : forall t (e : texp t) ts (s : vstack ts),
adamc@207 837 tprogDenote (tcompile e ts) s = (texpDenote e, s).
adamc@14 838 induction e; crush.
adamc@14 839 Qed.
adamc@14 840
adamc@20 841 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
adamc@20 842
adam@375 843 Hint Rewrite tcompile_correct'.
adamc@14 844
adamc@207 845 Theorem tcompile_correct : forall t (e : texp t),
adamc@207 846 tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@14 847 crush.
adamc@14 848 Qed.
adamc@22 849 (* end thide *)
adam@312 850
adam@399 851 (** It is probably worth emphasizing that we are doing more than building mathematical models. Our compilers are functional programs that can be executed efficiently. One strategy for doing so is based on%\index{program extraction}% _program extraction_, which generates OCaml code from Coq developments. For instance, we run a command to output the OCaml version of [tcompile]:%\index{Vernacular commands!Extraction}% *)
adam@312 852
adam@312 853 Extraction tcompile.
adam@312 854
adam@312 855 (** <<
adam@312 856 let rec tcompile t e ts =
adam@312 857 match e with
adam@312 858 | TNConst n ->
adam@312 859 TCons (ts, (Cons (Nat, ts)), (Cons (Nat, ts)), (TiNConst (ts, n)), (TNil
adam@312 860 (Cons (Nat, ts))))
adam@312 861 | TBConst b ->
adam@312 862 TCons (ts, (Cons (Bool, ts)), (Cons (Bool, ts)), (TiBConst (ts, b)),
adam@312 863 (TNil (Cons (Bool, ts))))
adam@312 864 | TBinop (t1, t2, t0, b, e1, e2) ->
adam@312 865 tconcat ts (Cons (t2, ts)) (Cons (t0, ts)) (tcompile t2 e2 ts)
adam@312 866 (tconcat (Cons (t2, ts)) (Cons (t1, (Cons (t2, ts)))) (Cons (t0, ts))
adam@312 867 (tcompile t1 e1 (Cons (t2, ts))) (TCons ((Cons (t1, (Cons (t2,
adam@312 868 ts)))), (Cons (t0, ts)), (Cons (t0, ts)), (TiBinop (t1, t2, t0, ts,
adam@312 869 b)), (TNil (Cons (t0, ts))))))
adam@312 870 >>
adam@312 871
adam@312 872 We can compile this code with the usual OCaml compiler and obtain an executable program with halfway decent performance.
adam@312 873
adam@312 874 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. *)