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@20
|
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@20
|
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@16
|
96 (** It is convenient to be able to test definitions before starting to prove things about them. We can verify that our semantics is sensible by evaluating some sample uses. *)
|
adamc@16
|
97
|
adamc@16
|
98 Eval simpl in expDenote (Const 42).
|
adamc@18
|
99 (** [[
|
adamc@18
|
100 = 42 : nat
|
adamc@18
|
101 ]] *)
|
adamc@16
|
102 Eval simpl in expDenote (Binop Plus (Const 2) (Const 2)).
|
adamc@18
|
103 (** [[
|
adamc@18
|
104 = 4 : nat
|
adamc@18
|
105 ]] *)
|
adamc@16
|
106 Eval simpl in expDenote (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
|
adamc@18
|
107 (** [[
|
adamc@18
|
108 = 28 : nat
|
adamc@18
|
109 ]] *)
|
adamc@9
|
110
|
adamc@20
|
111 (** ** Target Language *)
|
adamc@4
|
112
|
adamc@10
|
113 (** We will compile our source programs onto a simple stack machine, whose syntax is: *)
|
adamc@2
|
114
|
adamc@4
|
115 Inductive instr : Set :=
|
adamc@4
|
116 | IConst : nat -> instr
|
adamc@4
|
117 | IBinop : binop -> instr.
|
adamc@2
|
118
|
adamc@4
|
119 Definition prog := list instr.
|
adamc@4
|
120 Definition stack := list nat.
|
adamc@2
|
121
|
adamc@10
|
122 (** 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
|
123
|
adamc@10
|
124 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
|
125
|
adamc@4
|
126 Definition instrDenote (i : instr) (s : stack) : option stack :=
|
adamc@4
|
127 match i with
|
adamc@4
|
128 | IConst n => Some (n :: s)
|
adamc@4
|
129 | IBinop b =>
|
adamc@4
|
130 match s with
|
adamc@4
|
131 | arg1 :: arg2 :: s' => Some ((binopDenote b) arg1 arg2 :: s')
|
adamc@4
|
132 | _ => None
|
adamc@4
|
133 end
|
adamc@4
|
134 end.
|
adamc@2
|
135
|
adamc@10
|
136 (** With [instrDenote] defined, it is easy to define a function [progDenote], which iterates application of [instrDenote] through a whole program. *)
|
adamc@10
|
137
|
adamc@4
|
138 Fixpoint progDenote (p : prog) (s : stack) {struct p} : option stack :=
|
adamc@4
|
139 match p with
|
adamc@4
|
140 | nil => Some s
|
adamc@4
|
141 | i :: p' =>
|
adamc@4
|
142 match instrDenote i s with
|
adamc@4
|
143 | None => None
|
adamc@4
|
144 | Some s' => progDenote p' s'
|
adamc@4
|
145 end
|
adamc@4
|
146 end.
|
adamc@2
|
147
|
adamc@10
|
148 (** 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
|
149
|
adamc@4
|
150
|
adamc@9
|
151 (** ** Translation *)
|
adamc@4
|
152
|
adamc@10
|
153 (** Our compiler itself is now unsurprising. [++] is the list concatenation operator from the Coq standard library. *)
|
adamc@2
|
154
|
adamc@4
|
155 Fixpoint compile (e : exp) : prog :=
|
adamc@4
|
156 match e with
|
adamc@4
|
157 | Const n => IConst n :: nil
|
adamc@4
|
158 | Binop b e1 e2 => compile e2 ++ compile e1 ++ IBinop b :: nil
|
adamc@4
|
159 end.
|
adamc@2
|
160
|
adamc@2
|
161
|
adamc@16
|
162 (** 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
|
163
|
adamc@16
|
164 Eval simpl in compile (Const 42).
|
adamc@18
|
165 (** [[
|
adamc@18
|
166 = IConst 42 :: nil : prog
|
adamc@18
|
167 ]] *)
|
adamc@16
|
168 Eval simpl in compile (Binop Plus (Const 2) (Const 2)).
|
adamc@18
|
169 (** [[
|
adamc@18
|
170 = IConst 2 :: IConst 2 :: IBinop Plus :: nil : prog
|
adamc@18
|
171 ]] *)
|
adamc@16
|
172 Eval simpl in compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7)).
|
adamc@18
|
173 (** [[
|
adamc@18
|
174 = IConst 7 :: IConst 2 :: IConst 2 :: IBinop Plus :: IBinop Times :: nil : prog
|
adamc@18
|
175 ]] *)
|
adamc@16
|
176
|
adamc@16
|
177 (** We can also run our compiled programs and chedk that they give the right results. *)
|
adamc@16
|
178
|
adamc@16
|
179 Eval simpl in progDenote (compile (Const 42)) nil.
|
adamc@18
|
180 (** [[
|
adamc@18
|
181 = Some (42 :: nil) : option stack
|
adamc@18
|
182 ]] *)
|
adamc@16
|
183 Eval simpl in progDenote (compile (Binop Plus (Const 2) (Const 2))) nil.
|
adamc@18
|
184 (** [[
|
adamc@18
|
185 = Some (4 :: nil) : option stack
|
adamc@18
|
186 ]] *)
|
adamc@16
|
187 Eval simpl in progDenote (compile (Binop Times (Binop Plus (Const 2) (Const 2)) (Const 7))) nil.
|
adamc@18
|
188 (** [[
|
adamc@18
|
189 = Some (28 :: nil) : option stack
|
adamc@18
|
190 ]] *)
|
adamc@16
|
191
|
adamc@16
|
192
|
adamc@20
|
193 (** ** Translation Correctness *)
|
adamc@4
|
194
|
adamc@11
|
195 (** 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
|
196
|
adamc@11
|
197 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
|
adamc@11
|
198 (* begin hide *)
|
adamc@11
|
199 Abort.
|
adamc@11
|
200 (* end hide *)
|
adamc@11
|
201
|
adamc@11
|
202 (** 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
|
203 *)
|
adamc@2
|
204
|
adamc@15
|
205 Lemma compileCorrect' : forall e p s, progDenote (compile e ++ p) s = progDenote p (expDenote e :: s).
|
adamc@11
|
206
|
adamc@11
|
207 (** 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
|
208
|
adamc@11
|
209 [[
|
adamc@11
|
210 1 subgoal
|
adamc@11
|
211
|
adamc@11
|
212 ============================
|
adamc@15
|
213 forall (e : exp) (p : list instr) (s : stack),
|
adamc@15
|
214 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
|
adamc@11
|
215 ]]
|
adamc@11
|
216
|
adamc@11
|
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 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
|
adamc@11
|
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 hypotheses, if we had any. Below the line is the conclusion, which, in general, is to be proved from the hypotheses.
|
adamc@11
|
220
|
adamc@11
|
221 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
|
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
|
adamc@11
|
228 [[
|
adamc@11
|
229 2 subgoals
|
adamc@11
|
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)
|
adamc@11
|
236 ]]
|
adamc@11
|
237 [[
|
adamc@11
|
238 subgoal 2 is:
|
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@11
|
242 ]]
|
adamc@11
|
243
|
adamc@11
|
244 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
|
245
|
adamc@11
|
246 We begin the first case with another very common tactic.
|
adamc@11
|
247 *)
|
adamc@11
|
248
|
adamc@4
|
249 intros.
|
adamc@11
|
250
|
adamc@11
|
251 (** The current subgoal changes to:
|
adamc@11
|
252 [[
|
adamc@11
|
253
|
adamc@11
|
254 n : nat
|
adamc@11
|
255 s : stack
|
adamc@11
|
256 p : list instr
|
adamc@11
|
257 ============================
|
adamc@11
|
258 progDenote (compile (Const n) ++ p) s =
|
adamc@11
|
259 progDenote p (expDenote (Const n) :: s)
|
adamc@11
|
260 ]]
|
adamc@11
|
261
|
adamc@11
|
262 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
|
adamc@11
|
263
|
adamc@11
|
264 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
|
265 *)
|
adamc@11
|
266
|
adamc@4
|
267 unfold compile.
|
adamc@11
|
268
|
adamc@11
|
269 (** [[
|
adamc@11
|
270
|
adamc@11
|
271 n : nat
|
adamc@11
|
272 s : stack
|
adamc@11
|
273 p : list instr
|
adamc@11
|
274 ============================
|
adamc@11
|
275 progDenote ((IConst n :: nil) ++ p) s =
|
adamc@11
|
276 progDenote p (expDenote (Const n) :: s)
|
adamc@11
|
277 ]] *)
|
adamc@11
|
278
|
adamc@4
|
279 unfold expDenote.
|
adamc@11
|
280
|
adamc@11
|
281 (** [[
|
adamc@11
|
282
|
adamc@11
|
283 n : nat
|
adamc@11
|
284 s : stack
|
adamc@11
|
285 p : list instr
|
adamc@11
|
286 ============================
|
adamc@11
|
287 progDenote ((IConst n :: nil) ++ p) s = progDenote p (n :: s)
|
adamc@11
|
288 ]]
|
adamc@11
|
289
|
adamc@11
|
290 We only need to unfold the first occurrence of [progDenote] to prove the goal: *)
|
adamc@11
|
291
|
adamc@11
|
292 unfold progDenote at 1.
|
adamc@11
|
293
|
adamc@11
|
294 (** [[
|
adamc@11
|
295
|
adamc@11
|
296 n : nat
|
adamc@11
|
297 s : stack
|
adamc@11
|
298 p : list instr
|
adamc@11
|
299 ============================
|
adamc@11
|
300 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
|
adamc@11
|
301 option stack :=
|
adamc@11
|
302 match p0 with
|
adamc@11
|
303 | nil => Some s0
|
adamc@11
|
304 | i :: p' =>
|
adamc@11
|
305 match instrDenote i s0 with
|
adamc@11
|
306 | Some s' => progDenote p' s'
|
adamc@11
|
307 | None => None (A:=stack)
|
adamc@11
|
308 end
|
adamc@11
|
309 end) ((IConst n :: nil) ++ p) s =
|
adamc@11
|
310 progDenote p (n :: s)
|
adamc@11
|
311 ]]
|
adamc@11
|
312
|
adamc@11
|
313 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
|
314 *)
|
adamc@11
|
315
|
adamc@4
|
316 simpl.
|
adamc@11
|
317
|
adamc@11
|
318 (** [[
|
adamc@11
|
319
|
adamc@11
|
320 n : nat
|
adamc@11
|
321 s : stack
|
adamc@11
|
322 p : list instr
|
adamc@11
|
323 ============================
|
adamc@11
|
324 (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@11
|
334 ]]
|
adamc@11
|
335
|
adamc@11
|
336 Now we can unexpand the definition of [progDenote]:
|
adamc@11
|
337 *)
|
adamc@11
|
338
|
adamc@11
|
339 fold progDenote.
|
adamc@11
|
340
|
adamc@11
|
341 (** [[
|
adamc@11
|
342
|
adamc@11
|
343 n : nat
|
adamc@11
|
344 s : stack
|
adamc@11
|
345 p : list instr
|
adamc@11
|
346 ============================
|
adamc@11
|
347 progDenote p (n :: s) = progDenote p (n :: s)
|
adamc@11
|
348 ]]
|
adamc@11
|
349
|
adamc@11
|
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:
|
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
|
adamc@11
|
359 b : binop
|
adamc@11
|
360 e1 : exp
|
adamc@11
|
361 IHe1 : forall (s : stack) (p : list instr),
|
adamc@11
|
362 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
|
adamc@11
|
363 e2 : exp
|
adamc@11
|
364 IHe2 : forall (s : stack) (p : list instr),
|
adamc@11
|
365 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
|
adamc@11
|
366 ============================
|
adamc@11
|
367 forall (s : stack) (p : list instr),
|
adamc@11
|
368 progDenote (compile (Binop b e1 e2) ++ p) s =
|
adamc@11
|
369 progDenote p (expDenote (Binop b e1 e2) :: s)
|
adamc@11
|
370 ]]
|
adamc@11
|
371
|
adamc@11
|
372 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
|
373
|
adamc@11
|
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. *)
|
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@11
|
382 (** Now we arrive at a point where the tactics we have seen so far are insufficient:
|
adamc@11
|
383
|
adamc@11
|
384 [[
|
adamc@11
|
385
|
adamc@11
|
386 b : binop
|
adamc@11
|
387 e1 : exp
|
adamc@11
|
388 IHe1 : forall (s : stack) (p : list instr),
|
adamc@11
|
389 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
|
adamc@11
|
390 e2 : exp
|
adamc@11
|
391 IHe2 : forall (s : stack) (p : list instr),
|
adamc@11
|
392 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
|
adamc@11
|
393 s : stack
|
adamc@11
|
394 p : list instr
|
adamc@11
|
395 ============================
|
adamc@11
|
396 progDenote ((compile e2 ++ compile e1 ++ IBinop b :: nil) ++ p) s =
|
adamc@11
|
397 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
|
adamc@11
|
398 ]]
|
adamc@11
|
399
|
adamc@11
|
400 What we need is the associative law of list concatenation, available as a theorem [app_ass] in the standard library. *)
|
adamc@11
|
401
|
adamc@11
|
402 Check app_ass.
|
adamc@11
|
403
|
adamc@11
|
404 (** [[
|
adamc@11
|
405
|
adamc@11
|
406 app_ass
|
adamc@11
|
407 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
|
adamc@11
|
408 ]]
|
adamc@11
|
409
|
adamc@11
|
410 We use it to perform a rewrite: *)
|
adamc@11
|
411
|
adamc@4
|
412 rewrite app_ass.
|
adamc@11
|
413
|
adamc@11
|
414 (** changing the conclusion to: [[
|
adamc@11
|
415
|
adamc@11
|
416 progDenote (compile e2 ++ (compile e1 ++ IBinop b :: nil) ++ p) s =
|
adamc@11
|
417 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
|
adamc@11
|
418 ]]
|
adamc@11
|
419
|
adamc@11
|
420 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
|
421
|
adamc@4
|
422 rewrite IHe2.
|
adamc@11
|
423
|
adamc@11
|
424 (** [[
|
adamc@11
|
425
|
adamc@11
|
426 progDenote ((compile e1 ++ IBinop b :: nil) ++ p) (expDenote e2 :: s) =
|
adamc@11
|
427 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
|
adamc@11
|
428 ]]
|
adamc@11
|
429
|
adamc@11
|
430 The same process lets us apply the remaining hypothesis. *)
|
adamc@11
|
431
|
adamc@4
|
432 rewrite app_ass.
|
adamc@4
|
433 rewrite IHe1.
|
adamc@11
|
434
|
adamc@11
|
435 (** [[
|
adamc@11
|
436
|
adamc@11
|
437 progDenote ((IBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
|
adamc@11
|
438 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
|
adamc@11
|
439 ]]
|
adamc@11
|
440
|
adamc@11
|
441 Now we can apply a similar sequence of tactics to that that ended the proof of the first case.
|
adamc@11
|
442 *)
|
adamc@11
|
443
|
adamc@11
|
444 unfold progDenote at 1.
|
adamc@4
|
445 simpl.
|
adamc@11
|
446 fold progDenote.
|
adamc@4
|
447 reflexivity.
|
adamc@11
|
448
|
adamc@11
|
449 (** And the proof is completed, as indicated by the message:
|
adamc@11
|
450
|
adamc@11
|
451 [[
|
adamc@11
|
452 Proof completed.
|
adamc@11
|
453
|
adamc@11
|
454 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
|
455 *)
|
adamc@11
|
456
|
adamc@4
|
457 Abort.
|
adamc@2
|
458
|
adamc@4
|
459 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
|
adamc@4
|
460 progDenote p (expDenote e :: s).
|
adamc@4
|
461 induction e; crush.
|
adamc@4
|
462 Qed.
|
adamc@2
|
463
|
adamc@11
|
464 (** 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
|
465
|
adamc@11
|
466 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
|
467
|
adamc@4
|
468 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
|
adamc@11
|
469 intros.
|
adamc@11
|
470
|
adamc@11
|
471 (** [[
|
adamc@11
|
472
|
adamc@11
|
473 e : exp
|
adamc@11
|
474 ============================
|
adamc@11
|
475 progDenote (compile e) nil = Some (expDenote e :: nil)
|
adamc@11
|
476 ]]
|
adamc@11
|
477
|
adamc@11
|
478 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
|
479
|
adamc@11
|
480 Check app_nil_end.
|
adamc@11
|
481
|
adamc@11
|
482 (** [[
|
adamc@11
|
483
|
adamc@11
|
484 app_nil_end
|
adamc@11
|
485 : forall (A : Type) (l : list A), l = l ++ nil
|
adamc@11
|
486 ]] *)
|
adamc@11
|
487
|
adamc@4
|
488 rewrite (app_nil_end (compile e)).
|
adamc@11
|
489
|
adamc@11
|
490 (** 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
|
491
|
adamc@11
|
492 [[
|
adamc@11
|
493
|
adamc@11
|
494 e : exp
|
adamc@11
|
495 ============================
|
adamc@11
|
496 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
|
adamc@11
|
497 ]]
|
adamc@11
|
498
|
adamc@11
|
499 Now we can apply the lemma. *)
|
adamc@11
|
500
|
adamc@4
|
501 rewrite compileCorrect'.
|
adamc@11
|
502
|
adamc@11
|
503 (** [[
|
adamc@11
|
504
|
adamc@11
|
505 e : exp
|
adamc@11
|
506 ============================
|
adamc@11
|
507 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
|
adamc@11
|
508 ]]
|
adamc@11
|
509
|
adamc@11
|
510 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
|
511
|
adamc@4
|
512 reflexivity.
|
adamc@4
|
513 Qed.
|
adamc@14
|
514
|
adamc@14
|
515
|
adamc@20
|
516 (** * Typed Expressions *)
|
adamc@14
|
517
|
adamc@14
|
518 (** 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
|
519
|
adamc@20
|
520 (** ** Source Language *)
|
adamc@14
|
521
|
adamc@15
|
522 (** We define a trivial language of types to classify our expressions: *)
|
adamc@15
|
523
|
adamc@14
|
524 Inductive type : Set := Nat | Bool.
|
adamc@14
|
525
|
adamc@15
|
526 (** Now we define an expanded set of binary operators. *)
|
adamc@15
|
527
|
adamc@14
|
528 Inductive tbinop : type -> type -> type -> Set :=
|
adamc@14
|
529 | TPlus : tbinop Nat Nat Nat
|
adamc@14
|
530 | TTimes : tbinop Nat Nat Nat
|
adamc@14
|
531 | TEq : forall t, tbinop t t Bool
|
adamc@14
|
532 | TLt : tbinop Nat Nat Bool.
|
adamc@14
|
533
|
adamc@15
|
534 (** 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
|
535
|
adamc@15
|
536 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
|
537
|
adamc@15
|
538 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
|
539
|
adamc@15
|
540 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
|
541 *)
|
adamc@15
|
542
|
adamc@15
|
543 (** We can define a similar type family for typed expressions. *)
|
adamc@15
|
544
|
adamc@14
|
545 Inductive texp : type -> Set :=
|
adamc@14
|
546 | TNConst : nat -> texp Nat
|
adamc@14
|
547 | TBConst : bool -> texp Bool
|
adamc@14
|
548 | TBinop : forall arg1 arg2 res, tbinop arg1 arg2 res -> texp arg1 -> texp arg2 -> texp res.
|
adamc@14
|
549
|
adamc@15
|
550 (** 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
|
551
|
adamc@14
|
552 Definition typeDenote (t : type) : Set :=
|
adamc@14
|
553 match t with
|
adamc@14
|
554 | Nat => nat
|
adamc@14
|
555 | Bool => bool
|
adamc@14
|
556 end.
|
adamc@14
|
557
|
adamc@15
|
558 (** 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
|
559
|
adamc@15
|
560 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
|
561 *)
|
adamc@15
|
562
|
adamc@14
|
563 Definition eq_bool (b1 b2 : bool) : bool :=
|
adamc@14
|
564 match b1, b2 with
|
adamc@14
|
565 | true, true => true
|
adamc@14
|
566 | false, false => true
|
adamc@14
|
567 | _, _ => false
|
adamc@14
|
568 end.
|
adamc@14
|
569
|
adamc@14
|
570 Fixpoint eq_nat (n1 n2 : nat) {struct n1} : bool :=
|
adamc@14
|
571 match n1, n2 with
|
adamc@14
|
572 | O, O => true
|
adamc@14
|
573 | S n1', S n2' => eq_nat n1' n2'
|
adamc@14
|
574 | _, _ => false
|
adamc@14
|
575 end.
|
adamc@14
|
576
|
adamc@14
|
577 Fixpoint lt (n1 n2 : nat) {struct n1} : bool :=
|
adamc@14
|
578 match n1, n2 with
|
adamc@14
|
579 | O, S _ => true
|
adamc@14
|
580 | S n1', S n2' => lt n1' n2'
|
adamc@14
|
581 | _, _ => false
|
adamc@14
|
582 end.
|
adamc@14
|
583
|
adamc@15
|
584 (** Now we can interpret binary operators: *)
|
adamc@15
|
585
|
adamc@14
|
586 Definition tbinopDenote arg1 arg2 res (b : tbinop arg1 arg2 res)
|
adamc@14
|
587 : typeDenote arg1 -> typeDenote arg2 -> typeDenote res :=
|
adamc@14
|
588 match b in (tbinop arg1 arg2 res) return (typeDenote arg1 -> typeDenote arg2 -> typeDenote res) with
|
adamc@14
|
589 | TPlus => plus
|
adamc@14
|
590 | TTimes => mult
|
adamc@14
|
591 | TEq Nat => eq_nat
|
adamc@14
|
592 | TEq Bool => eq_bool
|
adamc@14
|
593 | TLt => lt
|
adamc@14
|
594 end.
|
adamc@14
|
595
|
adamc@15
|
596 (** 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
|
597
|
adamc@15
|
598 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
|
599
|
adamc@15
|
600 The same tricks suffice to define an expression denotation function in an unsurprising way:
|
adamc@15
|
601 *)
|
adamc@15
|
602
|
adamc@14
|
603 Fixpoint texpDenote t (e : texp t) {struct e} : typeDenote t :=
|
adamc@14
|
604 match e in (texp t) return (typeDenote t) with
|
adamc@14
|
605 | TNConst n => n
|
adamc@14
|
606 | TBConst b => b
|
adamc@14
|
607 | TBinop _ _ _ b e1 e2 => (tbinopDenote b) (texpDenote e1) (texpDenote e2)
|
adamc@14
|
608 end.
|
adamc@14
|
609
|
adamc@17
|
610 (** We can evaluate a few example programs to convince ourselves that this semantics is correct. *)
|
adamc@17
|
611
|
adamc@17
|
612 Eval simpl in texpDenote (TNConst 42).
|
adamc@18
|
613 (** [[
|
adamc@18
|
614 = 42 : typeDenote Nat
|
adamc@18
|
615 ]] *)
|
adamc@17
|
616 Eval simpl in texpDenote (TBConst true).
|
adamc@18
|
617 (** [[
|
adamc@18
|
618 = true : typeDenote Bool
|
adamc@18
|
619 ]] *)
|
adamc@17
|
620 Eval simpl in texpDenote (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
|
adamc@18
|
621 (** [[
|
adamc@18
|
622 = 28 : typeDenote Nat
|
adamc@18
|
623 ]] *)
|
adamc@17
|
624 Eval simpl in texpDenote (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
|
adamc@18
|
625 (** [[
|
adamc@18
|
626 = false : typeDenote Bool
|
adamc@18
|
627 ]] *)
|
adamc@17
|
628 Eval simpl in texpDenote (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)).
|
adamc@18
|
629 (** [[
|
adamc@18
|
630 = true : typeDenote Bool
|
adamc@18
|
631 ]] *)
|
adamc@17
|
632
|
adamc@14
|
633
|
adamc@20
|
634 (** ** Target Language *)
|
adamc@14
|
635
|
adamc@18
|
636 (** Now we want to define a suitable stack machine target for compilation. In the example of the untyped language, stack machine programs could encounter stack underflows and "get stuck." This was unfortunate, since we had to deal with this complication even though we proved that our compiler never produced underflowing programs. We could have used dependent types to force all stack machine programs to be underflow-free.
|
adamc@18
|
637
|
adamc@18
|
638 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
|
639
|
adamc@18
|
640 We start by defining stack types, which classify sets of possible stacks. *)
|
adamc@18
|
641
|
adamc@14
|
642 Definition tstack := list type.
|
adamc@14
|
643
|
adamc@18
|
644 (** 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
|
645
|
adamc@18
|
646 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
|
647
|
adamc@14
|
648 Inductive tinstr : tstack -> tstack -> Set :=
|
adamc@14
|
649 | TINConst : forall s, nat -> tinstr s (Nat :: s)
|
adamc@14
|
650 | TIBConst : forall s, bool -> tinstr s (Bool :: s)
|
adamc@14
|
651 | TIBinop : forall arg1 arg2 res s,
|
adamc@14
|
652 tbinop arg1 arg2 res
|
adamc@14
|
653 -> tinstr (arg1 :: arg2 :: s) (res :: s).
|
adamc@14
|
654
|
adamc@18
|
655 (** 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
|
656
|
adamc@14
|
657 Inductive tprog : tstack -> tstack -> Set :=
|
adamc@14
|
658 | TNil : forall s, tprog s s
|
adamc@14
|
659 | TCons : forall s1 s2 s3,
|
adamc@14
|
660 tinstr s1 s2
|
adamc@14
|
661 -> tprog s2 s3
|
adamc@14
|
662 -> tprog s1 s3.
|
adamc@14
|
663
|
adamc@18
|
664 (** 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
|
665
|
adamc@14
|
666 Fixpoint vstack (ts : tstack) : Set :=
|
adamc@14
|
667 match ts with
|
adamc@14
|
668 | nil => unit
|
adamc@14
|
669 | t :: ts' => typeDenote t * vstack ts'
|
adamc@14
|
670 end%type.
|
adamc@14
|
671
|
adamc@18
|
672 (** 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.
|
adamc@18
|
673
|
adamc@18
|
674 This idea of programming with types can take a while to internalize, but it enables a very simple definition of instruction denotation. We have the same kind of type annotations for the dependent [match], but everything else is like what you might expect from a Lisp-like version of ML that ignored type information. Nonetheless, the fact that [tinstrDenote] passes the type-checker guarantees that our stack machine programs can never go wrong. *)
|
adamc@18
|
675
|
adamc@14
|
676 Definition tinstrDenote ts ts' (i : tinstr ts ts') : vstack ts -> vstack ts' :=
|
adamc@14
|
677 match i in (tinstr ts ts') return (vstack ts -> vstack ts') with
|
adamc@14
|
678 | TINConst _ n => fun s => (n, s)
|
adamc@14
|
679 | TIBConst _ b => fun s => (b, s)
|
adamc@14
|
680 | TIBinop _ _ _ _ b => fun s =>
|
adamc@14
|
681 match s with
|
adamc@14
|
682 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
|
adamc@14
|
683 end
|
adamc@14
|
684 end.
|
adamc@14
|
685
|
adamc@18
|
686 (** 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
|
687
|
adamc@18
|
688 [[
|
adamc@18
|
689 Definition tinstrDenote ts ts' (i : tinstr ts ts') (s : vstack ts) : vstack ts' :=
|
adamc@18
|
690 match i in (tinstr ts ts') return (vstack ts') with
|
adamc@18
|
691 | TINConst _ n => (n, s)
|
adamc@18
|
692 | TIBConst _ b => (b, s)
|
adamc@18
|
693 | TIBinop _ _ _ _ b =>
|
adamc@18
|
694 match s with
|
adamc@18
|
695 (arg1, (arg2, s')) => ((tbinopDenote b) arg1 arg2, s')
|
adamc@18
|
696 end
|
adamc@18
|
697 end.
|
adamc@18
|
698
|
adamc@18
|
699 The Coq type-checker complains that:
|
adamc@18
|
700
|
adamc@18
|
701 [[
|
adamc@18
|
702 The term "(n, s)" has type "(nat * vstack ts)%type"
|
adamc@18
|
703 while it is expected to have type "vstack (Nat :: t)"
|
adamc@18
|
704 ]]
|
adamc@18
|
705
|
adamc@18
|
706 Recall from our earlier discussion of [match] annotations that we write the annotations to express to the type-checker the relationship between the type indices of the case object and the result type of the [match]. Coq chooses to assign to the wildcard [_] after [TINConst] the name [t], and the type error is telling us that the type checker cannot prove that [t] is the same as [ts]. By moving [s] out of the [match], we lose the ability to express, with [in] and [return] clauses, the relationship between the shared index [ts] of [s] and [i].
|
adamc@18
|
707
|
adamc@18
|
708 There %\textit{%#<i>#are#</i>#%}% reasonably general ways of getting around this problem without pushing binders inside [match]es. However, the alternatives are significantly more involved, and the technique we use here is almost certainly the best choice, whenever it applies.
|
adamc@18
|
709
|
adamc@18
|
710 *)
|
adamc@18
|
711
|
adamc@18
|
712 (** We finish the semantics with a straightforward definition of program denotation. *)
|
adamc@18
|
713
|
adamc@14
|
714 Fixpoint tprogDenote ts ts' (p : tprog ts ts') {struct p} : vstack ts -> vstack ts' :=
|
adamc@14
|
715 match p in (tprog ts ts') return (vstack ts -> vstack ts') with
|
adamc@14
|
716 | TNil _ => fun s => s
|
adamc@14
|
717 | TCons _ _ _ i p' => fun s => tprogDenote p' (tinstrDenote i s)
|
adamc@14
|
718 end.
|
adamc@14
|
719
|
adamc@14
|
720
|
adamc@14
|
721 (** ** Translation *)
|
adamc@14
|
722
|
adamc@19
|
723 (** To define our compilation, it is useful to have an auxiliary function for concatenating two stack machine programs. *)
|
adamc@19
|
724
|
adamc@14
|
725 Fixpoint tconcat ts ts' ts'' (p : tprog ts ts') {struct p} : tprog ts' ts'' -> tprog ts ts'' :=
|
adamc@14
|
726 match p in (tprog ts ts') return (tprog ts' ts'' -> tprog ts ts'') with
|
adamc@14
|
727 | TNil _ => fun p' => p'
|
adamc@14
|
728 | TCons _ _ _ i p1 => fun p' => TCons i (tconcat p1 p')
|
adamc@14
|
729 end.
|
adamc@14
|
730
|
adamc@19
|
731 (** With that function in place, the compilation is defined very similarly to how it was before, modulo the use of dependent typing. *)
|
adamc@19
|
732
|
adamc@14
|
733 Fixpoint tcompile t (e : texp t) (ts : tstack) {struct e} : tprog ts (t :: ts) :=
|
adamc@14
|
734 match e in (texp t) return (tprog ts (t :: ts)) with
|
adamc@14
|
735 | TNConst n => TCons (TINConst _ n) (TNil _)
|
adamc@14
|
736 | TBConst b => TCons (TIBConst _ b) (TNil _)
|
adamc@14
|
737 | TBinop _ _ _ b e1 e2 => tconcat (tcompile e2 _)
|
adamc@14
|
738 (tconcat (tcompile e1 _) (TCons (TIBinop _ b) (TNil _)))
|
adamc@14
|
739 end.
|
adamc@14
|
740
|
adamc@19
|
741 (** 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 arbitary terms, by writing underscores in place of specific values. You may have noticed that we have been calling functions without specifying all of their arguments. For instance, the recursive calls here to [tcompile] omit the [t] argument. Coq's %\textit{%#<i>#implicit argument#</i>#%}% mechanism automatically inserts underscores for arguments that it will probably be able to infer. Inference of such values is far from complete, though; generally, it only works in cases similar to those encountered with polymorphic type instantiation in Haskell and ML.
|
adamc@19
|
742
|
adamc@19
|
743 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
|
744
|
adamc@14
|
745 Print tcompile.
|
adamc@14
|
746
|
adamc@19
|
747 (** [[
|
adamc@19
|
748
|
adamc@19
|
749 tcompile =
|
adamc@19
|
750 fix tcompile (t : type) (e : texp t) (ts : tstack) {struct e} :
|
adamc@19
|
751 tprog ts (t :: ts) :=
|
adamc@19
|
752 match e in (texp t0) return (tprog ts (t0 :: ts)) with
|
adamc@19
|
753 | TNConst n => TCons (TINConst ts n) (TNil (Nat :: ts))
|
adamc@19
|
754 | TBConst b => TCons (TIBConst ts b) (TNil (Bool :: ts))
|
adamc@19
|
755 | TBinop arg1 arg2 res b e1 e2 =>
|
adamc@19
|
756 tconcat (tcompile arg2 e2 ts)
|
adamc@19
|
757 (tconcat (tcompile arg1 e1 (arg2 :: ts))
|
adamc@19
|
758 (TCons (TIBinop ts b) (TNil (res :: ts))))
|
adamc@19
|
759 end
|
adamc@19
|
760 : forall t : type, texp t -> forall ts : tstack, tprog ts (t :: ts)
|
adamc@19
|
761 ]] *)
|
adamc@19
|
762
|
adamc@19
|
763
|
adamc@19
|
764 (** We can check that the compiler generates programs that behave appropriately on our sample programs from above: *)
|
adamc@19
|
765
|
adamc@19
|
766 Eval simpl in tprogDenote (tcompile (TNConst 42) nil) tt.
|
adamc@19
|
767 (** [[
|
adamc@19
|
768 = (42, tt) : vstack (Nat :: nil)
|
adamc@19
|
769 ]] *)
|
adamc@19
|
770 Eval simpl in tprogDenote (tcompile (TBConst true) nil) tt.
|
adamc@19
|
771 (** [[
|
adamc@19
|
772 = (true, tt) : vstack (Bool :: nil)
|
adamc@19
|
773 ]] *)
|
adamc@19
|
774 Eval simpl in tprogDenote (tcompile (TBinop TTimes (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
|
adamc@19
|
775 (** [[
|
adamc@19
|
776 = (28, tt) : vstack (Nat :: nil)
|
adamc@19
|
777 ]] *)
|
adamc@19
|
778 Eval simpl in tprogDenote (tcompile (TBinop (TEq Nat) (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
|
adamc@19
|
779 (** [[
|
adamc@19
|
780 = (false, tt) : vstack (Bool :: nil)
|
adamc@19
|
781 ]] *)
|
adamc@19
|
782 Eval simpl in tprogDenote (tcompile (TBinop TLt (TBinop TPlus (TNConst 2) (TNConst 2)) (TNConst 7)) nil) tt.
|
adamc@19
|
783 (** [[
|
adamc@19
|
784 = (true, tt) : vstack (Bool :: nil)
|
adamc@19
|
785 ]] *)
|
adamc@19
|
786
|
adamc@14
|
787
|
adamc@20
|
788 (** ** Translation Correctness *)
|
adamc@20
|
789
|
adamc@20
|
790 (** We can state a correctness theorem similar to the last one. *)
|
adamc@20
|
791
|
adamc@20
|
792 Theorem tcompileCorrect : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
|
adamc@20
|
793 (* begin hide *)
|
adamc@20
|
794 Abort.
|
adamc@20
|
795 (* end hide *)
|
adamc@20
|
796
|
adamc@20
|
797 (** Again, we need to strengthen the theorem statement so that the induction will go through. This time, I will develop an alternative approach to this kind of proof, stating the key lemma as: *)
|
adamc@14
|
798
|
adamc@14
|
799 Lemma tcompileCorrect' : forall t (e : texp t)
|
adamc@14
|
800 ts (s : vstack ts),
|
adamc@14
|
801 tprogDenote (tcompile e ts) s
|
adamc@14
|
802 = (texpDenote e, s).
|
adamc@20
|
803
|
adamc@20
|
804 (** While lemma [compileCorrect'] quantified over a program that is the "continuation" for the expression we are considering, here we avoid drawing in any extra syntactic elements. In addition to the source expression and its type, we also quantify over an initial stack type and a stack compatible with it. Running the compilation of the program starting from that stack, we should arrive at a stack that differs only in having the program's denotation pushed onto it.
|
adamc@20
|
805
|
adamc@20
|
806 Let us try to prove this theorem in the same way that we settled on in the last section. *)
|
adamc@20
|
807
|
adamc@14
|
808 induction e; crush.
|
adamc@20
|
809
|
adamc@20
|
810 (** We are left with this unproved conclusion:
|
adamc@20
|
811
|
adamc@20
|
812 [[
|
adamc@20
|
813
|
adamc@20
|
814 tprogDenote
|
adamc@20
|
815 (tconcat (tcompile e2 ts)
|
adamc@20
|
816 (tconcat (tcompile e1 (arg2 :: ts))
|
adamc@20
|
817 (TCons (TIBinop ts t) (TNil (res :: ts))))) s =
|
adamc@20
|
818 (tbinopDenote t (texpDenote e1) (texpDenote e2), s)
|
adamc@20
|
819 ]]
|
adamc@20
|
820
|
adamc@20
|
821 We need an analogue to the [app_ass] theorem that we used to rewrite the goal in the last section. We can abort this proof and prove such a lemma about [tconcat].
|
adamc@20
|
822 *)
|
adamc@14
|
823 Abort.
|
adamc@14
|
824
|
adamc@14
|
825 Lemma tconcatCorrect : forall ts ts' ts'' (p : tprog ts ts') (p' : tprog ts' ts'')
|
adamc@14
|
826 (s : vstack ts),
|
adamc@14
|
827 tprogDenote (tconcat p p') s
|
adamc@14
|
828 = tprogDenote p' (tprogDenote p s).
|
adamc@14
|
829 induction p; crush.
|
adamc@14
|
830 Qed.
|
adamc@14
|
831
|
adamc@20
|
832 (** This one goes through completely automatically.
|
adamc@20
|
833
|
adamc@20
|
834 Some code behind the scenes registers [app_ass] for use by [crush]. We must register [tconcatCorrect] similarly to get the same effect: *)
|
adamc@20
|
835
|
adamc@14
|
836 Hint Rewrite tconcatCorrect : cpdt.
|
adamc@14
|
837
|
adamc@20
|
838 (** We ask that the lemma be used for left-to-right rewriting, and we ask for the hint to be added to the hint database called [cpdt], which is the database used by [crush]. Now we are ready to return to [tcompileCorrect'], proving it automatically this time. *)
|
adamc@20
|
839
|
adamc@14
|
840 Lemma tcompileCorrect' : forall t (e : texp t)
|
adamc@14
|
841 ts (s : vstack ts),
|
adamc@14
|
842 tprogDenote (tcompile e ts) s
|
adamc@14
|
843 = (texpDenote e, s).
|
adamc@14
|
844 induction e; crush.
|
adamc@14
|
845 Qed.
|
adamc@14
|
846
|
adamc@20
|
847 (** We can register this main lemma as another hint, allowing us to prove the final theorem trivially. *)
|
adamc@20
|
848
|
adamc@14
|
849 Hint Rewrite tcompileCorrect' : cpdt.
|
adamc@14
|
850
|
adamc@14
|
851 Theorem tcompileCorrect : forall t (e : texp t), tprogDenote (tcompile e nil) tt = (texpDenote e, tt).
|
adamc@14
|
852 crush.
|
adamc@14
|
853 Qed.
|