| author | wenzelm |
| Sat, 02 Sep 2000 21:49:51 +0200 | |
| changeset 9803 | bc883b390d91 |
| parent 9771 | 54c6a2c6e569 |
| child 9811 | 39ffdb8cab03 |
| permissions | -rw-r--r-- |
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
1 |
(* Title: HOL/Lambda/Type.thy |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
2 |
ID: $Id$ |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
3 |
Author: Stefan Berghofer |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
4 |
Copyright 2000 TU Muenchen |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
5 |
|
| 9622 | 6 |
Simply-typed lambda terms. Subject reduction and strong normalization |
7 |
of simply-typed lambda terms. Partly based on a paper proof by Ralph |
|
8 |
Matthes. |
|
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
9 |
*) |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
10 |
|
| 9622 | 11 |
theory Type = InductTermi: |
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
12 |
|
| 9641 | 13 |
datatype type = |
| 9622 | 14 |
Atom nat |
| 9641 | 15 |
| Fun type type (infixr "=>" 200) |
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
16 |
|
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
17 |
consts |
| 9641 | 18 |
typing :: "((nat => type) \<times> dB \<times> type) set" |
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
19 |
|
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
20 |
syntax |
| 9641 | 21 |
"_typing" :: "[nat => type, dB, type] => bool" ("_ |- _ : _" [50,50,50] 50)
|
22 |
"_funs" :: "[type list, type] => type" (infixl "=>>" 150) |
|
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
23 |
|
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
24 |
translations |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
25 |
"env |- t : T" == "(env, t, T) : typing" |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
26 |
"Ts =>> T" == "foldr Fun Ts T" |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
27 |
|
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
28 |
inductive typing |
| 9641 | 29 |
intros [intro!] |
| 9622 | 30 |
Var: "env x = T ==> env |- Var x : T" |
31 |
Abs: "(nat_case T env) |- t : U ==> env |- (Abs t) : (T => U)" |
|
32 |
App: "env |- s : T => U ==> env |- t : T ==> env |- (s $ t) : U" |
|
33 |
||
| 9641 | 34 |
inductive_cases [elim!]: |
35 |
"e |- Var i : T" |
|
36 |
"e |- t $ u : T" |
|
37 |
"e |- Abs t : T" |
|
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
38 |
|
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
39 |
consts |
| 9641 | 40 |
"types" :: "[nat => type, dB list, type list] => bool" |
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
41 |
primrec |
|
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
42 |
"types e [] Ts = (Ts = [])" |
| 9622 | 43 |
"types e (t # ts) Ts = |
44 |
(case Ts of |
|
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
45 |
[] => False |
| 9641 | 46 |
| T # Ts => e |- t : T \<and> types e ts Ts)" |
| 9622 | 47 |
|
48 |
inductive_cases [elim!]: |
|
49 |
"x # xs : lists S" |
|
50 |
||
| 9641 | 51 |
declare IT.intros [intro!] |
52 |
||
| 9622 | 53 |
|
54 |
text {* Some tests. *}
|
|
55 |
||
56 |
lemma "\<exists>T U. e |- Abs (Abs (Abs (Var 1 $ (Var 2 $ Var 1 $ Var 0)))) : T \<and> U = T" |
|
57 |
apply (intro exI conjI) |
|
| 9641 | 58 |
apply force |
| 9622 | 59 |
apply (rule refl) |
60 |
done |
|
61 |
||
62 |
lemma "\<exists>T U. e |- Abs (Abs (Abs (Var 2 $ Var 0 $ (Var 1 $ Var 0)))) : T \<and> U = T"; |
|
63 |
apply (intro exI conjI) |
|
| 9641 | 64 |
apply force |
| 9622 | 65 |
apply (rule refl) |
66 |
done |
|
67 |
||
68 |
||
69 |
text {* n-ary function types *}
|
|
70 |
||
71 |
lemma list_app_typeD [rulify]: |
|
72 |
"\<forall>t T. e |- t $$ ts : T --> (\<exists>Ts. e |- t : Ts =>> T \<and> types e ts Ts)" |
|
73 |
apply (induct_tac ts) |
|
74 |
apply simp |
|
75 |
apply (intro strip) |
|
76 |
apply simp |
|
77 |
apply (erule_tac x = "t $ a" in allE) |
|
78 |
apply (erule_tac x = T in allE) |
|
79 |
apply (erule impE) |
|
80 |
apply assumption |
|
81 |
apply (elim exE conjE) |
|
82 |
apply (ind_cases "e |- t $ u : T") |
|
83 |
apply (rule_tac x = "Ta # Ts" in exI) |
|
84 |
apply simp |
|
85 |
done |
|
86 |
||
87 |
lemma list_app_typeI [rulify]: |
|
88 |
"\<forall>t T Ts. e |- t : Ts =>> T --> types e ts Ts --> e |- t $$ ts : T" |
|
89 |
apply (induct_tac ts) |
|
90 |
apply (intro strip) |
|
91 |
apply simp |
|
92 |
apply (intro strip) |
|
93 |
apply (case_tac Ts) |
|
94 |
apply simp |
|
95 |
apply simp |
|
96 |
apply (erule_tac x = "t $ a" in allE) |
|
97 |
apply (erule_tac x = T in allE) |
|
98 |
apply (erule_tac x = lista in allE) |
|
99 |
apply (erule impE) |
|
100 |
apply (erule conjE) |
|
101 |
apply (erule typing.App) |
|
102 |
apply assumption |
|
103 |
apply blast |
|
104 |
done |
|
105 |
||
106 |
lemma lists_types [rulify]: |
|
107 |
"\<forall>Ts. types e ts Ts --> ts : lists {t. \<exists>T. e |- t : T}"
|
|
108 |
apply (induct_tac ts) |
|
109 |
apply (intro strip) |
|
110 |
apply (case_tac Ts) |
|
111 |
apply simp |
|
112 |
apply (rule lists.Nil) |
|
113 |
apply simp |
|
114 |
apply (intro strip) |
|
115 |
apply (case_tac Ts) |
|
116 |
apply simp |
|
117 |
apply simp |
|
118 |
apply (rule lists.Cons) |
|
119 |
apply blast |
|
120 |
apply blast |
|
121 |
done |
|
122 |
||
123 |
||
124 |
text {* lifting preserves termination and well-typedness *}
|
|
125 |
||
126 |
lemma lift_map [rulify, simp]: |
|
127 |
"\<forall>t. lift (t $$ ts) i = lift t i $$ map (\<lambda>t. lift t i) ts" |
|
128 |
apply (induct_tac ts) |
|
| 9641 | 129 |
apply simp_all |
| 9622 | 130 |
done |
131 |
||
132 |
lemma subst_map [rulify, simp]: |
|
133 |
"\<forall>t. subst (t $$ ts) u i = subst t u i $$ map (\<lambda>t. subst t u i) ts" |
|
134 |
apply (induct_tac ts) |
|
| 9641 | 135 |
apply simp_all |
| 9622 | 136 |
done |
137 |
||
138 |
lemma lift_IT [rulify, intro!]: |
|
139 |
"t : IT ==> \<forall>i. lift t i : IT" |
|
140 |
apply (erule IT.induct) |
|
141 |
apply (rule allI) |
|
142 |
apply (simp (no_asm)) |
|
143 |
apply (rule conjI) |
|
144 |
apply |
|
145 |
(rule impI, |
|
| 9716 | 146 |
rule IT.Var, |
| 9622 | 147 |
erule lists.induct, |
148 |
simp (no_asm), |
|
149 |
rule lists.Nil, |
|
150 |
simp (no_asm), |
|
151 |
erule IntE, |
|
152 |
rule lists.Cons, |
|
153 |
blast, |
|
154 |
assumption)+ |
|
155 |
apply auto |
|
156 |
done |
|
157 |
||
158 |
lemma lifts_IT [rulify]: |
|
159 |
"ts : lists IT --> map (\<lambda>t. lift t 0) ts : lists IT" |
|
160 |
apply (induct_tac ts) |
|
161 |
apply auto |
|
162 |
done |
|
163 |
||
164 |
||
165 |
lemma shift_env [simp]: |
|
166 |
"nat_case T |
|
167 |
(\<lambda>j. if j < i then e j else if j = i then Ua else e (j - 1)) = |
|
168 |
(\<lambda>j. if j < Suc i then nat_case T e j else if j = Suc i then Ua |
|
169 |
else nat_case T e (j - 1))" |
|
170 |
apply (rule ext) |
|
171 |
apply (case_tac j) |
|
172 |
apply simp |
|
173 |
apply (case_tac nat) |
|
| 9641 | 174 |
apply simp_all |
| 9622 | 175 |
done |
176 |
||
177 |
lemma lift_type' [rulify]: |
|
178 |
"e |- t : T ==> \<forall>i U. |
|
179 |
(\<lambda>j. if j < i then e j |
|
| 9641 | 180 |
else if j = i then U |
| 9622 | 181 |
else e (j - 1)) |- lift t i : T" |
182 |
apply (erule typing.induct) |
|
183 |
apply auto |
|
184 |
done |
|
185 |
||
186 |
lemma lift_type [intro!]: |
|
187 |
"e |- t : T ==> nat_case U e |- lift t 0 : T" |
|
188 |
apply (subgoal_tac |
|
189 |
"nat_case U e = |
|
190 |
(\<lambda>j. if j < 0 then e j |
|
191 |
else if j = 0 then U else e (j - 1))") |
|
192 |
apply (erule ssubst) |
|
193 |
apply (erule lift_type') |
|
194 |
apply (rule ext) |
|
195 |
apply (case_tac j) |
|
196 |
apply simp_all |
|
197 |
done |
|
198 |
||
199 |
lemma lift_types [rulify]: |
|
200 |
"\<forall>Ts. types e ts Ts --> |
|
201 |
types (\<lambda>j. if j < i then e j |
|
202 |
else if j = i then U |
|
203 |
else e (j - 1)) (map (\<lambda>t. lift t i) ts) Ts" |
|
204 |
apply (induct_tac ts) |
|
205 |
apply simp |
|
206 |
apply (intro strip) |
|
207 |
apply (case_tac Ts) |
|
208 |
apply simp_all |
|
209 |
apply (rule lift_type') |
|
210 |
apply (erule conjunct1) |
|
211 |
done |
|
212 |
||
213 |
||
214 |
text {* substitution lemma *}
|
|
215 |
||
216 |
lemma subst_lemma [rulify]: |
|
217 |
"e |- t : T ==> \<forall>e' i U u. |
|
218 |
e = (\<lambda>j. if j < i then e' j |
|
219 |
else if j = i then U |
|
220 |
else e' (j-1)) --> |
|
221 |
e' |- u : U --> e' |- t[u/i] : T" |
|
222 |
apply (erule typing.induct) |
|
223 |
apply (intro strip) |
|
224 |
apply (case_tac "x = i") |
|
225 |
apply simp |
|
226 |
apply (frule linorder_neq_iff [THEN iffD1]) |
|
227 |
apply (erule disjE) |
|
228 |
apply simp |
|
229 |
apply (rule typing.Var) |
|
230 |
apply assumption |
|
231 |
apply (frule order_less_not_sym) |
|
232 |
apply (simp only: subst_gt split: split_if add: if_False) |
|
233 |
apply (rule typing.Var) |
|
234 |
apply assumption |
|
235 |
apply fastsimp |
|
236 |
apply fastsimp |
|
237 |
done |
|
238 |
||
239 |
lemma substs_lemma [rulify]: |
|
240 |
"e |- u : T ==> |
|
241 |
\<forall>Ts. types (\<lambda>j. if j < i then e j |
|
242 |
else if j = i then T else e (j - 1)) ts Ts --> |
|
| 9641 | 243 |
types e (map (\<lambda>t. t[u/i]) ts) Ts" |
| 9622 | 244 |
apply (induct_tac ts) |
245 |
apply (intro strip) |
|
246 |
apply (case_tac Ts) |
|
247 |
apply simp |
|
248 |
apply simp |
|
249 |
apply (intro strip) |
|
250 |
apply (case_tac Ts) |
|
251 |
apply simp |
|
252 |
apply simp |
|
253 |
apply (erule conjE) |
|
254 |
apply (erule subst_lemma) |
|
| 9641 | 255 |
apply (rule refl) |
| 9622 | 256 |
apply assumption |
257 |
done |
|
258 |
||
259 |
||
260 |
text {* subject reduction *}
|
|
261 |
||
262 |
lemma subject_reduction [rulify]: |
|
263 |
"e |- t : T ==> \<forall>t'. t -> t' --> e |- t' : T" |
|
264 |
apply (erule typing.induct) |
|
265 |
apply blast |
|
266 |
apply blast |
|
267 |
apply (intro strip) |
|
268 |
apply (ind_cases "s $ t -> t'") |
|
269 |
apply hypsubst |
|
270 |
apply (ind_cases "env |- Abs t : T => U") |
|
271 |
apply (rule subst_lemma) |
|
272 |
apply assumption |
|
273 |
prefer 2 |
|
274 |
apply assumption |
|
275 |
apply (rule ext) |
|
276 |
apply (case_tac j) |
|
| 9641 | 277 |
apply auto |
| 9622 | 278 |
done |
279 |
||
280 |
text {* additional lemmas *}
|
|
281 |
||
282 |
lemma app_last: "(t $$ ts) $ u = t $$ (ts @ [u])" |
|
283 |
apply simp |
|
284 |
done |
|
285 |
||
286 |
lemma subst_Var_IT [rulify]: "r : IT ==> \<forall>i j. r[Var i/j] : IT" |
|
287 |
apply (erule IT.induct) |
|
| 9641 | 288 |
txt {* @{term Var} *}
|
| 9622 | 289 |
apply (intro strip) |
290 |
apply (simp (no_asm) add: subst_Var) |
|
291 |
apply |
|
292 |
((rule conjI impI)+, |
|
| 9716 | 293 |
rule IT.Var, |
| 9622 | 294 |
erule lists.induct, |
295 |
simp (no_asm), |
|
296 |
rule lists.Nil, |
|
297 |
simp (no_asm), |
|
298 |
erule IntE, |
|
299 |
erule CollectE, |
|
300 |
rule lists.Cons, |
|
301 |
fast, |
|
302 |
assumption)+ |
|
| 9641 | 303 |
txt {* @{term Lambda} *}
|
| 9622 | 304 |
apply (intro strip) |
305 |
apply simp |
|
| 9716 | 306 |
apply (rule IT.Lambda) |
| 9622 | 307 |
apply fast |
| 9641 | 308 |
txt {* @{term Beta} *}
|
| 9622 | 309 |
apply (intro strip) |
310 |
apply (simp (no_asm_use) add: subst_subst [symmetric]) |
|
| 9716 | 311 |
apply (rule IT.Beta) |
| 9622 | 312 |
apply auto |
313 |
done |
|
314 |
||
315 |
lemma Var_IT: "Var n \<in> IT" |
|
316 |
apply (subgoal_tac "Var n $$ [] \<in> IT") |
|
317 |
apply simp |
|
| 9716 | 318 |
apply (rule IT.Var) |
| 9622 | 319 |
apply (rule lists.Nil) |
320 |
done |
|
321 |
||
322 |
lemma app_Var_IT: "t : IT ==> t $ Var i : IT" |
|
323 |
apply (erule IT.induct) |
|
324 |
apply (subst app_last) |
|
| 9716 | 325 |
apply (rule IT.Var) |
| 9622 | 326 |
apply simp |
327 |
apply (rule lists.Cons) |
|
328 |
apply (rule Var_IT) |
|
329 |
apply (rule lists.Nil) |
|
| 9716 | 330 |
apply (rule IT.Beta [where ?ss = "[]", unfold foldl_Nil [THEN eq_reflection]]) |
| 9622 | 331 |
apply (erule subst_Var_IT) |
332 |
apply (rule Var_IT) |
|
333 |
apply (subst app_last) |
|
| 9716 | 334 |
apply (rule IT.Beta) |
| 9622 | 335 |
apply (subst app_last [symmetric]) |
336 |
apply assumption |
|
337 |
apply assumption |
|
338 |
done |
|
339 |
||
340 |
||
341 |
text {* Well-typed substitution preserves termination. *}
|
|
342 |
||
343 |
lemma subst_type_IT [rulify]: |
|
344 |
"\<forall>t. t : IT --> (\<forall>e T u i. |
|
345 |
(\<lambda>j. if j < i then e j |
|
346 |
else if j = i then U |
|
347 |
else e (j - 1)) |- t : T --> |
|
348 |
u : IT --> e |- u : U --> t[u/i] : IT)" |
|
349 |
apply (rule_tac f = size and a = U in measure_induct) |
|
350 |
apply (rule allI) |
|
351 |
apply (rule impI) |
|
352 |
apply (erule IT.induct) |
|
| 9641 | 353 |
txt {* @{term Var} *}
|
| 9622 | 354 |
apply (intro strip) |
355 |
apply (case_tac "n = i") |
|
| 9661 | 356 |
txt {* @{term "n = i"} *}
|
| 9622 | 357 |
apply (case_tac rs) |
358 |
apply simp |
|
359 |
apply simp |
|
360 |
apply (drule list_app_typeD) |
|
361 |
apply (elim exE conjE) |
|
362 |
apply (ind_cases "e |- t $ u : T") |
|
363 |
apply (ind_cases "e |- Var i : T") |
|
| 9641 | 364 |
apply (drule_tac s = "(?T::type) => ?U" in sym) |
| 9622 | 365 |
apply simp |
366 |
apply (subgoal_tac "lift u 0 $ Var 0 : IT") |
|
367 |
prefer 2 |
|
368 |
apply (rule app_Var_IT) |
|
369 |
apply (erule lift_IT) |
|
370 |
apply (subgoal_tac "(lift u 0 $ Var 0)[a[u/i]/0] : IT") |
|
371 |
apply (simp (no_asm_use)) |
|
| 9641 | 372 |
apply (subgoal_tac "(Var 0 $$ map (\<lambda>t. lift t 0) |
373 |
(map (\<lambda>t. t[u/i]) list))[(u $ a[u/i])/0] : IT") |
|
| 9771 | 374 |
apply (simp (no_asm_use) del: map_compose |
375 |
add: map_compose [symmetric] o_def) |
|
| 9622 | 376 |
apply (erule_tac x = "Ts =>> T" in allE) |
377 |
apply (erule impE) |
|
378 |
apply simp |
|
379 |
apply (erule_tac x = "Var 0 $$ |
|
| 9641 | 380 |
map (\<lambda>t. lift t 0) (map (\<lambda>t. t[u/i]) list)" in allE) |
| 9622 | 381 |
apply (erule impE) |
| 9716 | 382 |
apply (rule IT.Var) |
| 9622 | 383 |
apply (rule lifts_IT) |
384 |
apply (drule lists_types) |
|
385 |
apply |
|
386 |
(ind_cases "x # xs : lists (Collect P)", |
|
| 9641 | 387 |
erule lists_IntI [THEN lists.induct], |
388 |
assumption) |
|
389 |
apply fastsimp |
|
| 9622 | 390 |
apply fastsimp |
391 |
apply (erule_tac x = e in allE) |
|
392 |
apply (erule_tac x = T in allE) |
|
393 |
apply (erule_tac x = "u $ a[u/i]" in allE) |
|
394 |
apply (erule_tac x = 0 in allE) |
|
395 |
apply (fastsimp intro!: list_app_typeI lift_types subst_lemma substs_lemma) |
|
396 |
apply (erule_tac x = Ta in allE) |
|
397 |
apply (erule impE) |
|
398 |
apply simp |
|
399 |
apply (erule_tac x = "lift u 0 $ Var 0" in allE) |
|
400 |
apply (erule impE) |
|
401 |
apply assumption |
|
402 |
apply (erule_tac x = e in allE) |
|
403 |
apply (erule_tac x = "Ts =>> T" in allE) |
|
404 |
apply (erule_tac x = "a[u/i]" in allE) |
|
405 |
apply (erule_tac x = 0 in allE) |
|
406 |
apply (erule impE) |
|
407 |
apply (rule typing.App) |
|
408 |
apply (erule lift_type') |
|
409 |
apply (rule typing.Var) |
|
410 |
apply simp |
|
411 |
apply (fast intro!: subst_lemma) |
|
| 9641 | 412 |
txt {* @{term "n ~= i"} *}
|
| 9622 | 413 |
apply (drule list_app_typeD) |
414 |
apply (erule exE) |
|
415 |
apply (erule conjE) |
|
416 |
apply (drule lists_types) |
|
| 9641 | 417 |
apply (subgoal_tac "map (\<lambda>x. x[u/i]) rs : lists IT") |
| 9622 | 418 |
apply (simp add: subst_Var) |
419 |
apply fast |
|
420 |
apply (erule lists_IntI [THEN lists.induct]) |
|
421 |
apply assumption |
|
422 |
apply fastsimp |
|
423 |
apply fastsimp |
|
| 9641 | 424 |
txt {* @{term Lambda} *}
|
| 9622 | 425 |
apply fastsimp |
| 9641 | 426 |
txt {* @{term Beta} *}
|
| 9622 | 427 |
apply (intro strip) |
428 |
apply (simp (no_asm)) |
|
| 9716 | 429 |
apply (rule IT.Beta) |
| 9622 | 430 |
apply (simp (no_asm) del: subst_map add: subst_subst subst_map [symmetric]) |
431 |
apply (drule subject_reduction) |
|
432 |
apply (rule apps_preserves_beta) |
|
433 |
apply (rule beta.beta) |
|
434 |
apply fast |
|
435 |
apply (drule list_app_typeD) |
|
436 |
apply fast |
|
437 |
done |
|
438 |
||
439 |
||
440 |
text {* main theorem: well-typed terms are strongly normalizing *}
|
|
441 |
||
442 |
lemma type_implies_IT: "e |- t : T ==> t : IT" |
|
443 |
apply (erule typing.induct) |
|
444 |
apply (rule Var_IT) |
|
| 9716 | 445 |
apply (erule IT.Lambda) |
| 9622 | 446 |
apply (subgoal_tac "(Var 0 $ lift t 0)[s/0] : IT") |
447 |
apply simp |
|
448 |
apply (rule subst_type_IT) |
|
| 9771 | 449 |
apply (rule lists.Nil |
450 |
[THEN 2 lists.Cons [THEN IT.Var], unfold foldl_Nil [THEN eq_reflection] |
|
451 |
foldl_Cons [THEN eq_reflection]]) |
|
| 9622 | 452 |
apply (erule lift_IT) |
453 |
apply (rule typing.App) |
|
454 |
apply (rule typing.Var) |
|
455 |
apply simp |
|
456 |
apply (erule lift_type') |
|
457 |
apply assumption |
|
458 |
apply assumption |
|
459 |
done |
|
460 |
||
461 |
theorem type_implies_termi: "e |- t : T ==> t : termi beta" |
|
462 |
apply (rule IT_implies_termi) |
|
463 |
apply (erule type_implies_IT) |
|
464 |
done |
|
465 |
||
|
9114
de99e37effda
Subject reduction and strong normalization of simply-typed lambda terms.
berghofe
parents:
diff
changeset
|
466 |
end |