comparison book/src/StackMachine.v @ 11:e5501b9c965d

Nats StackMachine done
author Adam Chlipala <adamc@hcoop.net>
date Mon, 01 Sep 2008 12:52:30 -0400
parents d8363f2a3bef
children
comparison
equal deleted inserted replaced
10:d8363f2a3bef 11:e5501b9c965d
14 (* end hide *) 14 (* end hide *)
15 15
16 16
17 (** 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. 17 (** 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.
18 18
19 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. *) 19 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.
20
21 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. *)
20 22
21 23
22 (** * Arithmetic expressions over natural numbers *) 24 (** * Arithmetic expressions over natural numbers *)
23 25
24 (** We will begin with that staple of compiler textbooks, arithemtic expressions over a single type of numbers. *) 26 (** We will begin with that staple of compiler textbooks, arithemtic expressions over a single type of numbers. *)
141 end. 143 end.
142 144
143 145
144 (** ** Translation correctness *) 146 (** ** Translation correctness *)
145 147
148 (** 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: *)
149
150 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
151 (* begin hide *)
152 Abort.
153 (* end hide *)
154
155 (** 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:
156 *)
157
146 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s = 158 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
147 progDenote p (expDenote e :: s). 159 progDenote p (expDenote e :: s).
160
161 (** 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:
162
163 [[
164 1 subgoal
165
166 ============================
167 forall (e : exp) (s : stack) (p : list instr),
168 progDenote (compile e ++ p) s = progDenote p (expDenote e :: s)
169 ]]
170
171 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.
172
173 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.
174
175 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:
176 *)
177
148 induction e. 178 induction e.
149 179
180 (** 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:
181
182 [[
183 2 subgoals
184
185 n : nat
186 ============================
187 forall (s : stack) (p : list instr),
188 progDenote (compile (Const n) ++ p) s =
189 progDenote p (expDenote (Const n) :: s)
190 ]]
191 [[
192 subgoal 2 is:
193 forall (s : stack) (p : list instr),
194 progDenote (compile (Binop b e1 e2) ++ p) s =
195 progDenote p (expDenote (Binop b e1 e2) :: s)
196 ]]
197
198 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.
199
200 We begin the first case with another very common tactic.
201 *)
202
150 intros. 203 intros.
204
205 (** The current subgoal changes to:
206 [[
207
208 n : nat
209 s : stack
210 p : list instr
211 ============================
212 progDenote (compile (Const n) ++ p) s =
213 progDenote p (expDenote (Const n) :: s)
214 ]]
215
216 We see that [intros] changes [forall]-bound variables at the beginning of a goal into free variables.
217
218 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.
219 *)
220
151 unfold compile. 221 unfold compile.
222
223 (** [[
224
225 n : nat
226 s : stack
227 p : list instr
228 ============================
229 progDenote ((IConst n :: nil) ++ p) s =
230 progDenote p (expDenote (Const n) :: s)
231 ]] *)
232
152 unfold expDenote. 233 unfold expDenote.
234
235 (** [[
236
237 n : nat
238 s : stack
239 p : list instr
240 ============================
241 progDenote ((IConst n :: nil) ++ p) s = progDenote p (n :: s)
242 ]]
243
244 We only need to unfold the first occurrence of [progDenote] to prove the goal: *)
245
246 unfold progDenote at 1.
247
248 (** [[
249
250 n : nat
251 s : stack
252 p : list instr
253 ============================
254 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
255 option stack :=
256 match p0 with
257 | nil => Some s0
258 | i :: p' =>
259 match instrDenote i s0 with
260 | Some s' => progDenote p' s'
261 | None => None (A:=stack)
262 end
263 end) ((IConst n :: nil) ++ p) s =
264 progDenote p (n :: s)
265 ]]
266
267 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:
268 *)
269
153 simpl. 270 simpl.
271
272 (** [[
273
274 n : nat
275 s : stack
276 p : list instr
277 ============================
278 (fix progDenote (p0 : prog) (s0 : stack) {struct p0} :
279 option stack :=
280 match p0 with
281 | nil => Some s0
282 | i :: p' =>
283 match instrDenote i s0 with
284 | Some s' => progDenote p' s'
285 | None => None (A:=stack)
286 end
287 end) p (n :: s) = progDenote p (n :: s)
288 ]]
289
290 Now we can unexpand the definition of [progDenote]:
291 *)
292
293 fold progDenote.
294
295 (** [[
296
297 n : nat
298 s : stack
299 p : list instr
300 ============================
301 progDenote p (n :: s) = progDenote p (n :: s)
302 ]]
303
304 It looks like we are at the end of this case, since we have a trivial equality. Indeed, a single tactic finishes us off:
305 *)
306
154 reflexivity. 307 reflexivity.
308
309 (** On to the second inductive case:
310
311 [[
312
313 b : binop
314 e1 : exp
315 IHe1 : forall (s : stack) (p : list instr),
316 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
317 e2 : exp
318 IHe2 : forall (s : stack) (p : list instr),
319 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
320 ============================
321 forall (s : stack) (p : list instr),
322 progDenote (compile (Binop b e1 e2) ++ p) s =
323 progDenote p (expDenote (Binop b e1 e2) :: s)
324 ]]
325
326 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.
327
328 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. *)
155 329
156 intros. 330 intros.
157 unfold compile. 331 unfold compile.
158 fold compile. 332 fold compile.
159 unfold expDenote. 333 unfold expDenote.
160 fold expDenote. 334 fold expDenote.
335
336 (** Now we arrive at a point where the tactics we have seen so far are insufficient:
337
338 [[
339
340 b : binop
341 e1 : exp
342 IHe1 : forall (s : stack) (p : list instr),
343 progDenote (compile e1 ++ p) s = progDenote p (expDenote e1 :: s)
344 e2 : exp
345 IHe2 : forall (s : stack) (p : list instr),
346 progDenote (compile e2 ++ p) s = progDenote p (expDenote e2 :: s)
347 s : stack
348 p : list instr
349 ============================
350 progDenote ((compile e2 ++ compile e1 ++ IBinop b :: nil) ++ p) s =
351 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
352 ]]
353
354 What we need is the associative law of list concatenation, available as a theorem [app_ass] in the standard library. *)
355
356 Check app_ass.
357
358 (** [[
359
360 app_ass
361 : forall (A : Type) (l m n : list A), (l ++ m) ++ n = l ++ m ++ n
362 ]]
363
364 We use it to perform a rewrite: *)
365
161 rewrite app_ass. 366 rewrite app_ass.
367
368 (** changing the conclusion to: [[
369
370 progDenote (compile e2 ++ (compile e1 ++ IBinop b :: nil) ++ p) s =
371 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
372 ]]
373
374 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: *)
375
162 rewrite IHe2. 376 rewrite IHe2.
377
378 (** [[
379
380 progDenote ((compile e1 ++ IBinop b :: nil) ++ p) (expDenote e2 :: s) =
381 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
382 ]]
383
384 The same process lets us apply the remaining hypothesis. *)
385
163 rewrite app_ass. 386 rewrite app_ass.
164 rewrite IHe1. 387 rewrite IHe1.
388
389 (** [[
390
391 progDenote ((IBinop b :: nil) ++ p) (expDenote e1 :: expDenote e2 :: s) =
392 progDenote p (binopDenote b (expDenote e1) (expDenote e2) :: s)
393 ]]
394
395 Now we can apply a similar sequence of tactics to that that ended the proof of the first case.
396 *)
397
398 unfold progDenote at 1.
165 simpl. 399 simpl.
400 fold progDenote.
166 reflexivity. 401 reflexivity.
402
403 (** And the proof is completed, as indicated by the message:
404
405 [[
406 Proof completed.
407
408 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.
409 *)
410
167 Abort. 411 Abort.
168 412
169 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s = 413 Lemma compileCorrect' : forall e s p, progDenote (compile e ++ p) s =
170 progDenote p (expDenote e :: s). 414 progDenote p (expDenote e :: s).
171 induction e; crush. 415 induction e; crush.
172 Qed. 416 Qed.
173 417
418 (** 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.
419
420 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. *)
421
174 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil). 422 Theorem compileCorrect : forall e, progDenote (compile e) nil = Some (expDenote e :: nil).
175 intro. 423 intros.
424
425 (** [[
426
427 e : exp
428 ============================
429 progDenote (compile e) nil = Some (expDenote e :: nil)
430 ]]
431
432 At this point, we want to massage the lefthand side to match the statement of [compileCorrect']. A theorem from the standard library is useful: *)
433
434 Check app_nil_end.
435
436 (** [[
437
438 app_nil_end
439 : forall (A : Type) (l : list A), l = l ++ nil
440 ]] *)
441
176 rewrite (app_nil_end (compile e)). 442 rewrite (app_nil_end (compile e)).
443
444 (** 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.
445
446 [[
447
448 e : exp
449 ============================
450 progDenote (compile e ++ nil) nil = Some (expDenote e :: nil)
451 ]]
452
453 Now we can apply the lemma. *)
454
177 rewrite compileCorrect'. 455 rewrite compileCorrect'.
456
457 (** [[
458
459 e : exp
460 ============================
461 progDenote nil (expDenote e :: nil) = Some (expDenote e :: nil)
462 ]]
463
464 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. *)
465
178 reflexivity. 466 reflexivity.
179 Qed. 467 Qed.