adamc@152
|
1 (* Copyright (c) 2008, Adam Chlipala
|
adamc@152
|
2 *
|
adamc@152
|
3 * This work is licensed under a
|
adamc@152
|
4 * Creative Commons Attribution-Noncommercial-No Derivative Works 3.0
|
adamc@152
|
5 * Unported License.
|
adamc@152
|
6 * The license text is available at:
|
adamc@152
|
7 * http://creativecommons.org/licenses/by-nc-nd/3.0/
|
adamc@152
|
8 *)
|
adamc@152
|
9
|
adamc@152
|
10 (* begin hide *)
|
adamc@152
|
11 Require Import List String.
|
adamc@152
|
12
|
adamc@152
|
13 Require Import Tactics.
|
adamc@152
|
14
|
adamc@152
|
15 Set Implicit Arguments.
|
adamc@152
|
16 (* end hide *)
|
adamc@152
|
17
|
adamc@152
|
18
|
adamc@153
|
19 (** %\part{Formalizing Programming Languages and Compilers}
|
adamc@152
|
20
|
adamc@153
|
21 \chapter{First-Order Abstract Syntax}% *)
|
adamc@152
|
22
|
adamc@152
|
23 (** TODO: Prose for this chapter *)
|
adamc@152
|
24
|
adamc@152
|
25
|
adamc@152
|
26 (** * Concrete Binding *)
|
adamc@152
|
27
|
adamc@152
|
28 Module Concrete.
|
adamc@152
|
29
|
adamc@152
|
30 Definition var := string.
|
adamc@152
|
31 Definition var_eq := string_dec.
|
adamc@152
|
32
|
adamc@152
|
33 Inductive exp : Set :=
|
adamc@152
|
34 | Const : bool -> exp
|
adamc@152
|
35 | Var : var -> exp
|
adamc@152
|
36 | App : exp -> exp -> exp
|
adamc@152
|
37 | Abs : var -> exp -> exp.
|
adamc@152
|
38
|
adamc@152
|
39 Inductive type : Set :=
|
adamc@152
|
40 | Bool : type
|
adamc@152
|
41 | Arrow : type -> type -> type.
|
adamc@152
|
42
|
adamc@152
|
43 Infix "-->" := Arrow (right associativity, at level 60).
|
adamc@152
|
44
|
adamc@152
|
45 Definition ctx := list (var * type).
|
adamc@152
|
46
|
adamc@152
|
47 Reserved Notation "G |-v x : t" (no associativity, at level 90, x at next level).
|
adamc@152
|
48
|
adamc@152
|
49 Inductive lookup : ctx -> var -> type -> Prop :=
|
adamc@152
|
50 | First : forall x t G,
|
adamc@152
|
51 (x, t) :: G |-v x : t
|
adamc@152
|
52 | Next : forall x t x' t' G,
|
adamc@152
|
53 x <> x'
|
adamc@152
|
54 -> G |-v x : t
|
adamc@152
|
55 -> (x', t') :: G |-v x : t
|
adamc@152
|
56
|
adamc@152
|
57 where "G |-v x : t" := (lookup G x t).
|
adamc@152
|
58
|
adamc@152
|
59 Hint Constructors lookup.
|
adamc@152
|
60
|
adamc@152
|
61 Reserved Notation "G |-e e : t" (no associativity, at level 90, e at next level).
|
adamc@152
|
62
|
adamc@152
|
63 Inductive hasType : ctx -> exp -> type -> Prop :=
|
adamc@152
|
64 | TConst : forall G b,
|
adamc@152
|
65 G |-e Const b : Bool
|
adamc@152
|
66 | TVar : forall G v t,
|
adamc@152
|
67 G |-v v : t
|
adamc@152
|
68 -> G |-e Var v : t
|
adamc@152
|
69 | TApp : forall G e1 e2 dom ran,
|
adamc@152
|
70 G |-e e1 : dom --> ran
|
adamc@152
|
71 -> G |-e e2 : dom
|
adamc@152
|
72 -> G |-e App e1 e2 : ran
|
adamc@152
|
73 | TAbs : forall G x e' dom ran,
|
adamc@152
|
74 (x, dom) :: G |-e e' : ran
|
adamc@152
|
75 -> G |-e Abs x e' : dom --> ran
|
adamc@152
|
76
|
adamc@152
|
77 where "G |-e e : t" := (hasType G e t).
|
adamc@152
|
78
|
adamc@152
|
79 Hint Constructors hasType.
|
adamc@152
|
80
|
adamc@152
|
81 Notation "x ## G" := (forall t' : type, In (x, t') G -> False) (no associativity, at level 90).
|
adamc@152
|
82
|
adamc@152
|
83 Notation "G' # G" := (forall (x : var) (t : type), In (x, t) G -> x ## G') (no associativity, at level 90).
|
adamc@152
|
84
|
adamc@152
|
85 Lemma lookup_In : forall G x t,
|
adamc@152
|
86 G |-v x : t
|
adamc@152
|
87 -> In (x, t) G.
|
adamc@152
|
88 induction 1; crush.
|
adamc@152
|
89 Qed.
|
adamc@152
|
90
|
adamc@152
|
91 Hint Resolve lookup_In.
|
adamc@152
|
92
|
adamc@152
|
93 Lemma disjoint_invert1 : forall G x t G' x' t',
|
adamc@152
|
94 G |-v x : t
|
adamc@152
|
95 -> (x', t') :: G' # G
|
adamc@152
|
96 -> x <> x'.
|
adamc@152
|
97 crush; eauto.
|
adamc@152
|
98 Qed.
|
adamc@152
|
99
|
adamc@152
|
100 Lemma disjoint_invert2 : forall G' G p,
|
adamc@152
|
101 p :: G' # G
|
adamc@152
|
102 -> G' # G.
|
adamc@152
|
103 firstorder.
|
adamc@152
|
104 Qed.
|
adamc@152
|
105
|
adamc@152
|
106 Hint Resolve disjoint_invert1 disjoint_invert2.
|
adamc@152
|
107 Hint Extern 1 (_ <> _) => (intro; subst).
|
adamc@152
|
108
|
adamc@152
|
109 Lemma weaken_lookup' : forall G x t,
|
adamc@152
|
110 G |-v x : t
|
adamc@152
|
111 -> forall G', G' # G
|
adamc@152
|
112 -> G' ++ G |-v x : t.
|
adamc@152
|
113 induction G' as [ | [x' t'] tl ]; crush; eauto 9.
|
adamc@152
|
114 Qed.
|
adamc@152
|
115
|
adamc@152
|
116 Lemma weaken_lookup : forall G2 x t G',
|
adamc@152
|
117 G' # G2
|
adamc@152
|
118 -> forall G1, G1 ++ G2 |-v x : t
|
adamc@152
|
119 -> G1 ++ G' ++ G2 |-v x : t.
|
adamc@152
|
120 Hint Resolve weaken_lookup'.
|
adamc@152
|
121
|
adamc@152
|
122 induction G1 as [ | [x' t'] tl ]; crush;
|
adamc@152
|
123 match goal with
|
adamc@152
|
124 | [ H : _ |-v _ : _ |- _ ] => inversion H; crush
|
adamc@152
|
125 end.
|
adamc@152
|
126 Qed.
|
adamc@152
|
127
|
adamc@152
|
128 Hint Resolve weaken_lookup.
|
adamc@152
|
129
|
adamc@152
|
130 Lemma hasType_push : forall x t G1 G2 e t',
|
adamc@152
|
131 ((x, t) :: G1) ++ G2 |-e e : t'
|
adamc@152
|
132 -> (x, t) :: G1 ++ G2 |-e e : t'.
|
adamc@152
|
133 trivial.
|
adamc@152
|
134 Qed.
|
adamc@152
|
135
|
adamc@152
|
136 Hint Resolve hasType_push.
|
adamc@152
|
137
|
adamc@152
|
138 Theorem weaken_hasType' : forall G' G e t,
|
adamc@152
|
139 G |-e e : t
|
adamc@152
|
140 -> forall G1 G2, G = G1 ++ G2
|
adamc@152
|
141 -> G' # G2
|
adamc@152
|
142 -> G1 ++ G' ++ G2 |-e e : t.
|
adamc@152
|
143 induction 1; crush; eauto.
|
adamc@152
|
144 Qed.
|
adamc@152
|
145
|
adamc@152
|
146 Theorem weaken_hasType : forall G e t,
|
adamc@152
|
147 G |-e e : t
|
adamc@152
|
148 -> forall G', G' # G
|
adamc@152
|
149 -> G' ++ G |-e e : t.
|
adamc@152
|
150 intros; change (G' ++ G) with (nil ++ G' ++ G);
|
adamc@152
|
151 eapply weaken_hasType'; eauto.
|
adamc@152
|
152 Qed.
|
adamc@152
|
153
|
adamc@152
|
154 Theorem weaken_hasType_closed : forall e t,
|
adamc@152
|
155 nil |-e e : t
|
adamc@152
|
156 -> forall G, G |-e e : t.
|
adamc@152
|
157 intros; rewrite (app_nil_end G); apply weaken_hasType; auto.
|
adamc@152
|
158 Qed.
|
adamc@152
|
159
|
adamc@152
|
160 Theorem weaken_hasType1 : forall G e t,
|
adamc@152
|
161 G |-e e : t
|
adamc@152
|
162 -> forall x t', x ## G
|
adamc@152
|
163 -> (x, t') :: G |-e e : t.
|
adamc@152
|
164 intros; change ((x, t') :: G) with (((x, t') :: nil) ++ G);
|
adamc@152
|
165 apply weaken_hasType; crush;
|
adamc@152
|
166 match goal with
|
adamc@152
|
167 | [ H : (_, _) = (_, _) |- _ ] => injection H
|
adamc@152
|
168 end; crush; eauto.
|
adamc@152
|
169 Qed.
|
adamc@152
|
170
|
adamc@152
|
171 Hint Resolve weaken_hasType_closed weaken_hasType1.
|
adamc@152
|
172
|
adamc@152
|
173 Section subst.
|
adamc@152
|
174 Variable x : var.
|
adamc@152
|
175 Variable e1 : exp.
|
adamc@152
|
176
|
adamc@152
|
177 Fixpoint subst (e2 : exp) : exp :=
|
adamc@152
|
178 match e2 with
|
adamc@152
|
179 | Const b => Const b
|
adamc@152
|
180 | Var x' =>
|
adamc@152
|
181 if var_eq x' x
|
adamc@152
|
182 then e1
|
adamc@152
|
183 else Var x'
|
adamc@152
|
184 | App e1 e2 => App (subst e1) (subst e2)
|
adamc@152
|
185 | Abs x' e' =>
|
adamc@152
|
186 Abs x' (if var_eq x' x
|
adamc@152
|
187 then e'
|
adamc@152
|
188 else subst e')
|
adamc@152
|
189 end.
|
adamc@152
|
190
|
adamc@152
|
191 Variable xt : type.
|
adamc@152
|
192 Hypothesis Ht' : nil |-e e1 : xt.
|
adamc@152
|
193
|
adamc@152
|
194 Lemma subst_lookup' : forall G2 x' t,
|
adamc@152
|
195 x' ## G2
|
adamc@152
|
196 -> (x, xt) :: G2 |-v x' : t
|
adamc@152
|
197 -> t = xt.
|
adamc@152
|
198 inversion 2; crush; elimtype False; eauto.
|
adamc@152
|
199 Qed.
|
adamc@152
|
200
|
adamc@152
|
201 Lemma subst_lookup : forall x' t G2,
|
adamc@152
|
202 x <> x'
|
adamc@152
|
203 -> forall G1, G1 ++ (x, xt) :: G2 |-v x' : t
|
adamc@152
|
204 -> G1 ++ G2 |-v x' : t.
|
adamc@152
|
205 induction G1 as [ | [x'' t'] tl ]; crush;
|
adamc@152
|
206 match goal with
|
adamc@152
|
207 | [ H : _ |-v _ : _ |- _ ] => inversion H
|
adamc@152
|
208 end; crush.
|
adamc@152
|
209 Qed.
|
adamc@152
|
210
|
adamc@152
|
211 Hint Resolve subst_lookup.
|
adamc@152
|
212
|
adamc@152
|
213 Lemma subst_lookup'' : forall G2 x' t,
|
adamc@152
|
214 x' ## G2
|
adamc@152
|
215 -> forall G1, x' ## G1
|
adamc@152
|
216 -> G1 ++ (x, xt) :: G2 |-v x' : t
|
adamc@152
|
217 -> t = xt.
|
adamc@152
|
218 Hint Resolve subst_lookup'.
|
adamc@152
|
219
|
adamc@152
|
220 induction G1 as [ | [x'' t'] tl ]; crush; eauto;
|
adamc@152
|
221 match goal with
|
adamc@152
|
222 | [ H : _ |-v _ : _ |- _ ] => inversion H
|
adamc@152
|
223 end; crush; elimtype False; eauto.
|
adamc@152
|
224 Qed.
|
adamc@152
|
225
|
adamc@152
|
226 Implicit Arguments subst_lookup'' [G2 x' t G1].
|
adamc@152
|
227
|
adamc@152
|
228 Lemma disjoint_cons : forall x x' t (G : ctx),
|
adamc@152
|
229 x ## G
|
adamc@152
|
230 -> x' <> x
|
adamc@152
|
231 -> x ## (x', t) :: G.
|
adamc@152
|
232 firstorder;
|
adamc@152
|
233 match goal with
|
adamc@152
|
234 | [ H : (_, _) = (_, _) |- _ ] => injection H
|
adamc@152
|
235 end; crush.
|
adamc@152
|
236 Qed.
|
adamc@152
|
237
|
adamc@152
|
238 Hint Resolve disjoint_cons.
|
adamc@152
|
239
|
adamc@152
|
240 Lemma shadow_lookup : forall G2 v t t' G1,
|
adamc@152
|
241 G1 |-v x : t'
|
adamc@152
|
242 -> G1 ++ (x, xt) :: G2 |-v v : t
|
adamc@152
|
243 -> G1 ++ G2 |-v v : t.
|
adamc@152
|
244 induction G1 as [ | [x'' t''] tl ]; crush;
|
adamc@152
|
245 match goal with
|
adamc@152
|
246 | [ H : nil |-v _ : _ |- _ ] => inversion H
|
adamc@152
|
247 | [ H1 : _ |-v _ : _, H2 : _ |-v _ : _ |- _ ] =>
|
adamc@152
|
248 inversion H1; crush; inversion H2; crush
|
adamc@152
|
249 end.
|
adamc@152
|
250 Qed.
|
adamc@152
|
251
|
adamc@152
|
252 Lemma shadow_hasType' : forall G2 G e t,
|
adamc@152
|
253 G |-e e : t
|
adamc@152
|
254 -> forall G1, G = G1 ++ (x, xt) :: G2
|
adamc@152
|
255 -> forall t'', G1 |-v x : t''
|
adamc@152
|
256 -> G1 ++ G2 |-e e : t.
|
adamc@152
|
257 Hint Resolve shadow_lookup.
|
adamc@152
|
258
|
adamc@152
|
259 induction 1; crush; eauto;
|
adamc@152
|
260 match goal with
|
adamc@152
|
261 | [ H : (?x0, _) :: _ ++ (x, _) :: _ |-e _ : _ |- _ ] =>
|
adamc@152
|
262 destruct (var_eq x0 x); subst; eauto
|
adamc@152
|
263 end.
|
adamc@152
|
264 Qed.
|
adamc@152
|
265
|
adamc@152
|
266 Lemma shadow_hasType : forall G1 G2 e t t'',
|
adamc@152
|
267 G1 ++ (x, xt) :: G2 |-e e : t
|
adamc@152
|
268 -> G1 |-v x : t''
|
adamc@152
|
269 -> G1 ++ G2 |-e e : t.
|
adamc@152
|
270 intros; eapply shadow_hasType'; eauto.
|
adamc@152
|
271 Qed.
|
adamc@152
|
272
|
adamc@152
|
273 Hint Resolve shadow_hasType.
|
adamc@152
|
274
|
adamc@152
|
275 Theorem subst_hasType : forall G e2 t,
|
adamc@152
|
276 G |-e e2 : t
|
adamc@152
|
277 -> forall G1 G2, G = G1 ++ (x, xt) :: G2
|
adamc@152
|
278 -> x ## G1
|
adamc@152
|
279 -> x ## G2
|
adamc@152
|
280 -> G1 ++ G2 |-e subst e2 : t.
|
adamc@152
|
281 induction 1; crush;
|
adamc@152
|
282 try match goal with
|
adamc@152
|
283 | [ |- context[if ?E then _ else _] ] => destruct E
|
adamc@152
|
284 end; crush; eauto 6;
|
adamc@152
|
285 match goal with
|
adamc@152
|
286 | [ H1 : x ## _, H2 : x ## _, H3 : _ |-v x : _ |- _ ] =>
|
adamc@152
|
287 rewrite (subst_lookup'' H2 H1 H3)
|
adamc@152
|
288 end; crush.
|
adamc@152
|
289 Qed.
|
adamc@152
|
290
|
adamc@152
|
291 Theorem subst_hasType_closed : forall e2 t,
|
adamc@152
|
292 (x, xt) :: nil |-e e2 : t
|
adamc@152
|
293 -> nil |-e subst e2 : t.
|
adamc@152
|
294 intros; change (nil ++ nil |-e subst e2 : t);
|
adamc@152
|
295 eapply subst_hasType; eauto.
|
adamc@152
|
296 Qed.
|
adamc@152
|
297 End subst.
|
adamc@152
|
298
|
adamc@154
|
299 Hint Resolve subst_hasType_closed.
|
adamc@154
|
300
|
adamc@154
|
301 Notation "[ x ~> e1 ] e2" := (subst x e1 e2) (no associativity, at level 80).
|
adamc@154
|
302
|
adamc@154
|
303 Inductive val : exp -> Prop :=
|
adamc@154
|
304 | VConst : forall b, val (Const b)
|
adamc@154
|
305 | VAbs : forall x e, val (Abs x e).
|
adamc@154
|
306
|
adamc@154
|
307 Hint Constructors val.
|
adamc@154
|
308
|
adamc@154
|
309 Reserved Notation "e1 ==> e2" (no associativity, at level 90).
|
adamc@154
|
310
|
adamc@154
|
311 Inductive step : exp -> exp -> Prop :=
|
adamc@154
|
312 | Beta : forall x e1 e2,
|
adamc@154
|
313 App (Abs x e1) e2 ==> [x ~> e2] e1
|
adamc@154
|
314 | Cong1 : forall e1 e2 e1',
|
adamc@154
|
315 e1 ==> e1'
|
adamc@154
|
316 -> App e1 e2 ==> App e1' e2
|
adamc@154
|
317 | Cong2 : forall e1 e2 e2',
|
adamc@154
|
318 val e1
|
adamc@154
|
319 -> e2 ==> e2'
|
adamc@154
|
320 -> App e1 e2 ==> App e1 e2'
|
adamc@154
|
321
|
adamc@154
|
322 where "e1 ==> e2" := (step e1 e2).
|
adamc@154
|
323
|
adamc@154
|
324 Hint Constructors step.
|
adamc@154
|
325
|
adamc@154
|
326 Theorem progress : forall G e t, G |-e e : t
|
adamc@154
|
327 -> G = nil
|
adamc@154
|
328 -> val e \/ exists e', e ==> e'.
|
adamc@154
|
329 induction 1; crush; eauto;
|
adamc@154
|
330 try match goal with
|
adamc@154
|
331 | [ H : _ |-e _ : _ --> _ |- _ ] => inversion H
|
adamc@154
|
332 end;
|
adamc@154
|
333 repeat match goal with
|
adamc@154
|
334 | [ H : _ |- _ ] => solve [ inversion H; crush; eauto ]
|
adamc@154
|
335 end.
|
adamc@154
|
336 Qed.
|
adamc@154
|
337
|
adamc@154
|
338 Theorem preservation : forall G e t, G |-e e : t
|
adamc@154
|
339 -> G = nil
|
adamc@154
|
340 -> forall e', e ==> e'
|
adamc@154
|
341 -> G |-e e' : t.
|
adamc@154
|
342 induction 1; inversion 2; crush; eauto;
|
adamc@154
|
343 match goal with
|
adamc@154
|
344 | [ H : _ |-e Abs _ _ : _ |- _ ] => inversion H
|
adamc@154
|
345 end; eauto.
|
adamc@154
|
346 Qed.
|
adamc@154
|
347
|
adamc@152
|
348 End Concrete.
|