annotate src/Intensional.v @ 257:108ec446fbaf

Easy direction of Intensional
author Adam Chlipala <adamc@hcoop.net>
date Wed, 16 Dec 2009 15:54:50 -0500
parents 2a34c4dc6a10
children 4c9031b62cd0
rev   line source
adamc@223 1 (* Copyright (c) 2008-2009, Adam Chlipala
adamc@182 2 *
adamc@182 3 * This work is licensed under a
adamc@182 4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
adamc@182 5 * Unported License.
adamc@182 6 * The license text is available at:
adamc@182 7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
adamc@182 8 *)
adamc@182 9
adamc@257 10 (* begin hide *)
adamc@257 11 Require Import List.
adamc@257 12
adamc@257 13 Require Import Axioms DepList Tactics.
adamc@257 14
adamc@257 15 Set Implicit Arguments.
adamc@257 16 (* end hide *)
adamc@257 17
adamc@184 18
adamc@182 19 (** %\chapter{Intensional Transformations}% *)
adamc@182 20
adamc@257 21 (* begin hide *)
adamc@257 22
adamc@257 23 Inductive type : Type :=
adamc@257 24 | Nat : type
adamc@257 25 | Arrow : type -> type -> type.
adamc@257 26
adamc@257 27 Infix "-->" := Arrow (right associativity, at level 60).
adamc@257 28
adamc@257 29 Fixpoint typeDenote (t : type) : Set :=
adamc@257 30 match t with
adamc@257 31 | Nat => nat
adamc@257 32 | t1 --> t2 => typeDenote t1 -> typeDenote t2
adamc@257 33 end.
adamc@257 34
adamc@257 35 Module Phoas.
adamc@257 36 Section vars.
adamc@257 37 Variable var : type -> Type.
adamc@257 38
adamc@257 39 Inductive exp : type -> Type :=
adamc@257 40 | Var : forall t,
adamc@257 41 var t
adamc@257 42 -> exp t
adamc@257 43
adamc@257 44 | Const : nat -> exp Nat
adamc@257 45 | Plus : exp Nat -> exp Nat -> exp Nat
adamc@257 46
adamc@257 47 | App : forall t1 t2,
adamc@257 48 exp (t1 --> t2)
adamc@257 49 -> exp t1
adamc@257 50 -> exp t2
adamc@257 51 | Abs : forall t1 t2,
adamc@257 52 (var t1 -> exp t2)
adamc@257 53 -> exp (t1 --> t2).
adamc@257 54 End vars.
adamc@257 55
adamc@257 56 Definition Exp t := forall var, exp var t.
adamc@257 57
adamc@257 58 Implicit Arguments Var [var t].
adamc@257 59 Implicit Arguments Const [var].
adamc@257 60 Implicit Arguments Plus [var].
adamc@257 61 Implicit Arguments App [var t1 t2].
adamc@257 62 Implicit Arguments Abs [var t1 t2].
adamc@257 63
adamc@257 64 Notation "# v" := (Var v) (at level 70).
adamc@257 65
adamc@257 66 Notation "^ n" := (Const n) (at level 70).
adamc@257 67 Infix "+^" := Plus (left associativity, at level 79).
adamc@257 68
adamc@257 69 Infix "@" := App (left associativity, at level 77).
adamc@257 70 Notation "\ x , e" := (Abs (fun x => e)) (at level 78).
adamc@257 71 Notation "\ ! , e" := (Abs (fun _ => e)) (at level 78).
adamc@257 72
adamc@257 73 Fixpoint expDenote t (e : exp typeDenote t) : typeDenote t :=
adamc@257 74 match e with
adamc@257 75 | Var _ v => v
adamc@257 76
adamc@257 77 | Const n => n
adamc@257 78 | Plus e1 e2 => expDenote e1 + expDenote e2
adamc@257 79
adamc@257 80 | App _ _ e1 e2 => (expDenote e1) (expDenote e2)
adamc@257 81 | Abs _ _ e' => fun x => expDenote (e' x)
adamc@257 82 end.
adamc@257 83
adamc@257 84 Definition ExpDenote t (e : Exp t) := expDenote (e _).
adamc@257 85
adamc@257 86 Section exp_equiv.
adamc@257 87 Variables var1 var2 : type -> Type.
adamc@257 88
adamc@257 89 Inductive exp_equiv : list { t : type & var1 t * var2 t }%type
adamc@257 90 -> forall t, exp var1 t -> exp var2 t -> Prop :=
adamc@257 91 | EqVar : forall G t (v1 : var1 t) v2,
adamc@257 92 In (existT _ t (v1, v2)) G
adamc@257 93 -> exp_equiv G (#v1) (#v2)
adamc@257 94
adamc@257 95 | EqConst : forall G n,
adamc@257 96 exp_equiv G (^n) (^n)
adamc@257 97 | EqPlus : forall G x1 y1 x2 y2,
adamc@257 98 exp_equiv G x1 x2
adamc@257 99 -> exp_equiv G y1 y2
adamc@257 100 -> exp_equiv G (x1 +^ y1) (x2 +^ y2)
adamc@257 101
adamc@257 102 | EqApp : forall G t1 t2 (f1 : exp _ (t1 --> t2)) (x1 : exp _ t1) f2 x2,
adamc@257 103 exp_equiv G f1 f2
adamc@257 104 -> exp_equiv G x1 x2
adamc@257 105 -> exp_equiv G (f1 @ x1) (f2 @ x2)
adamc@257 106 | EqAbs : forall G t1 t2 (f1 : var1 t1 -> exp var1 t2) f2,
adamc@257 107 (forall v1 v2, exp_equiv (existT _ t1 (v1, v2) :: G) (f1 v1) (f2 v2))
adamc@257 108 -> exp_equiv G (Abs f1) (Abs f2).
adamc@257 109 End exp_equiv.
adamc@257 110 End Phoas.
adamc@257 111 (* end hide *)
adamc@257 112
adamc@257 113 Module DeBruijn.
adamc@257 114 Inductive exp : list type -> type -> Type :=
adamc@257 115 | Var : forall G t,
adamc@257 116 member t G
adamc@257 117 -> exp G t
adamc@257 118
adamc@257 119 | Const : forall G, nat -> exp G Nat
adamc@257 120 | Plus : forall G, exp G Nat -> exp G Nat -> exp G Nat
adamc@257 121
adamc@257 122 | App : forall G t1 t2,
adamc@257 123 exp G (t1 --> t2)
adamc@257 124 -> exp G t1
adamc@257 125 -> exp G t2
adamc@257 126 | Abs : forall G t1 t2,
adamc@257 127 exp (t1 :: G) t2
adamc@257 128 -> exp G (t1 --> t2).
adamc@257 129
adamc@257 130 Implicit Arguments Const [G].
adamc@257 131
adamc@257 132 Fixpoint expDenote G t (e : exp G t) : hlist typeDenote G -> typeDenote t :=
adamc@257 133 match e with
adamc@257 134 | Var _ _ v => fun s => hget s v
adamc@257 135
adamc@257 136 | Const _ n => fun _ => n
adamc@257 137 | Plus _ e1 e2 => fun s => expDenote e1 s + expDenote e2 s
adamc@257 138
adamc@257 139 | App _ _ _ e1 e2 => fun s => (expDenote e1 s) (expDenote e2 s)
adamc@257 140 | Abs _ _ _ e' => fun s x => expDenote e' (x ::: s)
adamc@257 141 end.
adamc@257 142 End DeBruijn.
adamc@257 143
adamc@257 144 Import Phoas DeBruijn.
adamc@257 145
adamc@257 146
adamc@257 147 (** * From De Bruijn to PHOAS *)
adamc@257 148
adamc@257 149 Section phoasify.
adamc@257 150 Variable var : type -> Type.
adamc@257 151
adamc@257 152 Fixpoint phoasify G t (e : DeBruijn.exp G t) : hlist var G -> Phoas.exp var t :=
adamc@257 153 match e with
adamc@257 154 | Var _ _ v => fun s => #(hget s v)
adamc@257 155
adamc@257 156 | Const _ n => fun _ => ^n
adamc@257 157 | Plus _ e1 e2 => fun s => phoasify e1 s +^ phoasify e2 s
adamc@257 158
adamc@257 159 | App _ _ _ e1 e2 => fun s => phoasify e1 s @ phoasify e2 s
adamc@257 160 | Abs _ _ _ e' => fun s => \x, phoasify e' (x ::: s)
adamc@257 161 end.
adamc@257 162 End phoasify.
adamc@257 163
adamc@257 164 Definition Phoasify t (e : DeBruijn.exp nil t) : Phoas.Exp t :=
adamc@257 165 fun _ => phoasify e HNil.
adamc@257 166
adamc@257 167 Theorem phoasify_sound : forall G t (e : DeBruijn.exp G t) s,
adamc@257 168 Phoas.expDenote (phoasify e s) = DeBruijn.expDenote e s.
adamc@257 169 induction e; crush; ext_eq; crush.
adamc@257 170 Qed.
adamc@257 171
adamc@257 172 Section vars.
adamc@257 173 Variables var1 var2 : type -> Type.
adamc@257 174
adamc@257 175 Fixpoint zip G (s1 : hlist var1 G) : hlist var2 G -> list {t : type & var1 t * var2 t}%type :=
adamc@257 176 match s1 with
adamc@257 177 | HNil => fun _ => nil
adamc@257 178 | HCons _ _ v1 s1' => fun s2 => existT _ _ (v1, hhd s2) :: zip s1' (htl s2)
adamc@257 179 end.
adamc@257 180
adamc@257 181 Lemma In_zip : forall t G (s1 : hlist _ G) s2 (m : member t G),
adamc@257 182 In (existT _ t (hget s1 m, hget s2 m)) (zip s1 s2).
adamc@257 183 induction s1; intro s2; dep_destruct s2; intro m; dep_destruct m; crush.
adamc@257 184 Qed.
adamc@257 185
adamc@257 186 Lemma unsimpl_zip : forall t (v1 : var1 t) (v2 : var2 t)
adamc@257 187 G (s1 : hlist _ G) s2 t' (e1 : Phoas.exp _ t') e2,
adamc@257 188 exp_equiv (zip (v1 ::: s1) (v2 ::: s2)) e1 e2
adamc@257 189 -> exp_equiv (existT _ _ (v1, v2) :: zip s1 s2) e1 e2.
adamc@257 190 trivial.
adamc@257 191 Qed.
adamc@257 192
adamc@257 193 Hint Resolve In_zip unsimpl_zip.
adamc@257 194
adamc@257 195 Theorem phoasify_wf : forall G t (e : DeBruijn.exp G t) s1 s2,
adamc@257 196 exp_equiv (zip s1 s2) (phoasify e s1) (phoasify e s2).
adamc@257 197 Hint Constructors exp_equiv.
adamc@257 198
adamc@257 199 induction e; crush.
adamc@257 200 Qed.
adamc@257 201 End vars.
adamc@257 202
adamc@257 203
adamc@257 204 (** * From PHOAS to De Bruijn *)