annotate src/StackMachine.v @ 15:d8c81e19e7d3

Source language semantics
author Adam Chlipala <adamc@hcoop.net>
date Wed, 03 Sep 2008 15:00:56 -0400
parents 8aed4562b62b
children 4d8ca6845e67
rev   line source
adamc@2 1 (* Copyright (c) 2008, 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 *)
adamc@2 11 Require Import List.
adamc@2 12
adamc@2 13 Require Import Tactics.
adamc@14 14
adamc@14 15 Set Implicit Arguments.
adamc@3 16 (* end hide *)
adamc@2 17
adamc@2 18
adamc@9 19 (** 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. I assume that you have installed Coq and Proof General.
adamc@9 20
adamc@11 21 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 a line [Require Import List Tactics] at the start of the file, to match some code hidden from the chapter source, and be sure to run the Coq binary %\texttt{%#<tt>#coqtop#</tt>#%}% with the command-line argument %\texttt{%#<tt>#-I SRC#</tt>#%}%, where %\texttt{%#<tt>#SRC#</tt>#%}% is the path to a directory containing the source for this book. In either case, if you have installed Proof General properly, it should start automatically when you visit a %\texttt{%#<tt>#.v#</tt>#%}% buffer in Emacs.
adamc@11 22
adamc@11 23 With Proof General, the portion of a buffer that Coq has processed is highlighted in some way, like being given a blue background. You step through Coq source files by positioning the point at the position you want Coq to run to and pressing C-C C-RET. This can be used both for normal step-by-step coding, by placing the point inside some command past the end of the highlighted region; and for undoing, by placing the point inside the highlighted region. *)
adamc@9 24
adamc@9 25
adamc@2 26 (** * Arithmetic expressions over natural numbers *)
adamc@2 27
adamc@9 28 (** We will begin with that staple of compiler textbooks, arithemtic expressions over a single type of numbers. *)
adamc@9 29
adamc@9 30 (** ** Source language *)
adamc@9 31
adamc@9 32 (** We begin with the syntax of the source language. *)
adamc@2 33
adamc@4 34 Inductive binop : Set := Plus | Times.
adamc@2 35
adamc@9 36 (** Our first line of Coq code should be unsurprising to ML and Haskell programmers. We define an algebraic datatype [binop] to stand for the binary operators of our source language. There are just two wrinkles compared to ML and Haskell. First, we use the keyword [Inductive], in place of %\texttt{%#<tt>#data#</tt>#%}%, %\texttt{%#<tt>#datatype#</tt>#%}%, or %\texttt{%#<tt>#type#</tt>#%}%. This is not just a trivial surface syntax difference; inductive types in Coq are much more expressive than garden variety algebraic datatypes, essentially enabling us to encode all of mathematics, though we begin humbly in this chapter. Second, there is the [: Set] fragment, which declares that we are defining a datatype that should be thought of as a constituent of programs. Later, we will see other options for defining datatypes in the universe of proofs or in an infinite hierarchy of universes, encompassing both programs and proofs, that is useful in higher-order constructions. *)
adamc@9 37
adamc@4 38 Inductive exp : Set :=
adamc@4 39 | Const : nat -> exp
adamc@4 40 | Binop : binop -> exp -> exp -> exp.
adamc@2 41
adamc@9 42 (** 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 43
adamc@9 44 A note for readers following along in the PDF version: coqdoc supports pretty-printing of tokens in LaTeX or HTML. Where you see a right arrow character, the source contains the ASCII text %\texttt{%#<tt>#->#</tt>#%}%. Other examples of this substitution appearing in this chapter are a double right arrow for %\texttt{%#<tt>#=>#</tt>#%}% and the inverted 'A' symbol for %\texttt{%#<tt>#forall#</tt>#%}%. When in doubt about the ASCII version of a symbol, you can consult the chapter source code.
adamc@9 45
adamc@9 46 %\medskip%
adamc@9 47
adamc@9 48 Now we are ready to say what these programs mean. We will do this by writing an interpreter that can be thought of as a trivial operational or denotational semantics. (If you are not familiar with these semantic techniques, no need to worry; we will stick to "common sense" constructions.) *)
adamc@9 49
adamc@4 50 Definition binopDenote (b : binop) : nat -> nat -> nat :=
adamc@4 51 match b with
adamc@4 52 | Plus => plus
adamc@4 53 | Times => mult
adamc@4 54 end.
adamc@2 55
adamc@9 56 (** 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 57
adamc@9 58 [[
adamc@9 59 Definition binopDenote : binop -> nat -> nat -> nat := fun (b : binop) =>
adamc@9 60 match b with
adamc@9 61 | Plus => plus
adamc@9 62 | Times => mult
adamc@9 63 end.
adamc@9 64
adamc@9 65 In this example, we could also omit all of the type annotations, arriving at:
adamc@9 66
adamc@9 67 [[
adamc@9 68 Definition binopDenote := fun b =>
adamc@9 69 match b with
adamc@9 70 | Plus => plus
adamc@9 71 | Times => mult
adamc@9 72 end.
adamc@9 73
adamc@9 74 Languages like Haskell and ML have a convenient %\textit{%#<i>#principal typing#</i>#%}% property, which gives us strong guarantees about how effective type inference will be. Unfortunately, Coq's type system is so expressive that any kind of "complete" type inference is impossible, and the task even seems to be hard heuristically in practice. Nonetheless, Coq includes some very helpful heuristics, many of them copying the workings of Haskell and ML type-checkers for programs that fall in simple fragments of Coq's language.
adamc@9 75
adamc@9 76 This is as good a time as any to mention the preponderance of different languages associated with Coq. The theoretical foundation of Coq is a formal system called the %\textit{%#<i>#Calculus of Inductive Constructions (CIC)#</i>#%}%, which is an extension of the older %\textit{%#<i>#Calculus of Constructions (CoC)#</i>#%}%. CIC is quite a spartan foundation, which is helpful for proving metatheory but not so helpful for real development. Still, it is nice to know that it has been proved that CIC enjoys properties like %\textit{%#<i>#strong normalization#</i>#%}%, meaning that every program (and, more importantly, every proof term) terminates; and %\textit{%#<i>#relative consistency#</i>#%}% with systems like versions of Zermelo Fraenkel set theory, which roughly means that you can believe that Coq proofs mean that the corresponding propositions are "really true," if you believe in set theory.
adamc@9 77
adamc@9 78 Coq is actually based on an extension of CIC called %\textit{%#<i>#Gallina#</i>#%}%. The text after the [:=] and before the period in the last code example is a term of Gallina. Gallina adds many useful features that are not compiled internalluy 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 79
adamc@9 80 Commands like [Inductive] and [Definition] are part of %\textit{%#<i>#the vernacular#</i>#%}%, which includes all sorts of useful queries and requests to the Coq system.
adamc@9 81
adamc@9 82 Finally, there is %\textit{%#<i>#Ltac#</i>#%}%, Coq's domain-specific language for writing proofs and decision procedures. We will see some basic examples of Ltac later in this chapter, and much of this book is devoted to more involved Ltac examples.
adamc@9 83
adamc@9 84 %\medskip%
adamc@9 85
adamc@9 86 We can give a simple definition of the meaning of an expression: *)
adamc@9 87
adamc@4 88 Fixpoint expDenote (e : exp) : nat :=
adamc@4 89 match e with
adamc@4 90 | Const n => n
adamc@4 91 | Binop b e1 e2 => (binopDenote b) (expDenote e1) (expDenote e2)
adamc@4 92 end.
adamc@2 93
adamc@9 94 (** We declare explicitly that this is a recursive definition, using the keyword [Fixpoint]. The rest should be old hat for functional programmers. *)
adamc@2 95
adamc@9 96
adamc@9 97 (** ** Target language *)
adamc@4 98
adamc@10 99 (** We will compile our source programs onto a simple stack machine, whose syntax is: *)
adamc@2 100
adamc@4 101 Inductive instr : Set :=
adamc@4 102 | IConst : nat -> instr
adamc@4 103 | IBinop : binop -> instr.
adamc@2 104
adamc@4 105 Definition prog := list instr.
adamc@4 106 Definition stack := list nat.
adamc@2 107
adamc@10 108 (** 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 109
adamc@10 110 We can give instructions meanings as functions from stacks to optional stacks, where running an instruction results in [None] in case of a stack underflow and results in [Some s'] when the result of execution is the new stack [s']. [::] is the "list cons" operator from the Coq standard library. *)
adamc@10 111
adamc@4 112 Definition instrDenote (i : instr) (s : stack) : option stack :=
adamc@4 113 match i with
adamc@4 114 | IConst n => Some (n :: s)
adamc@4 115 | IBinop b =>
adamc@4 116 match s with
adamc@4 117 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
adamc@4 118 | _ => None
adamc@4 119 end
adamc@4 120 end.
adamc@2 121
adamc@10 122 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *)
adamc@10 123
adamc@4 124 Fixpoint progDenote (p : prog) (s : stack) {struct p} : option stack :=
adamc@4 125 match p with
adamc@4 126 | nil => Some s
adamc@4 127 | i :: p' =>
adamc@4 128 match instrDenote i s with
adamc@4 129 | None => None
adamc@4 130 | Some s' => progDenote p' s'
adamc@4 131 end
adamc@4 132 end.
adamc@2 133
adamc@10 134 (** There is one interesting difference compared to our previous example of a [Fixpoint]. This recursive function takes two arguments, [p] and [s]. It is critical for the soundness of Coq that every program terminate, so a shallow syntactic termination check is imposed on every recursive function definition. One of the function parameters must be designated to decrease monotonically across recursive calls. That is, every recursive call must use a version of that argument that has been pulled out of the current argument by some number of [match] expressions. [expDenote] has only one argument, so we did not need to specify which of its arguments decreases. For [progDenote], we resolve the ambiguity by writing [{struct p}] to indicate that argument [p] decreases structurally. *)
adamc@2 135
adamc@4 136
adamc@9 137 (** ** Translation *)
adamc@4 138
adamc@10 139 (** Our compiler itself is now unsurprising. [++] is the list concatenation operator from the Coq standard library. *)
adamc@2 140
adamc@4 141 Fixpoint compile (e : exp) : prog :=
adamc@4 142 match e with
adamc@4 143 | Const n => IConst n :: nil
adamc@4 144 | Binop b e1 e2 => compile e2 ++ compile e1 ++ IBinop b :: nil
adamc@4 145 end.
adamc@2 146
adamc@2 147
adamc@9 148 (** ** Translation correctness *)
adamc@4 149
adamc@11 150 (** We are ready to prove that our compiler is implemented correctly. We can use a new vernacular command [Theorem] to start a correctness proof, in terms of the semantics we defined earlier: *)
adamc@11 151
adamc@11 152 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 153 (* begin hide *)
adamc@11 154 Abort.
adamc@11 155 (* end hide *)
adamc@11 156
adamc@11 157 (** Though a pencil-and-paper proof might clock out at this point, writing "by a routine induction on [e]," it turns out not to make sense to attack this proof directly. We need to use the standard trick of %\textit{%#<i>#strengthening the induction hypothesis#</i>#%}%. We do that by proving an auxiliary lemma:
adamc@11 158 *)
adamc@2 159
adamc@15 160 Lemma compileCorrect' : forall e p s, progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
adamc@11 161
adamc@11 162 (** After the period in the [Lemma] command, we are in %\textit{%#<i>#the interactive proof-editing mode#</i>#%}%. We find ourselves staring at this ominous screen of text:
adamc@11 163
adamc@11 164 [[
adamc@11 165 1 subgoal
adamc@11 166
adamc@11 167 ============================
adamc@15 168 forall (e : exp) (p : list instr) (s : stack),
adamc@15 169 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
adamc@11 170 ]]
adamc@11 171
adamc@11 172 Coq seems to be restating the lemma for us. What we are seeing is a limited case of a more general protocol for describing where we are in a proof. We are told that we have a single subgoal. In general, during a proof, we can have many pending subgoals, each of which is a logical proposition to prove. Subgoals can be proved in any order, but it usually works best to prove them in the order that Coq chooses.
adamc@11 173
adamc@11 174 Next in the output, we see our single subgoal described in full detail. There is a double-dashed line, above which would be our free variables and hypotheses, if we had any. Below the line is the conclusion, which, in general, is to be proved from the hypotheses.
adamc@11 175
adamc@11 176 We manipulate the proof state by running commands called %\textit{%#<i>#tactics#</i>#%}%. Let us start out by running one of the most important tactics:
adamc@11 177 *)
adamc@11 178
adamc@4 179 induction e.
adamc@2 180
adamc@11 181 (** 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 182
adamc@11 183 [[
adamc@11 184 2 subgoals
adamc@11 185
adamc@11 186 n : nat
adamc@11 187 ============================
adamc@11 188 forall (s : stack) (p : list instr),
adamc@11 189 progDenote (compile (Const n) ++ p) s =
adamc@11 190 progDenote p (expDenote (Const n) :: s)
adamc@11 191 ]]
adamc@11 192 [[
adamc@11 193 subgoal 2 is:
adamc@11 194 forall (s : stack) (p : list instr),
adamc@11 195 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 196 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@11 197 ]]
adamc@11 198
adamc@11 199 The first and current subgoal is displayed with the double-dashed line below free variables and hypotheses, while later subgoals are only summarized with their conclusions. We see an example of a free variable in the first subgoal; [n] is a free variable of type [nat]. The conclusion is the original theorem statement where [e] has been replaced by [Const n]. In a similar manner, the second case has [e] replaced by a generalized invocation of the [Binop] expression constructor. We can see that proving both cases corresponds to a standard proof by structural induction.
adamc@11 200
adamc@11 201 We begin the first case with another very common tactic.
adamc@11 202 *)
adamc@11 203
adamc@4 204 intros.
adamc@11 205
adamc@11 206 (** The current subgoal changes to:
adamc@11 207 [[
adamc@11 208
adamc@11 209 n : nat
adamc@11 210 s : stack
adamc@11 211 p : list instr
adamc@11 212 ============================
adamc@11 213 progDenote (compile (Const n) ++ p) s =
adamc@11 214 progDenote p (expDenote (Const n) :: s)
adamc@11 215 ]]
adamc@11 216
adamc@11 217 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
adamc@11 218
adamc@11 219 To progress further, we need to use the definitions of some of the functions appearing in the goal. The [unfold] tactic replaces an identifier with its definition.
adamc@11 220 *)
adamc@11 221
adamc@4 222 unfold compile.
adamc@11 223
adamc@11 224 (** [[
adamc@11 225
adamc@11 226 n : nat
adamc@11 227 s : stack
adamc@11 228 p : list instr
adamc@11 229 ============================
adamc@11 230 progDenote ((IConst n :: nil) ++ p) s =
adamc@11 231 progDenote p (expDenote (Const n) :: s)
adamc@11 232 ]] *)
adamc@11 233
adamc@4 234 unfold expDenote.
adamc@11 235
adamc@11 236 (** [[
adamc@11 237
adamc@11 238 n : nat
adamc@11 239 s : stack
adamc@11 240 p : list instr
adamc@11 241 ============================
adamc@11 242 progDenote ((IConst n :: nil) ++ p) s = progDenote p (n :: s)
adamc@11 243 ]]
adamc@11 244
adamc@11 245 We only need to unfold the first occurrence of [progDenote] to prove the goal: *)
adamc@11 246
adamc@11 247 unfold progDenote at 1.
adamc@11 248
adamc@11 249 (** [[
adamc@11 250
adamc@11 251 n : nat
adamc@11 252 s : stack
adamc@11 253 p : list instr
adamc@11 254 ============================
adamc@11 255 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 256 option stack :=
adamc@11 257 match p0 with
adamc@11 258 | nil => Some s0
adamc@11 259 | i :: p' =>
adamc@11 260 match instrDenote i s0 with
adamc@11 261 | Some s' => progDenote p' s'
adamc@11 262 | None => None (A:=stack)
adamc@11 263 end
adamc@11 264 end) ((IConst n :: nil) ++ p) s =
adamc@11 265 progDenote p (n :: s)
adamc@11 266 ]]
adamc@11 267
adamc@11 268 This last [unfold] has left us with an anonymous fixpoint version of [progDenote], which will generally happen when unfolding recursive definitions. Fortunately, in this case, we can eliminate such complications right away, since the structure of the argument [(IConst n :: nil) ++ p] is known, allowing us to simplify the internal pattern match with the [simpl] tactic:
adamc@11 269 *)
adamc@11 270
adamc@4 271 simpl.
adamc@11 272
adamc@11 273 (** [[
adamc@11 274
adamc@11 275 n : nat
adamc@11 276 s : stack
adamc@11 277 p : list instr
adamc@11 278 ============================
adamc@11 279 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
adamc@11 280 option stack :=
adamc@11 281 match p0 with
adamc@11 282 | nil => Some s0
adamc@11 283 | i :: p' =>
adamc@11 284 match instrDenote i s0 with
adamc@11 285 | Some s' => progDenote p' s'
adamc@11 286 | None => None (A:=stack)
adamc@11 287 end
adamc@11 288 end) p (n :: s) = progDenote p (n :: s)
adamc@11 289 ]]
adamc@11 290
adamc@11 291 Now we can unexpand the definition of [progDenote]:
adamc@11 292 *)
adamc@11 293
adamc@11 294 fold progDenote.
adamc@11 295
adamc@11 296 (** [[
adamc@11 297
adamc@11 298 n : nat
adamc@11 299 s : stack
adamc@11 300 p : list instr
adamc@11 301 ============================
adamc@11 302 progDenote p (n :: s) = progDenote p (n :: s)
adamc@11 303 ]]
adamc@11 304
adamc@11 305 It looks like we are at the end of this case, since we have a trivial equality. Indeed, a single tactic finishes us off:
adamc@11 306 *)
adamc@11 307
adamc@4 308 reflexivity.
adamc@2 309
adamc@11 310 (** On to the second inductive case:
adamc@11 311
adamc@11 312 [[
adamc@11 313
adamc@11 314 b : binop
adamc@11 315 e1 : exp
adamc@11 316 IHe1 : forall (s : stack) (p : list instr),
adamc@11 317 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 318 e2 : exp
adamc@11 319 IHe2 : forall (s : stack) (p : list instr),
adamc@11 320 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 321 ============================
adamc@11 322 forall (s : stack) (p : list instr),
adamc@11 323 progDenote (compile (Binop b e1 e2) ++ p) s =
adamc@11 324 progDenote p (expDenote (Binop b e1 e2) :: s)
adamc@11 325 ]]
adamc@11 326
adamc@11 327 We see our first example of hypotheses above the double-dashed line. They are the inductive hypotheses [IHe1] and [IHe2] corresponding to the subterms [e1] and [e2], respectively.
adamc@11 328
adamc@11 329 We start out the same way as before, introducing new free variables and unfolding and folding the appropriate definitions. The seemingly frivolous [unfold]/[fold] pairs are actually accomplishing useful work, because [unfold] will sometimes perform easy simplifications. *)
adamc@11 330
adamc@4 331 intros.
adamc@4 332 unfold compile.
adamc@4 333 fold compile.
adamc@4 334 unfold expDenote.
adamc@4 335 fold expDenote.
adamc@11 336
adamc@11 337 (** Now we arrive at a point where the tactics we have seen so far are insufficient:
adamc@11 338
adamc@11 339 [[
adamc@11 340
adamc@11 341 b : binop
adamc@11 342 e1 : exp
adamc@11 343 IHe1 : forall (s : stack) (p : list instr),
adamc@11 344 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
adamc@11 345 e2 : exp
adamc@11 346 IHe2 : forall (s : stack) (p : list instr),
adamc@11 347 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
adamc@11 348 s : stack
adamc@11 349 p : list instr
adamc@11 350 ============================
adamc@11 351 progDenote ((compile e2 ++ compile e1 ++ IBinop b :: nil) ++ p) s =
adamc@11 352 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@11 353 ]]
adamc@11 354
adamc@11 355 What we need is the associative law of list concatenation, available as a theorem [app_ass] in the standard library. *)
adamc@11 356
adamc@11 357 Check app_ass.
adamc@11 358
adamc@11 359 (** [[
adamc@11 360
adamc@11 361 app_ass
adamc@11 362 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
adamc@11 363 ]]
adamc@11 364
adamc@11 365 We use it to perform a rewrite: *)
adamc@11 366
adamc@4 367 rewrite app_ass.
adamc@11 368
adamc@11 369 (** changing the conclusion to: [[
adamc@11 370
adamc@11 371 progDenote (compile e2 ++ (compile e1 ++ IBinop b :: nil) ++ p) s =
adamc@11 372 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@11 373 ]]
adamc@11 374
adamc@11 375 Now we can notice that the lefthand side of the equality matches the lefthand side of the second inductive hypothesis, so we can rewrite with that hypothesis, too: *)
adamc@11 376
adamc@4 377 rewrite IHe2.
adamc@11 378
adamc@11 379 (** [[
adamc@11 380
adamc@11 381 progDenote ((compile e1 ++ IBinop b :: nil) ++ p) (expDenote e2 :: s) =
adamc@11 382 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@11 383 ]]
adamc@11 384
adamc@11 385 The same process lets us apply the remaining hypothesis. *)
adamc@11 386
adamc@4 387 rewrite app_ass.
adamc@4 388 rewrite IHe1.
adamc@11 389
adamc@11 390 (** [[
adamc@11 391
adamc@11 392 progDenote ((IBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
adamc@11 393 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
adamc@11 394 ]]
adamc@11 395
adamc@11 396 Now we can apply a similar sequence of tactics to that that ended the proof of the first case.
adamc@11 397 *)
adamc@11 398
adamc@11 399 unfold progDenote at 1.
adamc@4 400 simpl.
adamc@11 401 fold progDenote.
adamc@4 402 reflexivity.
adamc@11 403
adamc@11 404 (** And the proof is completed, as indicated by the message:
adamc@11 405
adamc@11 406 [[
adamc@11 407 Proof completed.
adamc@11 408
adamc@11 409 And there lies our first proof. Already, even for simple theorems like this, the final proof script is unstructured and not very enlightening to readers. If we extend this approach to more serious theorems, we arrive at the unreadable proof scripts that are the favorite complaints of opponents of tactic-based proving. Fortunately, Coq has rich support for scripted automation, and we can take advantage of such a scripted tactic (defined elsewhere) to make short work of this lemma. We abort the old proof attempt and start again.
adamc@11 410 *)
adamc@11 411
adamc@4 412 Abort.
adamc@2 413
adamc@4 414 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
adamc@4 415 progDenote p (expDenote e :: s).
adamc@4 416 induction e; crush.
adamc@4 417 Qed.
adamc@2 418
adamc@11 419 (** We need only to state the basic inductive proof scheme and call a tactic that automates the tedious reasoning in between. In contrast to the period tactic terminator from our last proof, the semicolon tactic separator supports structured, compositional proofs. The tactic [t1; t2] has the effect of running [t1] and then running [t2] on each remaining subgoal. The semicolon is one of the most fundamental building blocks of effective proof automation. The period terminator is very useful for exploratory proving, where you need to see intermediate proof states, but final proofs of any serious complexity should have just one period, terminating a single compound tactic that probably uses semicolons.
adamc@11 420
adamc@11 421 The proof of our main theorem is now easy. We prove it with four period-terminated tactics, though separating them with semicolons would work as well; the version here is easier to step through. *)
adamc@11 422
adamc@4 423 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
adamc@11 424 intros.
adamc@11 425
adamc@11 426 (** [[
adamc@11 427
adamc@11 428 e : exp
adamc@11 429 ============================
adamc@11 430 progDenote (compile e) nil = Some (expDenote e :: nil)
adamc@11 431 ]]
adamc@11 432
adamc@11 433 At this point, we want to massage the lefthand side to match the statement of [compileCorrect']. A theorem from the standard library is useful: *)
adamc@11 434
adamc@11 435 Check app_nil_end.
adamc@11 436
adamc@11 437 (** [[
adamc@11 438
adamc@11 439 app_nil_end
adamc@11 440 : forall (A : Type) (l : list A), l = l ++ nil
adamc@11 441 ]] *)
adamc@11 442
adamc@4 443 rewrite (app_nil_end (compile e)).
adamc@11 444
adamc@11 445 (** 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 446
adamc@11 447 [[
adamc@11 448
adamc@11 449 e : exp
adamc@11 450 ============================
adamc@11 451 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
adamc@11 452 ]]
adamc@11 453
adamc@11 454 Now we can apply the lemma. *)
adamc@11 455
adamc@4 456 rewrite compileCorrect'.
adamc@11 457
adamc@11 458 (** [[
adamc@11 459
adamc@11 460 e : exp
adamc@11 461 ============================
adamc@11 462 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
adamc@11 463 ]]
adamc@11 464
adamc@11 465 We are almost done. The lefthand and righthand sides can be seen to match by simple symbolic evaluation. That means we are in luck, because Coq identifies any pair of terms as equal whenever they normalize to the same result by symbolic evaluation. By the definition of [progDenote], that is the case here, but we do not need to worry about such details. A simple invocation of [reflexivity] does the normalization and checks that the two results are syntactically equal. *)
adamc@11 466
adamc@4 467 reflexivity.
adamc@4 468 Qed.
adamc@14 469
adamc@14 470
adamc@14 471 (** * Typed expressions *)
adamc@14 472
adamc@14 473 (** 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 474
adamc@14 475 (** ** Source language *)
adamc@14 476
adamc@15 477 (** We define a trivial language of types to classify our expressions: *)
adamc@15 478
adamc@14 479 Inductive type : Set := Nat | Bool.
adamc@14 480
adamc@15 481 (** Now we define an expanded set of binary operators. *)
adamc@15 482
adamc@14 483 Inductive tbinop : type -> type -> type -> Set :=
adamc@14 484 | TPlus : tbinop Nat Nat Nat
adamc@14 485 | TTimes : tbinop Nat Nat Nat
adamc@14 486 | TEq : forall t, tbinop t t Bool
adamc@14 487 | TLt : tbinop Nat Nat Bool.
adamc@14 488
adamc@15 489 (** The definition of [tbinop] is different from [binop] in an important way. Where we declared that [binop] has type [Set], here we declare that [tbinop] has type [type -> type -> type -> Set]. We define [tbinop] as an %\textit{%#<i>#indexed type family#</i>#%}%. Indexed inductive types are at the heart of Coq's expressive power; almost everything else of interest is defined in terms of them.
adamc@15 490
adamc@15 491 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 492
adamc@15 493 First, the indices of the range of each data constructor must be type variables bound at the top level of the datatype definition. There is no way to do what we did here, where we, for instance, say that [TPlus] is a constructor building a [tbinop] whose indices are all fixed at [Nat]. %\textit{%#<i>#Generalized algebraic datatypes (GADTs)#</i>#%}% are a popular feature in GHC Haskell and other languages that removes this first restriction.
adamc@15 494
adamc@15 495 The second restriction is not lifted by GADTs. In ML and Haskell, indices of types must be types and may not be %\textit{%#<i>#expressions#</i>#%}%. In Coq, types may be indiced by arbitrary Gallina terms. Type indices can live in the same universe as programs, and we can compute with them just like regular programs. Haskell supports a hobbled form of computation in type indices based on multi-parameter type classes, and recent extensions like type functions bring Haskell programming even closer to "real" functional programming with types, but, without dependent typing, there must always be a gap between how one programs with types and how one programs normally.
adamc@15 496 *)
adamc@15 497
adamc@15 498 (** We can define a similar type family for typed expressions. *)
adamc@15 499
adamc@14 500 Inductive texp : type -> Set :=
adamc@14 501 | TNConst : nat -> texp Nat
adamc@14 502 | TBConst : bool -> texp Bool
adamc@14 503 | TBinop : forall arg1 arg2 res, tbinop arg1 arg2 res -> texp arg1 -> texp arg2 -> texp res.
adamc@14 504
adamc@15 505 (** 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 506
adamc@14 507 Definition typeDenote (t : type) : Set :=
adamc@14 508 match t with
adamc@14 509 | Nat => nat
adamc@14 510 | Bool => bool
adamc@14 511 end.
adamc@14 512
adamc@15 513 (** It can take a few moments to come to terms with the fact that [Set], the type of types of programs, is itself a first-class type, and that we can write functions that return [Set]s. Past that wrinkle, the definition of [typeDenote] is trivial, relying on the [nat] and [bool] types from the Coq standard library.
adamc@15 514
adamc@15 515 We need to define a few auxiliary functions, implementing our boolean binary operators that do not appear with the right types in the standard library. They are entirely standard and ML-like, with the one caveat being that the Coq [nat] type uses a unary representation, where [O] is zero and [S n] is the successor of [n].
adamc@15 516 *)
adamc@15 517
adamc@14 518 Definition eq_bool (b1 b2 : bool) : bool :=
adamc@14 519 match b1, b2 with
adamc@14 520 | true, true => true
adamc@14 521 | false, false => true
adamc@14 522 | _, _ => false
adamc@14 523 end.
adamc@14 524
adamc@14 525 Fixpoint eq_nat (n1 n2 : nat) {struct n1} : bool :=
adamc@14 526 match n1, n2 with
adamc@14 527 | O, O => true
adamc@14 528 | S n1', S n2' => eq_nat n1' n2'
adamc@14 529 | _, _ => false
adamc@14 530 end.
adamc@14 531
adamc@14 532 Fixpoint lt (n1 n2 : nat) {struct n1} : bool :=
adamc@14 533 match n1, n2 with
adamc@14 534 | O, S _ => true
adamc@14 535 | S n1', S n2' => lt n1' n2'
adamc@14 536 | _, _ => false
adamc@14 537 end.
adamc@14 538
adamc@15 539 (** Now we can interpret binary operators: *)
adamc@15 540
adamc@14 541 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
adamc@14 542 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
adamc@14 543 match b in (tbinop arg1 arg2 res) return (typeDenote arg1 -> typeDenote arg2 -> typeDenote res) with
adamc@14 544 | TPlus => plus
adamc@14 545 | TTimes => mult
adamc@14 546 | TEq Nat => eq_nat
adamc@14 547 | TEq Bool => eq_bool
adamc@14 548 | TLt => lt
adamc@14 549 end.
adamc@14 550
adamc@15 551 (** This function has just a few differences from the denotation functions we saw earlier. First, [tbinop] is an indexed type, so its indices become additional arguments to [tbinopDenote]. Second, we need to perform a genuine %\textit{%#<i>#dependent pattern match#</i>#%}% to come up with a definition of this function that type-checks. In each branch of the [match], we need to use branch-specific information about the indices to [tbinop]. General type inference that takes such information into account is undecidable, and Coq avoids pursuing special heuristics for the problem, instead asking users to write annotations, like we see above on the line with [match].
adamc@15 552
adamc@15 553 The [in] annotation restates the type of the term being case-analyzed. Though we use the same names for the indices as we use in the type of the original argument binder, these are actually fresh variables, and they are %\textit{%#<i>#binding occcurrences#</i>#%}%. Their scope is the [return] clause. That is, [arg1], [arg2], and [arg3] are new bound variables bound only within the return clause [typeDenote arg1 -> typeDenote arg2 -> typeDenote res]. By being explicit about the functional relationship between the type indices and the match result, we regain decidable type inference.
adamc@15 554
adamc@15 555 The same tricks suffice to define an expression denotation function in an unsurprising way:
adamc@15 556 *)
adamc@15 557
adamc@14 558 Fixpoint texpDenote t (e : texp t) {struct e} : typeDenote t :=
adamc@14 559 match e in (texp t) return (typeDenote t) with
adamc@14 560 | TNConst n => n
adamc@14 561 | TBConst b => b
adamc@14 562 | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
adamc@14 563 end.
adamc@14 564
adamc@14 565
adamc@14 566 (** ** Target language *)
adamc@14 567
adamc@14 568 Definition tstack := list type.
adamc@14 569
adamc@14 570 Inductive tinstr : tstack -> tstack -> Set :=
adamc@14 571 | TINConst : forall s, nat -> tinstr s (Nat :: s)
adamc@14 572 | TIBConst : forall s, bool -> tinstr s (Bool :: s)
adamc@14 573 | TIBinop : forall arg1 arg2 res s,
adamc@14 574 tbinop arg1 arg2 res
adamc@14 575 -> tinstr (arg1 :: arg2 :: s) (res :: s).
adamc@14 576
adamc@14 577 Inductive tprog : tstack -> tstack -> Set :=
adamc@14 578 | TNil : forall s, tprog s s
adamc@14 579 | TCons : forall s1 s2 s3,
adamc@14 580 tinstr s1 s2
adamc@14 581 -> tprog s2 s3
adamc@14 582 -> tprog s1 s3.
adamc@14 583
adamc@14 584 Fixpoint vstack (ts : tstack) : Set :=
adamc@14 585 match ts with
adamc@14 586 | nil => unit
adamc@14 587 | t :: ts' => typeDenote t * vstack ts'
adamc@14 588 end%type.
adamc@14 589
adamc@14 590 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
adamc@14 591 match i in (tinstr ts ts') return (vstack ts -> vstack ts') with
adamc@14 592 | TINConst _ n => fun s => (n, s)
adamc@14 593 | TIBConst _ b => fun s => (b, s)
adamc@14 594 | TIBinop _ _ _ _ b => fun s =>
adamc@14 595 match s with
adamc@14 596 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
adamc@14 597 end
adamc@14 598 end.
adamc@14 599
adamc@14 600 Fixpoint tprogDenote ts ts' (p : tprog ts ts') {struct p} : vstack ts -> vstack ts' :=
adamc@14 601 match p in (tprog ts ts') return (vstack ts -> vstack ts') with
adamc@14 602 | TNil _ => fun s => s
adamc@14 603 | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
adamc@14 604 end.
adamc@14 605
adamc@14 606
adamc@14 607 (** ** Translation *)
adamc@14 608
adamc@14 609 Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') {struct p} : tprog ts' ts'' -> tprog ts ts'' :=
adamc@14 610 match p in (tprog ts ts') return (tprog ts' ts'' -> tprog ts ts'') with
adamc@14 611 | TNil _ => fun p' => p'
adamc@14 612 | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
adamc@14 613 end.
adamc@14 614
adamc@14 615 Fixpoint tcompile t (e : texp t) (ts : tstack) {struct e} : tprog ts (t :: ts) :=
adamc@14 616 match e in (texp t) return (tprog ts (t :: ts)) with
adamc@14 617 | TNConst n => TCons (TINConst _ n) (TNil _)
adamc@14 618 | TBConst b => TCons (TIBConst _ b) (TNil _)
adamc@14 619 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
adamc@14 620 (tconcat (tcompile e1 _) (TCons (TIBinop _ b) (TNil _)))
adamc@14 621 end.
adamc@14 622
adamc@14 623 Print tcompile.
adamc@14 624
adamc@14 625
adamc@14 626 (** ** Translation correctness *)
adamc@14 627
adamc@14 628 Lemma tcompileCorrect' : forall t (e : texp t)
adamc@14 629 ts (s : vstack ts),
adamc@14 630 tprogDenote (tcompile e ts) s
adamc@14 631 = (texpDenote e, s).
adamc@14 632 induction e; crush.
adamc@14 633 Abort.
adamc@14 634
adamc@14 635 Lemma tconcatCorrect : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
adamc@14 636 (s : vstack ts),
adamc@14 637 tprogDenote (tconcat p p') s
adamc@14 638 = tprogDenote p' (tprogDenote p s).
adamc@14 639 induction p; crush.
adamc@14 640 Qed.
adamc@14 641
adamc@14 642 Hint Rewrite tconcatCorrect : cpdt.
adamc@14 643
adamc@14 644 Lemma tcompileCorrect' : forall t (e : texp t)
adamc@14 645 ts (s : vstack ts),
adamc@14 646 tprogDenote (tcompile e ts) s
adamc@14 647 = (texpDenote e, s).
adamc@14 648 induction e; crush.
adamc@14 649 Qed.
adamc@14 650
adamc@14 651 Hint Rewrite tcompileCorrect' : cpdt.
adamc@14 652
adamc@14 653 Theorem tcompileCorrect : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
adamc@14 654 crush.
adamc@14 655 Qed.