comparison src/Intensional.v @ 260:9632f7316c29

Prosified Intensional
author Adam Chlipala <adamc@hcoop.net>
date Wed, 16 Dec 2009 17:24:58 -0500
parents 0c5fb678bfe2
children b441010125d4
comparison
equal deleted inserted replaced
259:0c5fb678bfe2 260:9632f7316c29
15 Set Implicit Arguments. 15 Set Implicit Arguments.
16 (* end hide *) 16 (* end hide *)
17 17
18 18
19 (** %\chapter{Intensional Transformations}% *) 19 (** %\chapter{Intensional Transformations}% *)
20
21 (** The essential benefit of higher-order encodings is that variable contexts are implicit. We represent the object language's contexts within the meta language's contexts. For translations like CPS conversion, this is a clear win, since such translations do not need to keep track of details of which variables are available. Other important translations, including closure conversion, need to work with variables as first-class, analyzable values.
22
23 Another example is conversion from PHOAS terms to de Bruijn terms. The output format makes the structure of variables explicit, so the translation requires explicit reasoning about variable identity. In this chapter, we implement verified translations in both directions between last chapter's PHOAS language and a de Bruijn version of it. Along the way, we show one approach to avoiding the use of axioms with PHOAS. *)
20 24
21 (* begin hide *) 25 (* begin hide *)
22 26
23 Inductive type : Type := 27 Inductive type : Type :=
24 | Nat : type 28 | Nat : type
109 End exp_equiv. 113 End exp_equiv.
110 114
111 Definition Wf t (E : Exp t) := forall var1 var2, exp_equiv nil (E var1) (E var2). 115 Definition Wf t (E : Exp t) := forall var1 var2, exp_equiv nil (E var1) (E var2).
112 End Phoas. 116 End Phoas.
113 (* end hide *) 117 (* end hide *)
118
119 (** The de Bruijn version of simply-typed lambda calculus is defined in a manner that should be old hat by now. *)
114 120
115 Module DeBruijn. 121 Module DeBruijn.
116 Inductive exp : list type -> type -> Type := 122 Inductive exp : list type -> type -> Type :=
117 | Var : forall G t, 123 | Var : forall G t,
118 member t G 124 member t G
146 Import Phoas DeBruijn. 152 Import Phoas DeBruijn.
147 153
148 154
149 (** * From De Bruijn to PHOAS *) 155 (** * From De Bruijn to PHOAS *)
150 156
157 (** The heart of the translation into PHOAS is this function [phoasify], which is parameterized by an [hlist] that represents a mapping from de Bruijn variables to PHOAS variables. *)
158
151 Section phoasify. 159 Section phoasify.
152 Variable var : type -> Type. 160 Variable var : type -> Type.
153 161
154 Fixpoint phoasify G t (e : DeBruijn.exp G t) : hlist var G -> Phoas.exp var t := 162 Fixpoint phoasify G t (e : DeBruijn.exp G t) : hlist var G -> Phoas.exp var t :=
155 match e with 163 match e with
164 End phoasify. 172 End phoasify.
165 173
166 Definition Phoasify t (e : DeBruijn.exp nil t) : Phoas.Exp t := 174 Definition Phoasify t (e : DeBruijn.exp nil t) : Phoas.Exp t :=
167 fun _ => phoasify e HNil. 175 fun _ => phoasify e HNil.
168 176
177 (** It turns out to be trivial to establish the translation's soundness. *)
178
169 Theorem phoasify_sound : forall G t (e : DeBruijn.exp G t) s, 179 Theorem phoasify_sound : forall G t (e : DeBruijn.exp G t) s,
170 Phoas.expDenote (phoasify e s) = DeBruijn.expDenote e s. 180 Phoas.expDenote (phoasify e s) = DeBruijn.expDenote e s.
171 induction e; crush; ext_eq; crush. 181 induction e; crush; ext_eq; crush.
172 Qed. 182 Qed.
173 183
184 (** We can prove that any output of [Phoasify] is well-formed, in a sense strong enough to let us avoid asserting last chapter's axiom. *)
185
186 Print Wf.
187 (** %\vspace{-.15in}% [[
188 Wf =
189 fun (t : type) (E : Exp t) =>
190 forall var1 var2 : type -> Type, exp_equiv nil (E var1) (E var2)
191 : forall t : type, Exp t -> Prop
192 ]] *)
193
174 Section vars. 194 Section vars.
175 Variables var1 var2 : type -> Type. 195 Variables var1 var2 : type -> Type.
176 196
177 Fixpoint zip G (s1 : hlist var1 G) : hlist var2 G -> list {t : type & var1 t * var2 t}%type := 197 (** In the course of proving well-formedness, we will need to translate back and forth between the de Bruijn and PHOAS representations of free variable information. The function [zip] combines two de Bruijn substitutions into a single PHOAS context. *)
198
199 Fixpoint zip G (s1 : hlist var1 G)
200 : hlist var2 G -> list {t : type & var1 t * var2 t}%type :=
178 match s1 with 201 match s1 with
179 | HNil => fun _ => nil 202 | HNil => fun _ => nil
180 | HCons _ _ v1 s1' => fun s2 => existT _ _ (v1, hhd s2) :: zip s1' (htl s2) 203 | HCons _ _ v1 s1' => fun s2 => existT _ _ (v1, hhd s2) :: zip s1' (htl s2)
181 end. 204 end.
205
206 (** Two simple lemmas about [zip] will make useful hints. *)
182 207
183 Lemma In_zip : forall t G (s1 : hlist _ G) s2 (m : member t G), 208 Lemma In_zip : forall t G (s1 : hlist _ G) s2 (m : member t G),
184 In (existT _ t (hget s1 m, hget s2 m)) (zip s1 s2). 209 In (existT _ t (hget s1 m, hget s2 m)) (zip s1 s2).
185 induction s1; intro s2; dep_destruct s2; intro m; dep_destruct m; crush. 210 induction s1; intro s2; dep_destruct s2; intro m; dep_destruct m; crush.
186 Qed. 211 Qed.
192 trivial. 217 trivial.
193 Qed. 218 Qed.
194 219
195 Hint Resolve In_zip unsimpl_zip. 220 Hint Resolve In_zip unsimpl_zip.
196 221
197 Theorem phoasify_wf : forall G t (e : DeBruijn.exp G t) s1 s2, 222 (** Now it is trivial to prove the main inductive lemma about well-formedness. *)
223
224 Lemma phoasify_wf : forall G t (e : DeBruijn.exp G t) s1 s2,
198 exp_equiv (zip s1 s2) (phoasify e s1) (phoasify e s2). 225 exp_equiv (zip s1 s2) (phoasify e s1) (phoasify e s2).
199 Hint Constructors exp_equiv. 226 Hint Constructors exp_equiv.
200 227
201 induction e; crush. 228 induction e; crush.
202 Qed. 229 Qed.
203 End vars. 230 End vars.
204 231
232 (** We apply [phoasify_wf] manually to prove the final theorem. *)
233
205 Theorem Phoasify_wf : forall t (e : DeBruijn.exp nil t), 234 Theorem Phoasify_wf : forall t (e : DeBruijn.exp nil t),
206 Wf (Phoasify e). 235 Wf (Phoasify e).
207 unfold Wf, Phoasify; intros; 236 unfold Wf, Phoasify; intros;
208 apply (phoasify_wf e (HNil (B := var1)) (HNil (B := var2))). 237 apply (phoasify_wf e (HNil (B := var1)) (HNil (B := var2))).
209 Qed. 238 Qed.
210 239
240 (** Now, if we compose [Phoasify] with any translation over PHOAS terms, we can verify the composed translation without relying on axioms. The conclusion of [Phoasify_wf] is robustly useful in verifying a wide variety of translations that use a wide variety of [var] instantiations. *)
241
211 242
212 (** * From PHOAS to De Bruijn *) 243 (** * From PHOAS to De Bruijn *)
213 244
214 Fixpoint lookup (G : list type) (n : nat) : option type := 245 (** The translation to de Bruijn terms is more involved. We will essentially be instantiating and using a PHOAS term following a convention isomorphic to %\textit{%#<i>#de Bruijn levels#</i>#%}%, which are different from the de Bruijn indices that we have treated so far. With levels, a given bound variable is referred to by the same number at each of its occurrences. In any expression, the binders that are not enclosed by other binders are assigned level 0, a binder with just one enclosing binder is assigned level 1, and so on. The uniformity of references to any binder will be critical to our translation, since it is compatible with the pattern of filling in all of a PHOAS variable's locations at once by applying a function.
215 match G with 246
247 We implement a special lookup function, for reading a numbered variable's type out of a de Bruijn level typing context. The last variable in the list is taken to have level 0, the next-to-last level 1, and so on. *)
248
249 Fixpoint lookup (ts : list type) (n : nat) : option type :=
250 match ts with
216 | nil => None 251 | nil => None
217 | t :: G' => if eq_nat_dec n (length G') then Some t else lookup G' n 252 | t :: ts' => if eq_nat_dec n (length ts') then Some t else lookup ts' n
218 end. 253 end.
219 254
220 Infix "##" := lookup (left associativity, at level 1). 255 Infix "##" := lookup (left associativity, at level 1).
256
257 (** With [lookup], we can define a notion of well-formedness for PHOAS expressions that we are treating according to the de Bruijn level convention. *)
221 258
222 Fixpoint wf (ts : list type) t (e : Phoas.exp (fun _ => nat) t) : Prop := 259 Fixpoint wf (ts : list type) t (e : Phoas.exp (fun _ => nat) t) : Prop :=
223 match e with 260 match e with
224 | Phoas.Var t n => ts ## n = Some t 261 | Phoas.Var t n => ts ## n = Some t
225 | Phoas.Const _ => True 262 | Phoas.Const _ => True
226 | Phoas.Plus e1 e2 => wf ts e1 /\ wf ts e2 263 | Phoas.Plus e1 e2 => wf ts e1 /\ wf ts e2
227 | Phoas.App _ _ e1 e2 => wf ts e1 /\ wf ts e2 264 | Phoas.App _ _ e1 e2 => wf ts e1 /\ wf ts e2
228 | Phoas.Abs t1 _ e1 => wf (t1 :: ts) (e1 (length ts)) 265 | Phoas.Abs t1 _ e1 => wf (t1 :: ts) (e1 (length ts))
229 end. 266 end.
230 267
268 (** ** Connecting Notions of Well-Formedness *)
269
270 (** Our first order of business now is to prove that any well-formed [Exp] instantiates to a well-formed de Bruijn level expression. We start by characterizing, as a function of de Bruijn level contexts, the set of PHOAS contexts that will occur in the proof, where we will be inducting over an [exp_equiv] derivation. *)
271
231 Fixpoint makeG (ts : list type) : list { t : type & nat * nat }%type := 272 Fixpoint makeG (ts : list type) : list { t : type & nat * nat }%type :=
232 match ts with 273 match ts with
233 | nil => nil 274 | nil => nil
234 | t :: ts' => existT _ t (length ts', length ts') :: makeG ts' 275 | t :: ts' => existT _ t (length ts', length ts') :: makeG ts'
235 end. 276 end.
277
278 (** Now we prove a connection between [lookup] and [makeG], by way of a lemma about [lookup]. *)
236 279
237 Opaque eq_nat_dec. 280 Opaque eq_nat_dec.
238 Hint Extern 1 (_ >= _) => omega. 281 Hint Extern 1 (_ >= _) => omega.
239 282
240 Lemma lookup_contra' : forall t ts n, 283 Lemma lookup_contra' : forall t ts n,
264 end; elimtype False; eauto. 307 end; elimtype False; eauto.
265 Qed. 308 Qed.
266 309
267 Hint Resolve lookup_In. 310 Hint Resolve lookup_In.
268 311
312 (** We can prove the main inductive lemma by induction over [exp_equiv] derivations. *)
313
269 Hint Extern 1 (_ :: _ = makeG (_ :: _)) => reflexivity. 314 Hint Extern 1 (_ :: _ = makeG (_ :: _)) => reflexivity.
270 315
271 Lemma Wf_wf' : forall G t e1 (e2 : Phoas.exp (fun _ => nat) t), 316 Lemma Wf_wf' : forall G t e1 (e2 : Phoas.exp (fun _ => nat) t),
272 exp_equiv G e1 e2 317 exp_equiv G e1 e2
273 -> forall ts, G = makeG ts 318 -> forall ts, G = makeG ts
279 Wf E 324 Wf E
280 -> wf nil (E (fun _ => nat)). 325 -> wf nil (E (fun _ => nat)).
281 intros; eapply Wf_wf'; eauto. 326 intros; eapply Wf_wf'; eauto.
282 Qed. 327 Qed.
283 328
329 (** ** The Translation *)
330
331 (** Implementing the translation itself will require some proofs. Our main function [dbify] will take [wf] proofs as arguments, and these proofs will be critical to constructing de Bruijn index terms. First, we use [congruence] to prove two basic theorems about [option]s. *)
332
284 Theorem None_Some : forall T (x : T), 333 Theorem None_Some : forall T (x : T),
285 None = Some x 334 None = Some x
286 -> False. 335 -> False.
287 congruence. 336 congruence.
288 Qed. 337 Qed.
291 Some x = Some y 340 Some x = Some y
292 -> x = y. 341 -> x = y.
293 congruence. 342 congruence.
294 Qed. 343 Qed.
295 344
345 (** We can use these theorems to implement [makeVar], which translates a proof about [lookup] into a de Bruijn index variable with a closely related type. *)
346
296 Fixpoint makeVar {ts n t} : ts ## n = Some t -> member t ts := 347 Fixpoint makeVar {ts n t} : ts ## n = Some t -> member t ts :=
297 match ts return ts ## n = Some t -> member t ts with 348 match ts with
298 | nil => fun Heq => match None_Some Heq with end 349 | nil => fun Heq => match None_Some Heq with end
299 | t' :: ts' => if eq_nat_dec n (length ts') as b 350 | t' :: ts' => if eq_nat_dec n (length ts') as b return (if b then _ else _) = _ -> _
300 return (if b then Some t' else lookup ts' n) = Some t -> member t (t' :: ts') 351 then fun Heq => match Some_Some Heq with refl_equal => HFirst end
301 then fun Heq => match Some_Some Heq in _ = t0 return member t0 (t' :: ts') with
302 | refl_equal => HFirst
303 end
304 else fun Heq => HNext (makeVar Heq) 352 else fun Heq => HNext (makeVar Heq)
305 end. 353 end.
306 354
307 Axiom cheat : forall T, T. 355 (** Now [dbify] is straightforward to define. We use the functions [proj1] and [proj2] to decompose proofs of conjunctions. *)
308 356
309 Fixpoint dbify {ts} t (e : Phoas.exp (fun _ => nat) t) : wf ts e -> DeBruijn.exp ts t := 357 Fixpoint dbify {ts} t (e : Phoas.exp (fun _ => nat) t) : wf ts e -> DeBruijn.exp ts t :=
310 match e in Phoas.exp _ t return wf ts e -> DeBruijn.exp ts t with 358 match e in Phoas.exp _ t return wf ts e -> DeBruijn.exp ts t with
311 | Phoas.Var _ n => fun wf => DeBruijn.Var (makeVar wf) 359 | Phoas.Var _ n => fun wf => DeBruijn.Var (makeVar wf)
312 360
313 | Phoas.Const n => fun _ => DeBruijn.Const n 361 | Phoas.Const n => fun _ => DeBruijn.Const n
314 | Phoas.Plus e1 e2 => fun wf => DeBruijn.Plus (dbify e1 (proj1 wf)) (dbify e2 (proj2 wf)) 362 | Phoas.Plus e1 e2 => fun wf =>
315 363 DeBruijn.Plus (dbify e1 (proj1 wf)) (dbify e2 (proj2 wf))
316 | Phoas.App _ _ e1 e2 => fun wf => DeBruijn.App (dbify e1 (proj1 wf)) (dbify e2 (proj2 wf)) 364
365 | Phoas.App _ _ e1 e2 => fun wf =>
366 DeBruijn.App (dbify e1 (proj1 wf)) (dbify e2 (proj2 wf))
317 | Phoas.Abs _ _ e1 => fun wf => DeBruijn.Abs (dbify (e1 (length ts)) wf) 367 | Phoas.Abs _ _ e1 => fun wf => DeBruijn.Abs (dbify (e1 (length ts)) wf)
318 end. 368 end.
369
370 (** We define the parametric translation [Dbify] by appealing to the well-formedness translation theorem [Wf_wf] that we proved earlier. *)
319 371
320 Definition Dbify t (E : Phoas.Exp t) (W : Wf E) : DeBruijn.exp nil t := 372 Definition Dbify t (E : Phoas.Exp t) (W : Wf E) : DeBruijn.exp nil t :=
321 dbify (E _) (Wf_wf W). 373 dbify (E _) (Wf_wf W).
322 374
323 Fixpoint makeG' ts (s : hlist typeDenote ts) : list { t : type & nat * typeDenote t }%type := 375 (** To prove soundness, it is helpful to classify a set of contexts which depends on a de Bruijn index substitution. *)
376
377 Fixpoint makeG' ts (s : hlist typeDenote ts)
378 : list { t : type & nat * typeDenote t }%type :=
324 match s with 379 match s with
325 | HNil => nil 380 | HNil => nil
326 | HCons _ ts' v s' => existT _ _ (length ts', v) :: makeG' s' 381 | HCons _ ts' v s' => existT _ _ (length ts', v) :: makeG' s'
327 end. 382 end.
383
384 (** We prove an analogous lemma to the one we proved connecting [makeG] and [lookup]. This time, we connect [makeG'] and [hget]. *)
328 385
329 Lemma In_makeG'_contra' : forall t v2 ts (s : hlist _ ts) n, 386 Lemma In_makeG'_contra' : forall t v2 ts (s : hlist _ ts) n,
330 In (existT _ t (n, v2)) (makeG' s) 387 In (existT _ t (n, v2)) (makeG' s)
331 -> n >= length ts 388 -> n >= length ts
332 -> False. 389 -> False.
347 induction s; crush; 404 induction s; crush;
348 match goal with 405 match goal with
349 | [ |- context[if ?E then _ else _] ] => destruct E; crush 406 | [ |- context[if ?E then _ else _] ] => destruct E; crush
350 end; 407 end;
351 repeat match goal with 408 repeat match goal with
352 | [ |- context[match ?pf with refl_equal => _ end] ] => rewrite (UIP_refl _ _ pf) 409 | [ |- context[match ?pf with refl_equal => _ end] ] =>
410 rewrite (UIP_refl _ _ pf)
353 end; crush; elimtype False; eauto. 411 end; crush; elimtype False; eauto.
354 Qed. 412 Qed.
355 413
356 Hint Resolve In_makeG'. 414 Hint Resolve In_makeG'.
415
416 (** Now the main inductive lemma can be stated and proved simply. *)
357 417
358 Lemma dbify_sound : forall G t (e1 : Phoas.exp _ t) (e2 : Phoas.exp _ t), 418 Lemma dbify_sound : forall G t (e1 : Phoas.exp _ t) (e2 : Phoas.exp _ t),
359 exp_equiv G e1 e2 419 exp_equiv G e1 e2
360 -> forall ts (w : wf ts e1) s, 420 -> forall ts (w : wf ts e1) s,
361 G = makeG' s 421 G = makeG' s
362 -> DeBruijn.expDenote (dbify e1 w) s = Phoas.expDenote e2. 422 -> DeBruijn.expDenote (dbify e1 w) s = Phoas.expDenote e2.
363 induction 1; crush; ext_eq; crush. 423 induction 1; crush; ext_eq; crush.
364 Qed. 424 Qed.
365 425
426 (** In the usual way, we wrap [dbify_sound] into the final soundness theorem, formally establishing the expressive equivalence of PHOAS and de Bruijn index terms. *)
427
366 Theorem Dbify_sound : forall t (E : Exp t) (W : Wf E), 428 Theorem Dbify_sound : forall t (E : Exp t) (W : Wf E),
367 DeBruijn.expDenote (Dbify W) HNil = Phoas.ExpDenote E. 429 DeBruijn.expDenote (Dbify W) HNil = Phoas.ExpDenote E.
368 unfold Dbify, Phoas.ExpDenote; intros; eapply dbify_sound; eauto. 430 unfold Dbify, Phoas.ExpDenote; intros; eapply dbify_sound; eauto.
369 Qed. 431 Qed.