author | Christian Urban <urbanc@in.tum.de> |
Wed, 04 Apr 2012 06:19:38 +0100 | |
changeset 3152 | da59c94bed7e |
parent 3132 | 87eca760dcba |
child 3192 | 14c7d7e29c44 |
permissions | -rw-r--r-- |
2701 | 1 |
|
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
theory Tutorial5 |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
imports Tutorial4 |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
begin |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
2701 | 6 |
section {* Type-Preservation and Progress Lemma*} |
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
|
2701 | 8 |
text {* |
9 |
The point of this tutorial is to prove the |
|
10 |
type-preservation and progress lemma. Since |
|
11 |
we now know that \<Down>, \<longrightarrow>cbv* and the machine |
|
12 |
correspond to each other, we only need to |
|
13 |
prove this property for one of them. We chose |
|
14 |
\<longrightarrow>cbv. |
|
15 |
||
16 |
First we need to establish two elimination |
|
17 |
properties and two auxiliary lemmas about contexts. |
|
18 |
*} |
|
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
20 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
lemma valid_elim: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
assumes a: "valid ((x, T) # \<Gamma>)" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
shows "atom x \<sharp> \<Gamma> \<and> valid \<Gamma>" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
using a by (cases) (auto) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
lemma valid_insert: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
assumes a: "valid (\<Delta> @ [(x, T)] @ \<Gamma>)" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
shows "valid (\<Delta> @ \<Gamma>)" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
using a |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
by (induct \<Delta>) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
31 |
(auto simp add: fresh_append fresh_Cons dest!: valid_elim) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
32 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
33 |
lemma fresh_list: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
shows "atom y \<sharp> xs = (\<forall>x \<in> set xs. atom y \<sharp> x)" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
by (induct xs) (simp_all add: fresh_Nil fresh_Cons) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
lemma context_unique: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
assumes a1: "valid \<Gamma>" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
and a2: "(x, T) \<in> set \<Gamma>" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
and a3: "(x, U) \<in> set \<Gamma>" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
shows "T = U" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
using a1 a2 a3 |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
by (induct) (auto simp add: fresh_list fresh_Pair fresh_at_base) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
|
2701 | 45 |
|
46 |
section {* EXERCISE 16 *} |
|
47 |
||
48 |
text {* |
|
49 |
Next we want to show the type substitution lemma. Unfortunately, |
|
50 |
we have to prove a slightly more general version of it, where |
|
51 |
the variable being substituted occurs somewhere inside the |
|
52 |
context. |
|
53 |
*} |
|
54 |
||
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
lemma type_substitution_aux: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
assumes a: "\<Delta> @ [(x, T')] @ \<Gamma> \<turnstile> e : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
and b: "\<Gamma> \<turnstile> e' : T'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
shows "\<Delta> @ \<Gamma> \<turnstile> e[x ::= e'] : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
59 |
using a b |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
proof (nominal_induct \<Gamma>'\<equiv>"\<Delta> @ [(x, T')] @ \<Gamma>" e T avoiding: x e' \<Delta> rule: typing.strong_induct) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
case (t_Var y T x e' \<Delta>) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
have a1: "valid (\<Delta> @ [(x, T')] @ \<Gamma>)" by fact |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
63 |
have a2: "(y,T) \<in> set (\<Delta> @ [(x, T')] @ \<Gamma>)" by fact |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
have a3: "\<Gamma> \<turnstile> e' : T'" by fact |
2701 | 65 |
|
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
from a1 have a4: "valid (\<Delta> @ \<Gamma>)" by (rule valid_insert) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
{ assume eq: "x = y" |
2701 | 68 |
|
69 |
have "\<Delta> @ \<Gamma> \<turnstile> Var y[x ::= e'] : T" sorry |
|
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
} |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
71 |
moreover |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
72 |
{ assume ineq: "x \<noteq> y" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
73 |
from a2 have "(y, T) \<in> set (\<Delta> @ \<Gamma>)" using ineq by simp |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
74 |
then have "\<Delta> @ \<Gamma> \<turnstile> Var y[x ::= e'] : T" using ineq a4 by auto |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
75 |
} |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
ultimately show "\<Delta> @ \<Gamma> \<turnstile> Var y[x::=e'] : T" by blast |
2701 | 77 |
next |
78 |
case (t_Lam y T1 t T2 x e' \<Delta>) |
|
79 |
have a1: "atom y \<sharp> e'" by fact |
|
80 |
have a2: "atom y \<sharp> \<Delta> @ [(x, T')] @ \<Gamma>" by fact |
|
81 |
have a3: "\<Gamma> \<turnstile> e' : T'" by fact |
|
82 |
have ih: "\<Gamma> \<turnstile> e' : T' \<Longrightarrow> ((y, T1) # \<Delta>) @ \<Gamma> \<turnstile> t [x ::= e'] : T2" |
|
83 |
using t_Lam(6)[of "(y, T1) # \<Delta>"] by auto |
|
84 |
||
85 |
||
86 |
show "\<Delta> @ \<Gamma> \<turnstile> (Lam [y]. t)[x ::= e'] : T1 \<rightarrow> T2" sorry |
|
87 |
next |
|
88 |
case (t_App t1 T1 T2 t2 x e' \<Delta>) |
|
89 |
have ih1: "\<Gamma> \<turnstile> e' : T' \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t1 [x ::= e'] : T1 \<rightarrow> T2" using t_App(2) by auto |
|
90 |
have ih2: "\<Gamma> \<turnstile> e' : T' \<Longrightarrow> \<Delta> @ \<Gamma> \<turnstile> t2 [x ::= e'] : T1" using t_App(4) by auto |
|
91 |
have a: "\<Gamma> \<turnstile> e' : T'" by fact |
|
92 |
||
93 |
show "\<Delta> @ \<Gamma> \<turnstile> App t1 t2 [x ::= e'] : T2" sorry |
|
94 |
qed |
|
95 |
||
96 |
text {* |
|
97 |
From this we can derive the usual version of the substitution |
|
98 |
lemma. |
|
99 |
*} |
|
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
101 |
corollary type_substitution: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
102 |
assumes a: "(x, T') # \<Gamma> \<turnstile> e : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
103 |
and b: "\<Gamma> \<turnstile> e' : T'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
104 |
shows "\<Gamma> \<turnstile> e[x ::= e'] : T" |
2701 | 105 |
using a b type_substitution_aux[of "[]"] |
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
106 |
by auto |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
107 |
|
2701 | 108 |
|
109 |
section {* Type Preservation *} |
|
110 |
||
111 |
text {* |
|
112 |
Finally we are in a position to establish the type preservation |
|
113 |
property. We just need the following two inversion rules for |
|
114 |
particualr typing instances. |
|
115 |
*} |
|
116 |
||
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
117 |
lemma t_App_elim: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
118 |
assumes a: "\<Gamma> \<turnstile> App t1 t2 : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
119 |
obtains T' where "\<Gamma> \<turnstile> t1 : T' \<rightarrow> T" "\<Gamma> \<turnstile> t2 : T'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
120 |
using a |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
121 |
by (cases) (auto simp add: lam.eq_iff lam.distinct) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
122 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
123 |
text {* we have not yet generated strong elimination rules *} |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
124 |
lemma t_Lam_elim: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
125 |
assumes ty: "\<Gamma> \<turnstile> Lam [x].t : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
126 |
and fc: "atom x \<sharp> \<Gamma>" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
127 |
obtains T1 T2 where "T = T1 \<rightarrow> T2" "(x, T1) # \<Gamma> \<turnstile> t : T2" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
128 |
using ty fc |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
129 |
apply(cases) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
130 |
apply(auto simp add: lam.eq_iff lam.distinct ty.eq_iff) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
131 |
apply(auto simp add: Abs1_eq_iff) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
132 |
apply(rotate_tac 3) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
133 |
apply(drule_tac p="(x \<leftrightarrow> xa)" in permute_boolI) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
134 |
apply(perm_simp) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
135 |
apply(auto simp add: flip_def swap_fresh_fresh ty_fresh) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
136 |
done |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
137 |
|
2701 | 138 |
|
139 |
section {* EXERCISE 17 *} |
|
140 |
||
141 |
text {* |
|
142 |
Fill in the gaps in the t_Lam case. You will need |
|
143 |
the type substitution lemma proved above. |
|
144 |
*} |
|
145 |
||
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
theorem cbv_type_preservation: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
assumes a: "t \<longrightarrow>cbv t'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
and b: "\<Gamma> \<turnstile> t : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
149 |
shows "\<Gamma> \<turnstile> t' : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
150 |
using a b |
2701 | 151 |
proof (nominal_induct avoiding: \<Gamma> T rule: cbv.strong_induct) |
152 |
case (cbv1 v x t \<Gamma> T) |
|
153 |
have fc: "atom x \<sharp> \<Gamma>" by fact |
|
154 |
have "\<Gamma> \<turnstile> App (Lam [x]. t) v : T" by fact |
|
155 |
then obtain T' where |
|
156 |
*: "\<Gamma> \<turnstile> Lam [x]. t : T' \<rightarrow> T" and |
|
157 |
**: "\<Gamma> \<turnstile> v : T'" by (rule t_App_elim) |
|
158 |
have "(x, T') # \<Gamma> \<turnstile> t : T" using * fc by (rule t_Lam_elim) (simp add: ty.eq_iff) |
|
159 |
||
160 |
show "\<Gamma> \<turnstile> t [x ::= v] : T " sorry |
|
161 |
qed (auto elim!: t_App_elim) |
|
162 |
||
163 |
text {* |
|
164 |
We can easily extend this to sequences of cbv* reductions. |
|
165 |
*} |
|
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
166 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
167 |
corollary cbvs_type_preservation: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
168 |
assumes a: "t \<longrightarrow>cbv* t'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
169 |
and b: "\<Gamma> \<turnstile> t : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
170 |
shows "\<Gamma> \<turnstile> t' : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
171 |
using a b |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
172 |
by (induct) (auto intro: cbv_type_preservation) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
173 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
174 |
text {* |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
175 |
The type-preservation property for the machine and |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
176 |
evaluation relation. |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
177 |
*} |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
178 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
179 |
theorem machine_type_preservation: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
180 |
assumes a: "<t, []> \<mapsto>* <t', []>" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
181 |
and b: "\<Gamma> \<turnstile> t : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
182 |
shows "\<Gamma> \<turnstile> t' : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
183 |
proof - |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
184 |
have "t \<longrightarrow>cbv* t'" using a machines_implies_cbvs by simp |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
185 |
then show "\<Gamma> \<turnstile> t' : T" using b cbvs_type_preservation by simp |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
186 |
qed |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
187 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
188 |
theorem eval_type_preservation: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
189 |
assumes a: "t \<Down> t'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
190 |
and b: "\<Gamma> \<turnstile> t : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
191 |
shows "\<Gamma> \<turnstile> t' : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
192 |
proof - |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
193 |
have "<t, []> \<mapsto>* <t', []>" using a eval_implies_machines by simp |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
194 |
then show "\<Gamma> \<turnstile> t' : T" using b machine_type_preservation by simp |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
195 |
qed |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
196 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
197 |
text {* The Progress Property *} |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
198 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
199 |
lemma canonical_tArr: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
200 |
assumes a: "[] \<turnstile> t : T1 \<rightarrow> T2" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
201 |
and b: "val t" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
202 |
obtains x t' where "t = Lam [x].t'" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
203 |
using b a by (induct) (auto) |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
204 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
205 |
theorem progress: |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
206 |
assumes a: "[] \<turnstile> t : T" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
207 |
shows "(\<exists>t'. t \<longrightarrow>cbv t') \<or> (val t)" |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
208 |
using a |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
209 |
by (induct \<Gamma>\<equiv>"[]::ty_ctx" t T) |
2698 | 210 |
(auto elim: canonical_tArr simp add: val.simps) |
2691
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
211 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
212 |
text {* |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
213 |
Done! Congratulations! |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
214 |
*} |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
215 |
|
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
216 |
end |
abb6c3ac2df2
separated type preservation and progress into a separate file
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
217 |