comparison src/Intensional.v @ 183:02569049397b

Closure conversion defined
author Adam Chlipala <adamc@hcoop.net>
date Sun, 16 Nov 2008 14:01:33 -0500
parents 24b99e025fe8
children 699fd70c04a7
comparison
equal deleted inserted replaced
182:24b99e025fe8 183:02569049397b
6 * The license text is available at: 6 * The license text is available at:
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/ 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
8 *) 8 *)
9 9
10 (* begin hide *) 10 (* begin hide *)
11 Require Import String List. 11 Require Import Arith Bool String List.
12 12
13 Require Import Axioms Tactics DepList. 13 Require Import Axioms Tactics DepList.
14 14
15 Set Implicit Arguments. 15 Set Implicit Arguments.
16 (* end hide *) 16 (* end hide *)
128 128
129 Axiom Exp_equiv : forall t (E : Exp t) var1 var2, 129 Axiom Exp_equiv : forall t (E : Exp t) var1 var2,
130 exp_equiv nil (E var1) (E var2). 130 exp_equiv nil (E var1) (E var2).
131 End Source. 131 End Source.
132 132
133 Section Closed. 133 Module Closed.
134 Inductive type : Type := 134 Inductive type : Type :=
135 | TNat : type 135 | TNat : type
136 | Arrow : type -> type -> type 136 | Arrow : type -> type -> type
137 | Code : type -> type -> type -> type 137 | Code : type -> type -> type -> type
138 | Prod : type -> type -> type 138 | Prod : type -> type -> type
179 | Fst : forall t1 t2, 179 | Fst : forall t1 t2,
180 exp (t1 ** t2) 180 exp (t1 ** t2)
181 -> exp t1 181 -> exp t1
182 | Snd : forall t1 t2, 182 | Snd : forall t1 t2,
183 exp (t1 ** t2) 183 exp (t1 ** t2)
184 -> exp t2
185
186 | Let : forall t1 t2,
187 exp t1
188 -> (var t1 -> exp t2)
184 -> exp t2. 189 -> exp t2.
185 190
186 Section funcs. 191 Section funcs.
187 Variable T : Type. 192 Variable T : Type.
188 193
209 Notation "# v" := (Var v) (at level 70) : cc_scope. 214 Notation "# v" := (Var v) (at level 70) : cc_scope.
210 215
211 Notation "^ n" := (Const n) (at level 70) : cc_scope. 216 Notation "^ n" := (Const n) (at level 70) : cc_scope.
212 Infix "+^" := Plus (left associativity, at level 79) : cc_scope. 217 Infix "+^" := Plus (left associativity, at level 79) : cc_scope.
213 218
214 Infix "@@" := App (no associativity, at level 75) : cc_scope. 219 Infix "@" := App (left associativity, at level 77) : cc_scope.
215 Infix "##" := Pack (no associativity, at level 71) : cc_scope. 220 Infix "##" := Pack (no associativity, at level 71) : cc_scope.
216 221
217 Notation "()" := EUnit : cc_scope. 222 Notation "()" := EUnit : cc_scope.
218 223
219 Notation "[ x1 , x2 ]" := (Pair x1 x2) (at level 73) : cc_scope. 224 Notation "[ x1 , x2 ]" := (Pair x1 x2) (at level 73) : cc_scope.
220 Notation "#1 x" := (Fst x) (at level 72) : cc_scope. 225 Notation "#1 x" := (Fst x) (at level 72) : cc_scope.
221 Notation "#2 x" := (Snd x) (at level 72) : cc_scope. 226 Notation "#2 x" := (Snd x) (at level 72) : cc_scope.
222 227
223 Notation "f <= \\ x , y , e ; fs" := 228 Notation "f <== \\ x , y , e ; fs" :=
224 (Abs (fun x y => e) (fun f => fs)) 229 (Abs (fun x y => e) (fun f => fs))
225 (right associativity, at level 80, e at next level) : cc_scope. 230 (right associativity, at level 80, e at next level) : cc_scope.
231
232 Notation "x <- e1 ; e2" := (Let e1 (fun x => e2))
233 (right associativity, at level 80, e1 at next level) : cc_scope.
226 234
227 Bind Scope cc_scope with exp funcs prog. 235 Bind Scope cc_scope with exp funcs prog.
228 236
229 Fixpoint typeDenote (t : type) : Set := 237 Fixpoint typeDenote (t : type) : Set :=
230 match t with 238 match t with
248 | EUnit => tt 256 | EUnit => tt
249 257
250 | Pair _ _ e1 e2 => (expDenote e1, expDenote e2) 258 | Pair _ _ e1 e2 => (expDenote e1, expDenote e2)
251 | Fst _ _ e' => fst (expDenote e') 259 | Fst _ _ e' => fst (expDenote e')
252 | Snd _ _ e' => snd (expDenote e') 260 | Snd _ _ e' => snd (expDenote e')
261
262 | Let _ _ e1 e2 => expDenote (e2 (expDenote e1))
253 end. 263 end.
254 264
255 Fixpoint funcsDenote T (fs : funcs typeDenote T) : T := 265 Fixpoint funcsDenote T (fs : funcs typeDenote T) : T :=
256 match fs with 266 match fs with
257 | Main v => v 267 | Main v => v
267 277
268 Definition ExpDenote t (E : Exp t) := expDenote (E _). 278 Definition ExpDenote t (E : Exp t) := expDenote (E _).
269 Definition ProgDenote t (P : Prog t) := progDenote (P _). 279 Definition ProgDenote t (P : Prog t) := progDenote (P _).
270 End Closed. 280 End Closed.
271 281
282 Import Source Closed.
283
284 Section splice.
285 Open Local Scope cc_scope.
286
287 Fixpoint spliceFuncs var T1 (fs : funcs var T1)
288 T2 (f : T1 -> funcs var T2) {struct fs} : funcs var T2 :=
289 match fs with
290 | Main v => f v
291 | Abs _ _ _ e fs' => Abs e (fun x => spliceFuncs (fs' x) f)
292 end.
293 End splice.
294
295 Notation "x <-- e1 ; e2" := (spliceFuncs e1 (fun x => e2))
296 (right associativity, at level 80, e1 at next level) : cc_scope.
297
298 Definition natvar (_ : Source.type) := nat.
299 Definition isfree := hlist (fun (_ : Source.type) => bool).
300
301 Ltac maybe_destruct E :=
302 match goal with
303 | [ x : _ |- _ ] =>
304 match E with
305 | x => idtac
306 end
307 | _ =>
308 match E with
309 | eq_nat_dec _ _ => idtac
310 end
311 end; destruct E.
312
313 Ltac my_crush :=
314 crush; repeat (match goal with
315 | [ x : (_ * _)%type |- _ ] => destruct x
316 | [ |- context[if ?B then _ else _] ] => maybe_destruct B
317 | [ _ : context[if ?B then _ else _] |- _ ] => maybe_destruct B
318 end; crush).
319
320 Section isfree.
321 Import Source.
322 Open Local Scope source_scope.
323
324 Hint Extern 3 False => omega.
325
326 Fixpoint lookup_type (envT : list Source.type) (n : nat) {struct envT}
327 : isfree envT -> option Source.type :=
328 match envT return (isfree envT -> _) with
329 | nil => fun _ => None
330 | first :: rest => fun fvs =>
331 if eq_nat_dec n (length rest)
332 then match fvs with
333 | (true, _) => Some first
334 | (false, _) => None
335 end
336 else lookup_type rest n (snd fvs)
337 end.
338
339 Implicit Arguments lookup_type [envT].
340
341 Notation ok := (fun (envT : list Source.type) (fvs : isfree envT)
342 (n : nat) (t : Source.type)
343 => lookup_type n fvs = Some t).
344
345 Fixpoint wfExp (envT : list Source.type) (fvs : isfree envT)
346 t (e : Source.exp natvar t) {struct e} : Prop :=
347 match e with
348 | Var t v => ok envT fvs v t
349
350 | Const _ => True
351 | Plus e1 e2 => wfExp envT fvs e1 /\ wfExp envT fvs e2
352
353 | App _ _ e1 e2 => wfExp envT fvs e1 /\ wfExp envT fvs e2
354 | Abs dom _ e' => wfExp (dom :: envT) (true ::: fvs) (e' (length envT))
355 end.
356
357 Implicit Arguments wfExp [envT t].
358
359 Theorem wfExp_weaken : forall t (e : exp natvar t) envT (fvs fvs' : isfree envT),
360 wfExp fvs e
361 -> (forall n t, ok _ fvs n t -> ok _ fvs' n t)
362 -> wfExp fvs' e.
363 Hint Extern 1 (lookup_type (envT := _ :: _) _ _ = Some _) =>
364 simpl in *; my_crush.
365
366 induction e; my_crush; eauto.
367 Defined.
368
369 Fixpoint isfree_none (envT : list Source.type) : isfree envT :=
370 match envT return (isfree envT) with
371 | nil => tt
372 | _ :: _ => (false, isfree_none _)
373 end.
374
375 Implicit Arguments isfree_none [envT].
376
377 Fixpoint isfree_one (envT : list Source.type) (n : nat) {struct envT} : isfree envT :=
378 match envT return (isfree envT) with
379 | nil => tt
380 | _ :: rest =>
381 if eq_nat_dec n (length rest)
382 then (true, isfree_none)
383 else (false, isfree_one _ n)
384 end.
385
386 Implicit Arguments isfree_one [envT].
387
388 Fixpoint isfree_merge (envT : list Source.type) : isfree envT -> isfree envT -> isfree envT :=
389 match envT return (isfree envT -> isfree envT -> isfree envT) with
390 | nil => fun _ _ => tt
391 | _ :: _ => fun fv1 fv2 => (fst fv1 || fst fv2, isfree_merge _ (snd fv1) (snd fv2))
392 end.
393
394 Implicit Arguments isfree_merge [envT].
395
396 Fixpoint fvsExp t (e : exp natvar t)
397 (envT : list Source.type) {struct e} : isfree envT :=
398 match e with
399 | Var _ n => isfree_one n
400
401 | Const _ => isfree_none
402 | Plus e1 e2 => isfree_merge (fvsExp e1 envT) (fvsExp e2 envT)
403
404 | App _ _ e1 e2 => isfree_merge (fvsExp e1 envT) (fvsExp e2 envT)
405 | Abs dom _ e' => snd (fvsExp (e' (length envT)) (dom :: envT))
406 end.
407
408 Lemma isfree_one_correct : forall t (v : natvar t) envT fvs,
409 ok envT fvs v t
410 -> ok envT (isfree_one (envT:=envT) v) v t.
411 induction envT; my_crush; eauto.
412 Defined.
413
414 Lemma isfree_merge_correct1 : forall t (v : natvar t) envT fvs1 fvs2,
415 ok envT fvs1 v t
416 -> ok envT (isfree_merge (envT:=envT) fvs1 fvs2) v t.
417 induction envT; my_crush; eauto.
418 Defined.
419
420 Hint Rewrite orb_true_r : cpdt.
421
422 Lemma isfree_merge_correct2 : forall t (v : natvar t) envT fvs1 fvs2,
423 ok envT fvs2 v t
424 -> ok envT (isfree_merge (envT:=envT) fvs1 fvs2) v t.
425 induction envT; my_crush; eauto.
426 Defined.
427
428 Hint Resolve isfree_one_correct isfree_merge_correct1 isfree_merge_correct2.
429
430 Lemma fvsExp_correct : forall t (e : exp natvar t)
431 envT (fvs : isfree envT), wfExp fvs e
432 -> forall (fvs' : isfree envT),
433 (forall v t, ok envT (fvsExp e envT) v t -> ok envT fvs' v t)
434 -> wfExp fvs' e.
435 Hint Extern 1 (_ = _) =>
436 match goal with
437 | [ H : lookup_type _ (fvsExp ?X ?Y) = _ |- _ ] =>
438 destruct (fvsExp X Y); my_crush
439 end.
440
441 induction e; my_crush; eauto.
442 Defined.
443
444 Lemma lookup_type_unique : forall v t1 t2 envT (fvs1 fvs2 : isfree envT),
445 lookup_type v fvs1 = Some t1
446 -> lookup_type v fvs2 = Some t2
447 -> t1 = t2.
448 induction envT; my_crush; eauto.
449 Defined.
450
451 Implicit Arguments lookup_type_unique [v t1 t2 envT fvs1 fvs2].
452
453 Hint Extern 2 (lookup_type _ _ = Some _) =>
454 match goal with
455 | [ H1 : lookup_type ?v _ = Some _,
456 H2 : lookup_type ?v _ = Some _ |- _ ] =>
457 (generalize (lookup_type_unique H1 H2); intro; subst)
458 || rewrite <- (lookup_type_unique H1 H2)
459 end.
460
461 Lemma lookup_none : forall v t envT,
462 lookup_type (envT:=envT) v (isfree_none (envT:=envT)) = Some t
463 -> False.
464 induction envT; my_crush.
465 Defined.
466
467 Hint Extern 2 (_ = _) => elimtype False; eapply lookup_none; eassumption.
468
469 Lemma lookup_one : forall v' v t envT,
470 lookup_type (envT:=envT) v' (isfree_one (envT:=envT) v) = Some t
471 -> v' = v.
472 induction envT; my_crush.
473 Defined.
474
475 Implicit Arguments lookup_one [v' v t envT].
476
477 Hint Extern 2 (lookup_type _ _ = Some _) =>
478 match goal with
479 | [ H : lookup_type _ _ = Some _ |- _ ] =>
480 generalize (lookup_one H); intro; subst
481 end.
482
483 Lemma lookup_merge : forall v t envT (fvs1 fvs2 : isfree envT),
484 lookup_type v (isfree_merge fvs1 fvs2) = Some t
485 -> lookup_type v fvs1 = Some t
486 \/ lookup_type v fvs2 = Some t.
487 induction envT; my_crush.
488 Defined.
489
490 Implicit Arguments lookup_merge [v t envT fvs1 fvs2].
491
492 Lemma lookup_bound : forall v t envT (fvs : isfree envT),
493 lookup_type v fvs = Some t
494 -> v < length envT.
495 Hint Resolve lt_S.
496 induction envT; my_crush; eauto.
497 Defined.
498
499 Hint Resolve lookup_bound.
500
501 Lemma lookup_bound_contra : forall t envT (fvs : isfree envT),
502 lookup_type (length envT) fvs = Some t
503 -> False.
504 intros; assert (length envT < length envT); eauto.
505 Defined.
506
507 Hint Resolve lookup_bound_contra.
508
509 Hint Extern 3 (_ = _) => elimtype False; omega.
510
511 Lemma lookup_push_drop : forall v t t' envT fvs,
512 v < length envT
513 -> lookup_type (envT := t :: envT) v (true, fvs) = Some t'
514 -> lookup_type (envT := envT) v fvs = Some t'.
515 my_crush.
516 Defined.
517
518 Lemma lookup_push_add : forall v t t' envT fvs,
519 lookup_type (envT := envT) v (snd fvs) = Some t'
520 -> lookup_type (envT := t :: envT) v fvs = Some t'.
521 my_crush; elimtype False; eauto.
522 Defined.
523
524 Hint Resolve lookup_bound lookup_push_drop lookup_push_add.
525
526 Theorem fvsExp_minimal : forall t (e : exp natvar t)
527 envT (fvs : isfree envT), wfExp fvs e
528 -> (forall v t, ok envT (fvsExp e envT) v t -> ok envT fvs v t).
529 Hint Extern 1 (_ = _) =>
530 match goal with
531 | [ H : lookup_type _ (isfree_merge _ _) = Some _ |- _ ] =>
532 destruct (lookup_merge H); clear H; eauto
533 end.
534
535 induction e; my_crush; eauto.
536 Defined.
537
538 Fixpoint ccType (t : Source.type) : Closed.type :=
539 match t with
540 | Nat%source => Nat
541 | (dom --> ran)%source => ccType dom --> ccType ran
542 end%cc.
543
544 Open Local Scope cc_scope.
545
546 Fixpoint envType (envT : list Source.type) : isfree envT -> Closed.type :=
547 match envT return (isfree envT -> _) with
548 | nil => fun _ => Unit
549 | t :: _ => fun tup =>
550 if fst tup
551 then ccType t ** envType _ (snd tup)
552 else envType _ (snd tup)
553 end.
554
555 Implicit Arguments envType [envT].
556
557 Fixpoint envOf (var : Closed.type -> Set) (envT : list Source.type) {struct envT}
558 : isfree envT -> Set :=
559 match envT return (isfree envT -> _) with
560 | nil => fun _ => unit
561 | first :: rest => fun fvs =>
562 match fvs with
563 | (true, fvs') => (var (ccType first) * envOf var rest fvs')%type
564 | (false, fvs') => envOf var rest fvs'
565 end
566 end.
567
568 Implicit Arguments envOf [envT].
569
570 Notation "var <| to" := (match to with
571 | None => unit
572 | Some t => var (ccType t)
573 end) (no associativity, at level 70).
574
575 Fixpoint lookup (var : Closed.type -> Set) (envT : list Source.type) :
576 forall (n : nat) (fvs : isfree envT), envOf var fvs -> var <| lookup_type n fvs :=
577 match envT return (forall (n : nat) (fvs : isfree envT), envOf var fvs
578 -> var <| lookup_type n fvs) with
579 | nil => fun _ _ _ => tt
580 | first :: rest => fun n fvs =>
581 match (eq_nat_dec n (length rest)) as Heq return
582 (envOf var (envT := first :: rest) fvs
583 -> var <| (if Heq
584 then match fvs with
585 | (true, _) => Some first
586 | (false, _) => None
587 end
588 else lookup_type n (snd fvs))) with
589 | left _ =>
590 match fvs return (envOf var (envT := first :: rest) fvs
591 -> var <| (match fvs with
592 | (true, _) => Some first
593 | (false, _) => None
594 end)) with
595 | (true, _) => fun env => fst env
596 | (false, _) => fun _ => tt
597 end
598 | right _ =>
599 match fvs return (envOf var (envT := first :: rest) fvs
600 -> var <| (lookup_type n (snd fvs))) with
601 | (true, fvs') => fun env => lookup var rest n fvs' (snd env)
602 | (false, fvs') => fun env => lookup var rest n fvs' env
603 end
604 end
605 end.
606
607 Theorem lok : forall var n t envT (fvs : isfree envT),
608 lookup_type n fvs = Some t
609 -> var <| lookup_type n fvs = var (ccType t).
610 crush.
611 Defined.
612 End isfree.
613
614 Implicit Arguments lookup_type [envT].
615 Implicit Arguments lookup [envT fvs].
616 Implicit Arguments wfExp [t envT].
617 Implicit Arguments envType [envT].
618 Implicit Arguments envOf [envT].
619 Implicit Arguments lok [var n t envT fvs].
620
621 Section lookup_hints.
622 Hint Resolve lookup_bound_contra.
623
624 (*Ltac my_chooser T k :=
625 match T with
626 | ptype =>
627 match goal with
628 | [ H : lookup _ _ = Some ?t |- _ ] => k t
629 end
630 | _ => default_chooser T k
631 end.
632
633 Ltac my_matching := matching equation my_chooser.*)
634
635 Hint Resolve lookup_bound_contra.
636
637 Lemma lookup_type_push : forall t' envT (fvs1 fvs2 : isfree envT) b1 b2,
638 (forall (n : nat) (t : Source.type),
639 lookup_type (envT := t' :: envT) n (b1, fvs1) = Some t ->
640 lookup_type (envT := t' :: envT) n (b2, fvs2) = Some t)
641 -> (forall (n : nat) (t : Source.type),
642 lookup_type n fvs1 = Some t ->
643 lookup_type n fvs2 = Some t).
644 intros until b2; intro H; intros n t;
645 generalize (H n t); my_crush; elimtype False; eauto.
646 Defined.
647
648 Lemma lookup_type_push_contra : forall t' envT (fvs1 fvs2 : isfree envT),
649 (forall (n : nat) (t : Source.type),
650 lookup_type (envT := t' :: envT) n (true, fvs1) = Some t ->
651 lookup_type (envT := t' :: envT) n (false, fvs2) = Some t)
652 -> False.
653 intros until fvs2; intro H; generalize (H (length envT) t'); my_crush.
654 Defined.
655 End lookup_hints.
656
657 Section packing.
658 Open Local Scope cc_scope.
659
660 Hint Resolve lookup_type_push lookup_type_push_contra.
661
662 Definition packExp (var : Closed.type -> Set) (envT : list Source.type)
663 (fvs1 fvs2 : isfree envT)
664 : (forall n t, lookup_type n fvs1 = Some t -> lookup_type n fvs2 = Some t)
665 -> envOf var fvs2 -> exp var (envType fvs1).
666 refine (fix packExp (var : Closed.type -> Set) (envT : list Source.type) {struct envT}
667 : forall fvs1 fvs2 : isfree envT,
668 (forall n t, lookup_type n fvs1 = Some t -> lookup_type n fvs2 = Some t)
669 -> envOf var fvs2 -> exp var (envType fvs1) :=
670 match envT return (forall fvs1 fvs2 : isfree envT,
671 (forall n t, lookup_type n fvs1 = Some t -> lookup_type n fvs2 = Some t)
672 -> envOf var fvs2
673 -> exp var (envType fvs1)) with
674 | nil => fun _ _ _ _ => ()
675 | first :: rest => fun fvs1 =>
676 match fvs1 return (forall fvs2 : isfree (first :: rest),
677 (forall n t, lookup_type (envT := first :: rest) n fvs1 = Some t
678 -> lookup_type n fvs2 = Some t)
679 -> envOf var fvs2
680 -> exp var (envType (envT := first :: rest) fvs1)) with
681 | (false, fvs1') => fun fvs2 =>
682 match fvs2 return ((forall n t, lookup_type (envT := first :: rest) n (false, fvs1') = Some t
683 -> lookup_type (envT := first :: rest) n fvs2 = Some t)
684 -> envOf (envT := first :: rest) var fvs2
685 -> exp var (envType (envT := first :: rest) (false, fvs1'))) with
686 | (false, fvs2') => fun Hmin env =>
687 packExp var _ fvs1' fvs2' _ env
688 | (true, fvs2') => fun Hmin env =>
689 packExp var _ fvs1' fvs2' _ (snd env)
690 end
691 | (true, fvs1') => fun fvs2 =>
692 match fvs2 return ((forall n t, lookup_type (envT := first :: rest) n (true, fvs1') = Some t
693 -> lookup_type (envT := first :: rest) n fvs2 = Some t)
694 -> envOf (envT := first :: rest) var fvs2
695 -> exp var (envType (envT := first :: rest) (true, fvs1'))) with
696 | (false, fvs2') => fun Hmin env =>
697 False_rect _ _
698 | (true, fvs2') => fun Hmin env =>
699 [#(fst env), packExp var _ fvs1' fvs2' _ (snd env)]
700 end
701 end
702 end); eauto.
703 Defined.
704
705 Hint Resolve fvsExp_correct fvsExp_minimal.
706 Hint Resolve lookup_push_drop lookup_bound lookup_push_add.
707
708 Implicit Arguments packExp [var envT].
709
710 Fixpoint unpackExp (var : Closed.type -> Set) t (envT : list Source.type) {struct envT}
711 : forall fvs : isfree envT,
712 exp var (envType fvs)
713 -> (envOf var fvs -> exp var t) -> exp var t :=
714 match envT return (forall fvs : isfree envT,
715 exp var (envType fvs)
716 -> (envOf var fvs -> exp var t) -> exp var t) with
717 | nil => fun _ _ f => f tt
718 | first :: rest => fun fvs =>
719 match fvs return (exp var (envType (envT := first :: rest) fvs)
720 -> (envOf var (envT := first :: rest) fvs -> exp var t)
721 -> exp var t) with
722 | (false, fvs') => fun p f =>
723 unpackExp rest fvs' p f
724 | (true, fvs') => fun p f =>
725 x <- #1 p;
726 unpackExp rest fvs' (#2 p)
727 (fun env => f (x, env))
728 end
729 end.
730
731 Implicit Arguments unpackExp [var t envT fvs].
732
733 Theorem wfExp_lax : forall t t' envT (fvs : isfree envT) (e : Source.exp natvar t),
734 wfExp (envT := t' :: envT) (true, fvs) e
735 -> wfExp (envT := t' :: envT) (true, snd (fvsExp e (t' :: envT))) e.
736 Hint Extern 1 (_ = _) =>
737 match goal with
738 | [ H : lookup_type _ (fvsExp ?X ?Y) = _ |- _ ] =>
739 destruct (fvsExp X Y); my_crush
740 end.
741 eauto.
742 Defined.
743
744 Implicit Arguments wfExp_lax [t t' envT fvs e].
745
746 Lemma inclusion : forall t t' envT fvs (e : Source.exp natvar t),
747 wfExp (envT := t' :: envT) (true, fvs) e
748 -> (forall n t, lookup_type n (snd (fvsExp e (t' :: envT))) = Some t
749 -> lookup_type n fvs = Some t).
750 eauto.
751 Defined.
752
753 Implicit Arguments inclusion [t t' envT fvs e].
754
755 Definition env_prog var t envT (fvs : isfree envT) :=
756 funcs var (envOf var fvs -> Closed.exp var t).
757
758 Implicit Arguments env_prog [envT].
759
760 Axiom cheat : forall T, T.
761
762 Import Source.
763 Open Local Scope cc_scope.
764
765 Fixpoint ccExp var t (e : Source.exp natvar t)
766 (envT : list Source.type) (fvs : isfree envT)
767 {struct e} : wfExp fvs e -> env_prog var (ccType t) fvs :=
768 match e in (Source.exp _ t) return (wfExp fvs e -> env_prog var (ccType t) fvs) with
769 | Const n => fun _ => Main (fun _ => ^n)
770 | Plus e1 e2 => fun wf =>
771 n1 <-- ccExp var e1 _ fvs (proj1 wf);
772 n2 <-- ccExp var e2 _ fvs (proj2 wf);
773 Main (fun env => n1 env +^ n2 env)
774
775 | Var _ n => fun wf =>
776 Main (fun env => #(match lok wf in _ = T return T with
777 | refl_equal => lookup var n env
778 end))
779
780 | App _ _ f x => fun wf =>
781 f' <-- ccExp var f _ fvs (proj1 wf);
782 x' <-- ccExp var x _ fvs (proj2 wf);
783 Main (fun env => f' env @ x' env)
784
785 | Abs dom _ b => fun wf =>
786 b' <-- ccExp var (b (length envT)) (dom :: envT) _ (wfExp_lax wf);
787 f <== \\env, arg, unpackExp (#env) (fun env => b' (arg, env));
788 Main (fun env => #f ##
789 packExp
790 (snd (fvsExp (b (length envT)) (dom :: envT)))
791 fvs (inclusion wf) env)
792 end.
793 End packing.
794
795 Implicit Arguments packExp [var envT].
796 Implicit Arguments unpackExp [var t envT fvs].
797
798 Implicit Arguments ccExp [var t envT].
799
800 Fixpoint map_funcs var result T1 T2 (f : T1 -> T2) (fs : cfuncs var result T1) {struct fs}
801 : cfuncs var result T2 :=
802 match fs with
803 | CMain v => CMain (f v)
804 | CAbs _ _ e fs' => CAbs e (fun x => map_funcs f (fs' x))
805 end.
806
807 Definition CcTerm' result (E : Pterm result) (Hwf : wfTerm (envT := nil) tt (E _)) : Cprog result :=
808 fun _ => map_funcs (fun f => f tt) (ccTerm (E _) (envT := nil) tt Hwf).
809
810
811 (** * Correctness *)
812
813 Scheme pterm_equiv_mut := Minimality for pterm_equiv Sort Prop
814 with pprimop_equiv_mut := Minimality for pprimop_equiv Sort Prop.
815
816 Section splicePrim_correct.
817 Variables result t t' : ptype.
818 Variable ps' : ctypeDenote ([< t >]) -> cprimops ctypeDenote t'.
819
820 Theorem splicePrim_correct : forall (ps : cprimops ctypeDenote t),
821 cprimopsDenote (splicePrim ps ps')
822 = cprimopsDenote (ps' (cprimopsDenote ps)).
823 induction ps; equation.
824 Qed.
825 End splicePrim_correct.
826
827 Section spliceTerm_correct.
828 Variables result t : ptype.
829 Variable e : ctypeDenote ([< t >]) -> cterm ctypeDenote result.
830 Variable k : ptypeDenote result -> bool.
831
832 Theorem spliceTerm_correct : forall (ps : cprimops ctypeDenote t),
833 ctermDenote (spliceTerm ps e) k
834 = ctermDenote (e (cprimopsDenote ps)) k.
835 induction ps; equation.
836 Qed.
837 End spliceTerm_correct.
838
839 Section spliceFuncs'_correct.
840 Variable result : ptype.
841 Variables T1 T2 : Type.
842 Variable f : T1 -> T2.
843 Variable k : ptypeDenote result -> bool.
844
845 Theorem spliceFuncs'_correct : forall fs,
846 cfuncsDenote (spliceFuncs' fs f) k
847 = f (cfuncsDenote fs k).
848 induction fs; equation.
849 Qed.
850 End spliceFuncs'_correct.
851
852 Section spliceFuncs_correct.
853 Variable result : ptype.
854 Variables T1 T2 T3 : Type.
855 Variable fs2 : cfuncs ctypeDenote result T2.
856 Variable f : T1 -> T2 -> T3.
857 Variable k : ptypeDenote result -> bool.
858
859 Hint Rewrite spliceFuncs'_correct : ltamer.
860
861 Theorem spliceFuncs_correct : forall fs1,
862 cfuncsDenote (spliceFuncs fs1 fs2 f) k
863 = f (cfuncsDenote fs1 k) (cfuncsDenote fs2 k).
864 induction fs1; equation.
865 Qed.
866 End spliceFuncs_correct.
867
868 Section inside_correct.
869 Variable result : ptype.
870 Variables T1 T2 : Type.
871 Variable fs2 : T1 -> cfuncs ctypeDenote result T2.
872 Variable k : ptypeDenote result -> bool.
873
874 Theorem inside_correct : forall fs1,
875 cfuncsDenote (inside fs1 fs2) k
876 = cfuncsDenote (fs2 (cfuncsDenote fs1 k)) k.
877 induction fs1; equation.
878 Qed.
879 End inside_correct.
880
881
882 Notation "var <| to" := (match to with
883 | None => unit
884 | Some t => var ([< t >])%cc
885 end) (no associativity, at level 70).
886
887 Section packing_correct.
888 Variable result : ptype.
889
890 Hint Rewrite splicePrim_correct spliceTerm_correct : ltamer.
891
892 Ltac my_matching := matching my_equation default_chooser.
893
894 Fixpoint makeEnv (envT : list ptype) : forall (fvs : isfree envT),
895 ptypeDenote (envType fvs)
896 -> envOf ctypeDenote fvs :=
897 match envT return (forall (fvs : isfree envT),
898 ptypeDenote (envType fvs)
899 -> envOf ctypeDenote fvs) with
900 | nil => fun _ _ => tt
901 | first :: rest => fun fvs =>
902 match fvs return (ptypeDenote (envType (envT := first :: rest) fvs)
903 -> envOf (envT := first :: rest) ctypeDenote fvs) with
904 | (false, fvs') => fun env => makeEnv rest fvs' env
905 | (true, fvs') => fun env => (fst env, makeEnv rest fvs' (snd env))
906 end
907 end.
908
909 Theorem unpackExp_correct : forall (envT : list ptype) (fvs : isfree envT)
910 (ps : cprimops ctypeDenote (envType fvs))
911 (e : envOf ctypeDenote fvs -> cterm ctypeDenote result) k,
912 ctermDenote (unpackExp ps e) k
913 = ctermDenote (e (makeEnv _ _ (cprimopsDenote ps))) k.
914 induction envT; my_matching.
915 Qed.
916
917 Lemma lookup_type_more : forall v2 envT (fvs : isfree envT) t b v,
918 (v2 = length envT -> False)
919 -> lookup_type v2 (envT := t :: envT) (b, fvs) = v
920 -> lookup_type v2 fvs = v.
921 equation.
922 Qed.
923
924 Lemma lookup_type_less : forall v2 t envT (fvs : isfree (t :: envT)) v,
925 (v2 = length envT -> False)
926 -> lookup_type v2 (snd fvs) = v
927 -> lookup_type v2 (envT := t :: envT) fvs = v.
928 equation.
929 Qed.
930
931 Lemma lookup_bound_contra_eq : forall t envT (fvs : isfree envT) v,
932 lookup_type v fvs = Some t
933 -> v = length envT
934 -> False.
935 simpler; eapply lookup_bound_contra; eauto.
936 Defined.
937
938 Lemma lookup_type_inner : forall result t envT v t' (fvs : isfree envT) e,
939 wfTerm (envT := t :: envT) (true, fvs) e
940 -> lookup_type v (snd (fvsTerm (result := result) e (t :: envT))) = Some t'
941 -> lookup_type v fvs = Some t'.
942 Hint Resolve lookup_bound_contra_eq fvsTerm_minimal
943 lookup_type_more lookup_type_less.
944 Hint Extern 2 (Some _ = Some _) => contradictory.
945
946 eauto 6.
947 Qed.
948
949 Lemma cast_irrel : forall T1 T2 x (H1 H2 : T1 = T2),
950 (x :? H1) = (x :? H2).
951 equation.
952 Qed.
953
954 Hint Immediate cast_irrel.
955
956 Lemma cast_irrel_unit : forall T1 T2 x y (H1 : T1 = T2) (H2 : unit = T2),
957 (x :? H1) = (y :? H2).
958 intros; generalize H1;
959 rewrite <- H2 in H1;
960 equation.
961 Qed.
962
963 Hint Immediate cast_irrel_unit.
964
965 Hint Extern 3 (_ = _) =>
966 match goal with
967 | [ |- context[False_rect _ ?H] ] =>
968 apply False_rect; exact H
969 end.
970
971 Theorem packExp_correct : forall v2 t envT (fvs1 fvs2 : isfree envT)
972 Heq1 (Heq2 : ctypeDenote <| lookup_type v2 fvs2 = ptypeDenote t)
973 Hincl env,
974 (lookup ctypeDenote v2 env :? Heq2)
975 = (lookup ctypeDenote v2
976 (makeEnv envT fvs1
977 (cprimopsDenote
978 (packExp fvs1 fvs2 Hincl env))) :? Heq1).
979 induction envT; my_equation.
980 Qed.
981 End packing_correct.
982
983 Lemma look : forall envT n (fvs : isfree envT) t,
984 lookup_type n fvs = Some t
985 -> ctypeDenote <| lookup_type n fvs = ptypeDenote t.
986 equation.
987 Qed.
988
989 Implicit Arguments look [envT n fvs t].
990
991
992 Theorem ccTerm_correct : forall result G
993 (e1 : pterm ptypeDenote result)
994 (e2 : pterm natvar result),
995 pterm_equiv G e1 e2
996 -> forall (envT : list ptype) (fvs : isfree envT)
997 (env : envOf ctypeDenote fvs) (Hwf : wfTerm fvs e2) k,
998 (forall t (v1 : ptypeDenote t) (v2 : natvar t),
999 In (vars (x := t) (v1, v2)) G
1000 -> v2 < length envT)
1001 -> (forall t (v1 : ptypeDenote t) (v2 : natvar t),
1002 In (vars (x := t) (v1, v2)) G
1003 -> lookup_type v2 fvs = Some t
1004 -> forall Heq, (lookup ctypeDenote v2 env :? Heq) = v1)
1005 -> ptermDenote e1 k
1006 = ctermDenote (cfuncsDenote (ccTerm e2 fvs Hwf) k env) k.
1007
1008 Hint Rewrite splicePrim_correct spliceTerm_correct
1009 spliceFuncs_correct inside_correct : ltamer.
1010
1011 Hint Rewrite unpackExp_correct : ltamer.
1012 Hint Resolve packExp_correct lookup_type_inner.
1013
1014 Hint Extern 1 (_ = _) => push_vars.
1015
1016 Hint Extern 1 (_ = _) =>
1017 match goal with
1018 | [ Hvars : forall t v1 v2, _,
1019 Hin : In _ _ |- _ ] =>
1020 rewrite (Hvars _ _ _ Hin)
1021 end.
1022
1023 Hint Extern 1 (wfPrimop _ _) => tauto.
1024
1025 Hint Extern 1 (_ < _) =>
1026 match goal with
1027 | [ Hvars : forall t v1 v2, _,
1028 Hin : In _ _ |- _ ] =>
1029 exact (Hvars _ _ _ Hin)
1030 end.
1031
1032 Hint Extern 1 (lookup_type _ _ = _) => tauto.
1033
1034 Hint Extern 1 (_ = _) =>
1035 match goal with
1036 | [ Hvars : forall t v1 v2, _,
1037 Hin : In (vars (_, length ?envT)) _ |- _ ] =>
1038 contradictory; assert (length envT < length envT); [auto | omega]
1039 end.
1040
1041 Hint Extern 5 (_ = _) => symmetry.
1042
1043 Hint Extern 1 (_ = _) =>
1044 match goal with
1045 | [ H : lookup_type ?v _ = Some ?t, fvs : isfree _ |- (lookup _ _ _ :? _) = _ ] =>
1046 let Hty := fresh "Hty" in
1047 (assert (Hty : lookup_type v fvs = Some t);
1048 [eauto
1049 | eapply (trans_cast (look Hty))])
1050 end.
1051
1052 Hint Extern 3 (ptermDenote _ _ = ctermDenote _ _) =>
1053 match goal with
1054 | [ H : _ |- ptermDenote (_ ?v1) _
1055 = ctermDenote (cfuncsDenote (ccTerm (_ ?v2) (envT := ?envT) ?fvs _) _ _) _ ] =>
1056 apply (H v1 v2 envT fvs); my_simpler
1057 end.
1058
1059 intro.
1060 apply (pterm_equiv_mut
1061 (fun G (e1 : pterm ptypeDenote result) (e2 : pterm natvar result) =>
1062 forall (envT : list ptype) (fvs : isfree envT)
1063 (env : envOf ctypeDenote fvs) (Hwf : wfTerm fvs e2) k,
1064 (forall t (v1 : ptypeDenote t) (v2 : natvar t),
1065 In (vars (x := t) (v1, v2)) G
1066 -> v2 < length envT)
1067 -> (forall t (v1 : ptypeDenote t) (v2 : natvar t),
1068 In (vars (x := t) (v1, v2)) G
1069 -> lookup_type v2 fvs = Some t
1070 -> forall Heq, (lookup ctypeDenote v2 env :? Heq) = v1)
1071 -> ptermDenote e1 k
1072 = ctermDenote (cfuncsDenote (ccTerm e2 fvs Hwf) k env) k)
1073 (fun G t (p1 : pprimop ptypeDenote result t) (p2 : pprimop natvar result t) =>
1074 forall (envT : list ptype) (fvs : isfree envT)
1075 (env : envOf ctypeDenote fvs) (Hwf : wfPrimop fvs p2) Hwf k,
1076 (forall t (v1 : ptypeDenote t) (v2 : natvar t),
1077 In (vars (x := t) (v1, v2)) G
1078 -> v2 < length envT)
1079 -> (forall t (v1 : ptypeDenote t) (v2 : natvar t),
1080 In (vars (x := t) (v1, v2)) G
1081 -> lookup_type v2 fvs = Some t
1082 -> forall Heq, (lookup ctypeDenote v2 env :? Heq) = v1)
1083 -> pprimopDenote p1 k
1084 = cprimopsDenote (cfuncsDenote (ccPrimop p2 fvs Hwf) k env)));
1085 my_simpler; eauto.
1086 Qed.
1087
1088
1089 (** * Parametric version *)
1090
1091 Section wf.
1092 Variable result : ptype.
1093
1094 Lemma Pterm_wf' : forall G (e1 e2 : pterm natvar result),
1095 pterm_equiv G e1 e2
1096 -> forall envT (fvs : isfree envT),
1097 (forall t (v1 v2 : natvar t), In (vars (v1, v2)) G
1098 -> lookup_type v1 fvs = Some t)
1099 -> wfTerm fvs e1.
1100 Hint Extern 3 (Some _ = Some _) => contradictory; eapply lookup_bound_contra; eauto.
1101
1102 apply (pterm_equiv_mut
1103 (fun G (e1 e2 : pterm natvar result) =>
1104 forall envT (fvs : isfree envT),
1105 (forall t (v1 v2 : natvar t), In (vars (v1, v2)) G
1106 -> lookup_type v1 fvs = Some t)
1107 -> wfTerm (envT:=envT) fvs e1)
1108 (fun G t (p1 p2 : pprimop natvar result t) =>
1109 forall envT (fvs : isfree envT),
1110 (forall t (v1 v2 : natvar t), In (vars (v1, v2)) G
1111 -> lookup_type v1 fvs = Some t)
1112 -> wfPrimop (envT:=envT) fvs p1));
1113 simpler;
1114 match goal with
1115 | [ envT : list ptype, H : _ |- _ ] =>
1116 apply (H (length envT) (length envT)); simpler
1117 end.
1118 Qed.
1119
1120 Theorem Pterm_wf : forall (E : Pterm result),
1121 wfTerm (envT := nil) tt (E _).
1122 intros; eapply Pterm_wf';
1123 [apply Pterm_equiv
1124 | simpler].
1125 Qed.
1126 End wf.
1127
1128 Definition CcTerm result (E : Pterm result) : Cprog result :=
1129 CcTerm' E (Pterm_wf E).
1130
1131 Lemma map_funcs_correct : forall result T1 T2 (f : T1 -> T2) (fs : cfuncs ctypeDenote result T1) k,
1132 cfuncsDenote (map_funcs f fs) k = f (cfuncsDenote fs k).
1133 induction fs; equation.
1134 Qed.
1135
1136 Theorem CcTerm_correct : forall result (E : Pterm result) k,
1137 PtermDenote E k
1138 = CprogDenote (CcTerm E) k.
1139 Hint Rewrite map_funcs_correct : ltamer.
1140
1141 unfold PtermDenote, CprogDenote, CcTerm, CcTerm', cprogDenote;
1142 simpler;
1143 apply (ccTerm_correct (result := result)
1144 (G := nil)
1145 (e1 := E _)
1146 (e2 := E _)
1147 (Pterm_equiv _ _ _)
1148 nil
1149 tt
1150 tt);
1151 simpler.
1152 Qed.