comparison src/DeBruijn.v @ 265:dce88a5c170c

Import DeBruijn
author Adam Chlipala <adamc@hcoop.net>
date Wed, 30 Dec 2009 13:38:13 -0500
parents
children da9ccc6bf572
comparison
equal deleted inserted replaced
264:aaf532c80729 265:dce88a5c170c
1 (* Copyright (c) 2009, Adam Chlipala
2 *
3 * This work is licensed under a
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
5 * Unported License.
6 * The license text is available at:
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
8 *)
9
10 (* begin hide *)
11 Set Implicit Arguments.
12
13 Require Import Eqdep.
14
15 Require Import DepList Tactics.
16 (* end hide *)
17
18 (** remove printing * *)
19 (** printing || $\mid\mid$ *)
20
21 (** %\chapter{Dependent De Bruijn Indices}% *)
22
23 (** The previous chapter introduced the most common form of de Bruijn indices, without essential use of dependent types. In earlier chapters, we used dependent de Bruijn indices to illustrate tricks for working with dependent types. This chapter presents one complete case study with dependent de Bruijn indices, focusing on producing the most maintainable proof possible of a classic theorem about lambda calculus.
24
25 The proof that follows does not provide a complete guide to all kinds of formalization with de Bruijn indices. Rather, it is intended as an example of some simple design patterns that can make proving standard theorems much easier.
26
27 We will prove commutativity of capture-avoiding substitution for basic untyped lambda calculus:
28
29 %$$x_1 \neq x_2 \Rightarrow [e_1/x_1][e_2/x_2]e = [e_2/x_2][[e_2/x_2]e_1/x_1]e$$%
30 #<tt>x1 <> x2 => [e1/x1][e2/x2]e = [e2/x2][[e2/x2]e1/x1]e</tt>#
31 *)
32
33 (** * Defining Syntax and Its Associated Operations *)
34
35 (** Our definition of expression syntax should be unsurprising. An expression of type [exp n] may refer to [n] different free variables. *)
36
37 Inductive exp : nat -> Type :=
38 | Var : forall n, fin n -> exp n
39 | App : forall n, exp n -> exp n -> exp n
40 | Abs : forall n, exp (S n) -> exp n.
41
42 (** The classic implementation of substitution in de Bruijn terms requires an auxiliary operation, %\emph{%lifting%}%, which increments the indices of all free variables in an expression. We need to lift whenever we %``%#"#go under a binder.#"#%''% It is useful to write an auxiliary function [liftVar] that lifts a variable; that is, [liftVar x y] will return [y + 1] if [y >= x], and it will return [y] otherwise. This simple description uses numbers rather than our dependent [fin] family, so the actual specification is more involved.
43
44 Combining a number of dependent types tricks, we wind up with this concrete realization. *)
45
46 Fixpoint liftVar n (x : fin n) : fin (pred n) -> fin n :=
47 match x with
48 | First _ => fun y => Next y
49 | Next _ x' => fun y =>
50 match y in fin n' return fin n' -> (fin (pred n') -> fin n')
51 -> fin (S n') with
52 | First _ => fun x' _ => First
53 | Next _ y' => fun _ fx' => Next (fx' y')
54 end x' (liftVar x')
55 end.
56
57 (** Now it is easy to implement the main lifting operation. *)
58
59 Fixpoint lift n (e : exp n) : fin (S n) -> exp (S n) :=
60 match e with
61 | Var _ f' => fun f => Var (liftVar f f')
62 | App _ e1 e2 => fun f => App (lift e1 f) (lift e2 f)
63 | Abs _ e1 => fun f => Abs (lift e1 (Next f))
64 end.
65
66 (** To define substitution itself, we will need to apply some explicit type casts, based on equalities between types. A single equality will suffice for all of our casts. Its statement is somewhat strange: it quantifies over a variable [f] of type [fin n], but then never mentions [f]. Rather, quantifying over [f] is useful because [fin] is a dependent type that is inhabited or not depending on its index. The body of the theorem, [S (pred n) = n], is true only for [n] $> 0$, but we can prove it by contradiction when [n = 0], because we have around a value [f] of the uninhabited type [fin 0]. *)
67
68 Theorem nzf : forall n (f : fin n), S (pred n) = n.
69 destruct 1; trivial.
70 Qed.
71
72 (** Now we define a notation to streamline our cast expressions. The code [[f return n, r for e]] denotes a cast of expression [e] whose type can be obtained by substituting some number [n1] for [n] in [r]. [f] should be a proof that [n1 = n2], for any [n2]. In that case, the type of the cast expression is [r] with [n2] substituted for [n]. *)
73
74 Notation "[ f 'return' n , r 'for' e ]" :=
75 match f in _ = n return r with
76 | refl_equal => e
77 end.
78
79 (** This notation is useful in defining a variable substitution operation. The idea is that [substVar x y] returns [None] if [x = y]; otherwise, it returns a %``%#"#squished#"#%''% version of [y] with a smaller [fin] index, reflecting that variable [x] has been substituted away. Without dependent types, this would be a simple definition. With dependency, it is reasonably intricate, and our main task in automating proofs about it will be hiding that intricacy. *)
80
81 Fixpoint substVar n (x : fin n) : fin n -> option (fin (pred n)) :=
82 match x with
83 | First _ => fun y =>
84 match y in fin n' return option (fin (pred n')) with
85 | First _ => None
86 | Next _ f' => Some f'
87 end
88 | Next _ x' => fun y =>
89 match y in fin n'
90 return fin (pred n') -> (fin (pred n') -> option (fin (pred (pred n'))))
91 -> option (fin (pred n')) with
92 | First _ => fun x' _ => Some [nzf x' return n, fin n for First]
93 | Next _ y' => fun _ fx' =>
94 match fx' y' with
95 | None => None
96 | Some f => Some [nzf y' return n, fin n for Next f]
97 end
98 end x' (substVar x')
99 end.
100
101 (** It is now easy to define our final substitution function. The abstraction case involves two casts, where one uses the [sym_eq] function to convert a proof of [n1 = n2] into a proof of [n2 = n1]. *)
102
103 Fixpoint subst n (e : exp n) : fin n -> exp (pred n) -> exp (pred n) :=
104 match e with
105 | Var _ f' => fun f v => match substVar f f' with
106 | None => v
107 | Some f'' => Var f''
108 end
109 | App _ e1 e2 => fun f v => App (subst e1 f v) (subst e2 f v)
110 | Abs _ e1 => fun f v => Abs [sym_eq (nzf f) return n, exp n for
111 subst e1 (Next f) [nzf f return n, exp n for lift v First]]
112 end.
113
114 (** Our final commutativity theorem is about [subst], but our proofs will rely on a few more auxiliary definitions. First, we will want an operation [more] that increments the index of a [fin] while preserving its interpretation as a number. *)
115
116 Fixpoint more n (f : fin n) : fin (S n) :=
117 match f with
118 | First _ => First
119 | Next _ f' => Next (more f')
120 end.
121
122 (** Second, we will want a kind of inverse to [liftVar]. *)
123
124 Fixpoint unliftVar n (f : fin n) : fin (pred n) -> fin (pred n) :=
125 match f with
126 | First _ => fun g => [nzf g return n, fin n for First]
127 | Next _ f' => fun g =>
128 match g in fin n'
129 return fin n' -> (fin (pred n') -> fin (pred n')) -> fin n' with
130 | First _ => fun f' _ => f'
131 | Next _ g' => fun _ unlift => Next (unlift g')
132 end f' (unliftVar f')
133 end.
134
135
136 (** * Custom Tactics *)
137
138 (* begin hide *)
139 Ltac simp := repeat progress (crush; try discriminate;
140 repeat match goal with
141 | [ |- context[match ?pf with refl_equal => _ end] ] => rewrite (UIP_refl _ _ pf)
142 | [ _ : context[match ?pf with refl_equal => _ end] |- _ ] => rewrite (UIP_refl _ _ pf) in *
143
144 | [ H : Next _ = Next _ |- _ ] => injection H; clear H
145 | [ H : ?E = _, H' : ?E = _ |- _ ] => rewrite H in H'
146 | [ H : ?P |- _ ] => rewrite H in *; [match P with
147 | forall x, _ => clear H
148 | _ => idtac
149 end]
150 end).
151 (* end hide *)
152
153 (** Less than a page of tactic code will be sufficient to automate our proof of commuativity. We start by defining a workhorse simplification tactic [simp], which extends [crush] in a few ways.
154
155 [[
156 Ltac simp := repeat progress (crush; try discriminate;
157
158 ]]
159
160 We enter an inner loop of applying hints specific to our domain.
161
162 [[
163 repeat match goal with
164
165 ]]
166
167 Our first two hints find places where equality proofs are pattern-matched on. The first hint matches pattern-matches in the conclusion, while the second hint matches pattern-matches in hypotheses. In each case, we apply the library theorem [UIP_refl], which says that any proof of a fact like [e = e] is itself equal to [refl_equal]. Rewriting with this fact enables reduction of the pattern-match that we found.
168
169 [[
170 | [ |- context[match ?pf with refl_equal => _ end] ] =>
171 rewrite (UIP_refl _ _ pf)
172 | [ _ : context[match ?pf with refl_equal => _ end] |- _ ] =>
173 rewrite (UIP_refl _ _ pf) in *
174
175 ]]
176
177 The next hint finds an opportunity to invert a [fin] equality hypothesis.
178
179 [[
180 | [ H : Next _ = Next _ |- _ ] => injection H; clear H
181
182 ]]
183
184 If we have two equality hypotheses that share a lefthand side, we can use one to rewrite the other, bringing the hypotheses' righthand sides together in a single equation.
185
186 [[
187 | [ H : ?E = _, H' : ?E = _ |- _ ] => rewrite H in H'
188
189 ]]
190
191 Finally, we would like automatic use of quantified equality hypotheses to perform rewriting. We pattern-match a hypothesis [H] asserting proposition [P]. We try to use [H] to perform rewriting everywhere in our goal. The rewrite succeeds if it generates no additional hypotheses, and, to prevent infinite loops in proof search, we clear [H] if it begins with universal quantification.
192
193 [[
194 | [ H : ?P |- _ ] => rewrite H in *; [match P with
195 | forall x, _ => clear H
196 | _ => idtac
197 end]
198 end).
199
200 ]]
201
202 In implementing another level of automation, it will be useful to mark which free variables we generated with tactics, as opposed to which were present in the original theorem statement. We use a dummy marker predicate [Generated] to record that information. A tactic [not_generated] fails if and only if its argument is a generated variable, and a tactic [generate] records that its argument is generated. *)
203
204 Definition Generated n (_ : fin n) := True.
205
206 Ltac not_generated x :=
207 match goal with
208 | [ _ : Generated x |- _ ] => fail 1
209 | _ => idtac
210 end.
211
212 Ltac generate x := assert (Generated x); [ constructor | ].
213
214 (** A tactic [destructG] performs case analysis on [fin] values. The built-in case analysis tactics are not smart enough to handle all situations, and we also want to mark new variables as generated, to avoid infinite loops of case analysis. Our [destructG] tactic will only proceed if its argument is not generated. *)
215
216 Theorem fin_inv : forall n (f : fin (S n)), f = First \/ exists f', f = Next f'.
217 intros; dep_destruct f; eauto.
218 Qed.
219
220 Ltac destructG E :=
221 not_generated E; let x := fresh "x" in
222 (destruct (fin_inv E) as [ | [x]] || destruct E as [ | ? x]);
223 [ | generate x ].
224
225 (* begin hide *)
226 Ltac dester := simp;
227 repeat (match goal with
228 | [ x : fin _, H : _ = _, IH : forall f : fin _, _ |- _ ] =>
229 generalize (IH x); clear IH; intro IH; rewrite H in IH
230 | [ x : fin _, y : fin _, H : _ = _, IH : forall (f : fin _) (g : fin _), _ |- _ ] =>
231 generalize (IH x y); clear IH; intro IH; rewrite H in IH
232 | [ |- context[match ?E with First _ => _ | Next _ _ => _ end] ] => destructG E
233 | [ _ : context[match ?E with First _ => _ | Next _ _ => _ end] |- _ ] => destructG E
234 | [ |- context[more ?E] ] => destructG E
235 | [ x : fin ?n |- _ ] =>
236 match goal with
237 | [ |- context[nzf x] ] =>
238 destruct n; [ inversion x | ]
239 end
240 | [ x : fin (pred ?n), y : fin ?n |- _ ] =>
241 match goal with
242 | [ |- context[nzf x] ] =>
243 destruct n; [ inversion y | ]
244 end
245 | [ |- context[match ?E with None => _ | Some _ => _ end] ] =>
246 match E with
247 | match _ with None => _ | Some _ => _ end => fail 1
248 | _ => case_eq E; firstorder
249 end
250 end; simp); eauto.
251 (* end hide *)
252
253 (** Our most powerful workhorse tactic will be [dester], which incorporates all of [simp]'s simplifications and adds heuristics for automatic case analysis and automatic quantifier instantiation.
254
255 [[
256 Ltac dester := simp;
257 repeat (match goal with
258
259 ]]
260
261 The first hint expresses our main insight into quantifier instantiation. We identify a hypothesis [IH] that begins with quantification over [fin] values. We also identify a free [fin] variable [x] and an arbitrary equality hypothesis [H]. Given these, we try instantiating [IH] with [x]. We know we chose correctly if the instantiated proposition includes an opportunity to rewrite using [H].
262
263 [[
264 | [ x : fin _, H : _ = _, IH : forall f : fin _, _ |- _ ] =>
265 generalize (IH x); clear IH; intro IH; rewrite H in IH
266
267 ]]
268
269 This basic idea suffices for all of our explicit quantifier instantiation. We add one more variant that handles cases where an opportunity for rewriting is only exposed if two different quantifiers are instantiated at once.
270
271 [[
272 | [ x : fin _, y : fin _, H : _ = _,
273 IH : forall (f : fin _) (g : fin _), _ |- _ ] =>
274 generalize (IH x y); clear IH; intro IH; rewrite H in IH
275
276 ]]
277
278 We want to case-analyze on any [fin] expression that is the discriminee of a [match] expression or an argument to [more].
279
280 [[
281 | [ |- context[match ?E with First _ => _ | Next _ _ => _ end] ] =>
282 destructG E
283 | [ _ : context[match ?E with First _ => _ | Next _ _ => _ end] |- _ ] =>
284 destructG E
285 | [ |- context[more ?E] ] => destructG E
286
287 ]]
288
289 Recall that [simp] will simplify equality proof terms of facts like [e = e]. The proofs in question will either be of [n = S (pred n)] or [S (pred n) = n], for some [n]. These equations do not have syntactically equal sides. We can get to the point where they %\emph{%#<i>#do#</i>#%}% have equal sides by performing case analysis on [n]. Whenever we do so, the [n = 0] case will be contradictory, allowing us to discharge it by finding a free variable of type [fin 0] and performing inversion on it. In the [n = S n'] case, the sides of these equalities will simplify to equal values, as needed. The next two hints identify [n] values that are good candidates for such case analysis.
290
291 [[
292 | [ x : fin ?n |- _ ] =>
293 match goal with
294 | [ |- context[nzf x] ] =>
295 destruct n; [ inversion x | ]
296 end
297 | [ x : fin (pred ?n), y : fin ?n |- _ ] =>
298 match goal with
299 | [ |- context[nzf x] ] =>
300 destruct n; [ inversion y | ]
301 end
302
303 ]]
304
305 Finally, we find [match] discriminees of [option] type, enforcing that we do not destruct any discriminees that are themselves [match] expressions. Crucially, we do these case analyses with [case_eq] instead of [destruct]. The former adds equality hypotheses to record the relationships between old variables and their new deduced forms. These equalities will be used by our quantifier instantiation heuristic.
306
307 [[
308 | [ |- context[match ?E with None => _ | Some _ => _ end] ] =>
309 match E with
310 | match _ with None => _ | Some _ => _ end => fail 1
311 | _ => case_eq E; firstorder
312 end
313
314 ]]
315
316 Each iteration of the loop ends by calling [simp] again, and, after no more progress can be made, we finish by calling [eauto].
317
318 [[
319 end; simp); eauto.
320
321 ]] *)
322
323
324 (** * Theorems *)
325
326 (** We are now ready to prove our main theorem, by way of a progression of lemmas.
327
328 The first pair of lemmas characterizes the interaction of substitution and lifting at the variable level. *)
329
330 Lemma substVar_unliftVar : forall n (f0 : fin n) f g,
331 match substVar f0 f, substVar (liftVar f0 g) f with
332 | Some f1, Some f2 => exists f', substVar g f1 = Some f'
333 /\ substVar (unliftVar f0 g) f2 = Some f'
334 | Some f1, None => substVar g f1 = None
335 | None, Some f2 => substVar (unliftVar f0 g) f2 = None
336 | None, None => False
337 end.
338 induction f0; dester.
339 Qed.
340
341 Lemma substVar_liftVar : forall n (f0 : fin n) f,
342 substVar f0 (liftVar f0 f) = Some f.
343 induction f0; dester.
344 Qed.
345
346 (** Next, we define a notion of %``%#"#greater-than-or-equal#"#%''% for [fin] values, prove an inversion theorem for it, and add that theorem as a hint. *)
347
348 Inductive fin_ge : forall n1, fin n1 -> forall n2, fin n2 -> Prop :=
349 | GeO : forall n1 (f1 : fin n1) n2,
350 fin_ge f1 (First : fin (S n2))
351 | GeS : forall n1 (f1 : fin n1) n2 (f2 : fin n2),
352 fin_ge f1 f2
353 -> fin_ge (Next f1) (Next f2).
354
355 Hint Constructors fin_ge.
356
357 Lemma fin_ge_inv' : forall n1 n2 (f1 : fin n1) (f2 : fin n2),
358 fin_ge f1 f2
359 -> match f1, f2 with
360 | Next _ f1', Next _ f2' => fin_ge f1' f2'
361 | _, _ => True
362 end.
363 destruct 1; dester.
364 Qed.
365
366 Lemma fin_ge_inv : forall n1 n2 (f1 : fin n1) (f2 : fin n2),
367 fin_ge (Next f1) (Next f2)
368 -> fin_ge f1 f2.
369 intros; generalize (fin_ge_inv' (f1 := Next f1) (f2 := Next f2)); dester.
370 Qed.
371
372 Hint Resolve fin_ge_inv.
373
374 (** A congruence lemma for the [fin] constructor [Next] is similarly useful. *)
375
376 Lemma Next_cong : forall n (f1 f2 : fin n),
377 f1 = f2
378 -> Next f1 = Next f2.
379 dester.
380 Qed.
381
382 Hint Resolve Next_cong.
383
384 (** We prove a crucial lemma about [liftVar] in terms of [fin_ge]. *)
385
386 Lemma liftVar_more : forall n (f : fin n) (f0 : fin (S n)) g,
387 fin_ge g f0
388 -> match liftVar f0 f in fin n'
389 return fin n' -> (fin (pred n') -> fin n') -> fin (S n') with
390 | First n0 => fun _ _ => First
391 | Next n0 y' => fun _ fx' => Next (fx' y')
392 end g (liftVar g) = liftVar (more f0) (liftVar g f).
393 induction f; inversion 1; dester.
394 Qed.
395
396 Hint Resolve liftVar_more.
397
398 (** We suggest a particular way of changing the form of a goal, so that other hints are able to match. *)
399
400 Hint Extern 1 (_ = lift _ (Next (more ?f))) =>
401 change (Next (more f)) with (more (Next f)).
402
403 (** We suggest applying the [f_equal] tactic to simplify equalities over expressions. For instance, this would reduce a goal [App f1 x1 = App f2 x2] to two goals [f1 = f2] and [x1 = x2]. *)
404
405 Hint Extern 1 (eq (A := exp _) _ _) => f_equal.
406
407 (** Our consideration of lifting in isolation finishes with another hint lemma. The auxiliary lemma with a strengthened induction hypothesis is where we put [fin_ge] to use, and we do not need to mention that predicate again afteward. *)
408
409 Lemma double_lift' : forall n (e : exp n) f g,
410 fin_ge g f
411 -> lift (lift e f) (Next g) = lift (lift e g) (more f).
412 induction e; dester.
413 Qed.
414
415 Lemma double_lift : forall n (e : exp n) g,
416 lift (lift e First) (Next g) = lift (lift e g) First.
417 intros; apply double_lift'; dester.
418 Qed.
419
420 Hint Resolve double_lift.
421
422 (** Now we characterize the interaction of substitution and lifting on variables. We start with a more general form [substVar_lift'] of the final lemma [substVar_lift], with the latter proved as a direct corollary of the former. *)
423
424 Lemma substVar_lift' : forall n (f0 : fin n) f g,
425 substVar [nzf f0 return n, fin (S n) for
426 liftVar (more g) [sym_eq (nzf f0) return n, fin n for f0]]
427 (liftVar (liftVar (Next f0) [nzf f0 return n, fin n for g]) f)
428 = match substVar f0 f with
429 | Some f'' => Some [nzf f0 return n, fin n for liftVar g f'']
430 | None => None
431 end.
432 induction f0; dester.
433 Qed.
434
435 Lemma substVar_lift : forall n (f0 f g : fin (S n)),
436 substVar (liftVar (more g) f0) (liftVar (liftVar (Next f0) g) f)
437 = match substVar f0 f with
438 | Some f'' => Some (liftVar g f'')
439 | None => None
440 end.
441 intros; generalize (substVar_lift' f0 f g); dester.
442 Qed.
443
444 (** We follow a similar decomposition for the expression-level theorem about substitution and lifting. *)
445
446 Lemma lift_subst' : forall n (e1 : exp n) f g e2,
447 lift (subst e1 f e2) g
448 = [sym_eq (nzf f) return n, exp n for
449 subst
450 (lift e1 (liftVar (Next f) [nzf f return n, fin n for g]))
451 [nzf f return n, fin (S n) for
452 liftVar (more g) [sym_eq (nzf f) return n, fin n for f]]
453 [nzf f return n, exp n for lift e2 g]].
454 induction e1; generalize substVar_lift; dester.
455 Qed.
456
457 Lemma lift_subst : forall n g (e2 : exp (S n)) e3,
458 subst (lift e2 First) (Next g) (lift e3 First) = lift (n := n) (subst e2 g e3) First.
459 intros; generalize (lift_subst' e2 g First e3); dester.
460 Qed.
461
462 Hint Resolve lift_subst.
463
464 (** Our last auxiliary lemma characterizes a situation where substitution can undo the effects of lifting. *)
465
466 Lemma undo_lift' : forall n (e1 : exp n) e2 f,
467 subst (lift e1 f) f e2 = e1.
468 induction e1; generalize substVar_liftVar; dester.
469 Qed.
470
471 Lemma undo_lift : forall n e2 e3 (f0 : fin (S (S n))) g,
472 e3 = subst (lift e3 (unliftVar f0 g)) (unliftVar f0 g)
473 (subst (n := S n) e2 g e3).
474 generalize undo_lift'; dester.
475 Qed.
476
477 Hint Resolve undo_lift.
478
479 (** Finally, we arrive at the substitution commutativity theorem. *)
480
481 Lemma subst_comm' : forall n (e1 : exp n) f g e2 e3,
482 subst (subst e1 f e2) g e3
483 = subst
484 (subst e1 (liftVar f g) [nzf g return n, exp n for
485 lift e3 [sym_eq (nzf g) return n, fin n for unliftVar f g]])
486 (unliftVar f g)
487 (subst e2 g e3).
488 induction e1; generalize (substVar_unliftVar (n := n)); dester.
489 Qed.
490
491 Theorem subst_comm : forall (e1 : exp 2) e2 e3,
492 subst (subst e1 First e2) First e3
493 = subst (subst e1 (Next First) (lift e3 First)) First (subst e2 First e3).
494 intros; generalize (subst_comm' e1 First First e2 e3); dester.
495 Qed.
496
497 (** The final theorem is specialized to the case of substituting in an expression with exactly two free variables, which yields a statement that is readable enough, as statements about de Bruijn indices go.
498
499 This proof script is resilient to specification changes. It is easy to add new constructors to the language being treated. The proofs adapt automatically to the addition of any constructor whose subterms each involve zero or one new bound variables. That is, to add such a constructor, we only need to add it to the definition of [exp] and add (quite obvious) cases for it in the definitions of [lift] and [subst]. *)