22073
|
1 |
(* "$Id$" *)
|
|
2 |
(* *)
|
22082
|
3 |
(* Formalisation of the chapter on Logical Relations *)
|
|
4 |
(* and a Case Study in Equivalence Checking *)
|
|
5 |
(* by Karl Crary from the book on Advanced Topics in *)
|
|
6 |
(* Types and Programming Languages, MIT Press 2005 *)
|
|
7 |
|
|
8 |
(* The formalisation was done by Julien Narboux and *)
|
|
9 |
(* Christian Urban *)
|
22073
|
10 |
|
|
11 |
theory Crary
|
|
12 |
imports "Nominal"
|
|
13 |
begin
|
|
14 |
|
|
15 |
atom_decl name
|
|
16 |
|
|
17 |
nominal_datatype ty = TBase
|
|
18 |
| TUnit
|
|
19 |
| Arrow "ty" "ty" ("_\<rightarrow>_" [100,100] 100)
|
|
20 |
|
|
21 |
nominal_datatype trm = Unit
|
|
22 |
| Var "name"
|
|
23 |
| Lam "\<guillemotleft>name\<guillemotright>trm" ("Lam [_]._" [100,100] 100)
|
|
24 |
| App "trm" "trm"
|
|
25 |
| Const "nat"
|
|
26 |
|
22082
|
27 |
(* The next 3 lemmas should be in the nominal library *)
|
22073
|
28 |
lemma eq_eqvt:
|
|
29 |
fixes pi::"name prm"
|
|
30 |
and x::"'a::pt_name"
|
|
31 |
shows "pi\<bullet>(x=y) = (pi\<bullet>x=pi\<bullet>y)"
|
|
32 |
apply(simp add: perm_bool perm_bij)
|
|
33 |
done
|
|
34 |
|
|
35 |
lemma in_eqvt:
|
|
36 |
fixes pi::"name prm"
|
|
37 |
and x::"'a::pt_name"
|
|
38 |
assumes "x\<in>X"
|
|
39 |
shows "pi\<bullet>x \<in> pi\<bullet>X"
|
|
40 |
using assms by (perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
|
|
41 |
|
|
42 |
lemma set_eqvt:
|
|
43 |
fixes pi::"name prm"
|
|
44 |
and xs::"('a::pt_name) list"
|
|
45 |
shows "pi\<bullet>(set xs) = set (pi\<bullet>xs)"
|
|
46 |
by (perm_simp add: pt_list_set_pi[OF pt_name_inst])
|
22082
|
47 |
(* end of library *)
|
22073
|
48 |
|
|
49 |
lemma perm_ty[simp]:
|
|
50 |
fixes T::"ty"
|
|
51 |
and pi::"name prm"
|
|
52 |
shows "pi\<bullet>T = T"
|
|
53 |
by (induct T rule: ty.induct_weak) (simp_all)
|
|
54 |
|
|
55 |
lemma fresh_ty[simp]:
|
|
56 |
fixes x::"name"
|
|
57 |
and T::"ty"
|
|
58 |
shows "x\<sharp>T"
|
|
59 |
by (simp add: fresh_def supp_def)
|
|
60 |
|
|
61 |
lemma ty_cases:
|
|
62 |
fixes T::ty
|
|
63 |
shows "(\<exists> T1 T2. T=T1\<rightarrow>T2) \<or> T=TUnit \<or> T=TBase"
|
22082
|
64 |
by (induct T rule:ty.induct_weak) (auto)
|
22073
|
65 |
|
|
66 |
text {* Size Functions *}
|
|
67 |
|
|
68 |
instance ty :: size ..
|
|
69 |
|
|
70 |
nominal_primrec
|
|
71 |
"size (TBase) = 1"
|
|
72 |
"size (TUnit) = 1"
|
|
73 |
"size (T1\<rightarrow>T2) = size T1 + size T2"
|
|
74 |
by (rule TrueI)+
|
|
75 |
|
|
76 |
instance trm :: size ..
|
|
77 |
|
|
78 |
nominal_primrec
|
|
79 |
"size (Unit) = 1"
|
|
80 |
"size (Var x) = 1"
|
|
81 |
"size (Lam [x].t) = size t + 1"
|
|
82 |
"size (App t1 t2) = size t1 + size t2"
|
|
83 |
"size (Const n) = 1"
|
|
84 |
apply(finite_guess add: fs_name1 perm_nat_def)+
|
|
85 |
apply(perm_full_simp add: perm_nat_def)
|
|
86 |
apply(simp add: fs_name1)
|
|
87 |
apply(rule TrueI)+
|
|
88 |
apply(simp add: fresh_nat)+
|
|
89 |
apply(fresh_guess add: fs_name1 perm_nat_def)+
|
|
90 |
done
|
|
91 |
|
|
92 |
lemma ty_size_greater_zero[simp]:
|
|
93 |
fixes T::"ty"
|
|
94 |
shows "size T > 0"
|
|
95 |
by (nominal_induct rule:ty.induct) (simp_all)
|
|
96 |
|
|
97 |
text {* valid contexts *}
|
|
98 |
|
|
99 |
inductive2
|
|
100 |
valid :: "(name \<times> 'a::pt_name) list \<Rightarrow> bool"
|
|
101 |
where
|
22082
|
102 |
v_nil[intro]: "valid []"
|
22073
|
103 |
| v_cons[intro]: "\<lbrakk>valid \<Gamma>;a\<sharp>\<Gamma>\<rbrakk> \<Longrightarrow> valid ((a,\<sigma>)#\<Gamma>)"
|
|
104 |
|
|
105 |
lemma valid_eqvt:
|
|
106 |
fixes pi:: "name prm"
|
|
107 |
assumes a: "valid \<Gamma>"
|
|
108 |
shows "valid (pi\<bullet>\<Gamma>)"
|
|
109 |
using a
|
|
110 |
by (induct) (auto simp add: fresh_bij)
|
|
111 |
|
|
112 |
inductive_cases2
|
|
113 |
valid_cons_elim_auto[elim]:"valid ((x,T)#\<Gamma>)"
|
|
114 |
|
22082
|
115 |
text {* typing judgements for terms *}
|
22073
|
116 |
|
|
117 |
inductive2
|
|
118 |
typing :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" (" _ \<turnstile> _ : _ " [60,60,60] 60)
|
|
119 |
where
|
|
120 |
t_Var[intro]: "\<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> Var x : T"
|
|
121 |
| t_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> e1 : T1\<rightarrow>T2; \<Gamma> \<turnstile> e2 : T1\<rbrakk>\<Longrightarrow> \<Gamma> \<turnstile> App e1 e2 : T2"
|
|
122 |
| t_Lam[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T1)#\<Gamma> \<turnstile> t : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x].t : T1\<rightarrow>T2"
|
|
123 |
| t_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n : TBase"
|
|
124 |
| t_Unit[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Unit : TUnit"
|
|
125 |
|
|
126 |
lemma typing_valid:
|
|
127 |
assumes "\<Gamma> \<turnstile> t : T"
|
|
128 |
shows "valid \<Gamma>"
|
|
129 |
using assms
|
|
130 |
by (induct) (auto)
|
|
131 |
|
|
132 |
lemma typing_eqvt :
|
|
133 |
fixes pi::"name prm"
|
|
134 |
assumes "\<Gamma> \<turnstile> t : T"
|
|
135 |
shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>t) : T"
|
|
136 |
using assms
|
|
137 |
apply(induct)
|
|
138 |
apply(auto simp add: fresh_bij set_eqvt valid_eqvt)
|
|
139 |
apply(rule t_Var)
|
|
140 |
apply(drule valid_eqvt)
|
|
141 |
apply(assumption)
|
|
142 |
apply(drule in_eqvt)
|
|
143 |
apply(simp add: set_eqvt)
|
|
144 |
done
|
|
145 |
|
|
146 |
text {* Inversion lemmas for the typing judgment *}
|
|
147 |
|
|
148 |
declare trm.inject [simp add]
|
|
149 |
declare ty.inject [simp add]
|
|
150 |
|
|
151 |
inductive_cases2 t_Lam_elim_auto[elim]: "\<Gamma> \<turnstile> Lam [x].t : T"
|
|
152 |
inductive_cases2 t_Var_elim_auto[elim]: "\<Gamma> \<turnstile> Var x : T"
|
|
153 |
inductive_cases2 t_App_elim_auto[elim]: "\<Gamma> \<turnstile> App x y : T"
|
|
154 |
inductive_cases2 t_Const_elim_auto[elim]: "\<Gamma> \<turnstile> Const n : T"
|
|
155 |
inductive_cases2 t_Unit_elim_auto[elim]: "\<Gamma> \<turnstile> Unit : TUnit"
|
|
156 |
|
|
157 |
declare trm.inject [simp del]
|
|
158 |
declare ty.inject [simp del]
|
|
159 |
|
|
160 |
lemma typing_induct_strong
|
|
161 |
[consumes 1, case_names t_Var t_App t_Lam t_Const t_Unit]:
|
|
162 |
fixes P::"'a::fs_name \<Rightarrow> (name\<times>ty) list \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow>bool"
|
|
163 |
and x :: "'a::fs_name"
|
|
164 |
assumes a: "\<Gamma> \<turnstile> e : T"
|
|
165 |
and a1: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x,T)\<in>set \<Gamma>\<rbrakk> \<Longrightarrow> P c \<Gamma> (Var x) T"
|
|
166 |
and a2: "\<And>\<Gamma> e1 T1 T2 e2 c.
|
|
167 |
\<lbrakk>\<Gamma> \<turnstile> e1 : T1\<rightarrow>T2 ; \<And>c. P c \<Gamma> e1 (T1\<rightarrow>T2); \<Gamma> \<turnstile> e2 : T1 ; \<And>c. P c \<Gamma> e2 T1\<rbrakk>
|
|
168 |
\<Longrightarrow> P c \<Gamma> (App e1 e2) T2"
|
|
169 |
and a3: "\<And>x \<Gamma> T1 t T2 c.
|
|
170 |
\<lbrakk>x\<sharp>(\<Gamma>,c); (x,T1)#\<Gamma> \<turnstile> t : T2 ; \<And>c. P c ((x,T1)#\<Gamma>) t T2\<rbrakk>
|
|
171 |
\<Longrightarrow> P c \<Gamma> (Lam [x].t) (T1\<rightarrow>T2)"
|
|
172 |
and a4: "\<And>\<Gamma> n c. valid \<Gamma> \<Longrightarrow> P c \<Gamma> (Const n) TBase"
|
|
173 |
and a5: "\<And>\<Gamma> c. valid \<Gamma> \<Longrightarrow> P c \<Gamma> Unit TUnit"
|
|
174 |
shows "P c \<Gamma> e T"
|
|
175 |
proof -
|
|
176 |
from a have "\<And>(pi::name prm) c. P c (pi\<bullet>\<Gamma>) (pi\<bullet>e) T"
|
|
177 |
proof (induct)
|
|
178 |
case (t_Var \<Gamma> x T pi c)
|
|
179 |
have "valid \<Gamma>" by fact
|
|
180 |
then have "valid (pi\<bullet>\<Gamma>)" by (simp only: valid_eqvt)
|
|
181 |
moreover
|
|
182 |
have "(x,T)\<in>set \<Gamma>" by fact
|
|
183 |
then have "pi\<bullet>(x,T)\<in>pi\<bullet>(set \<Gamma>)" by (simp only: pt_set_bij[OF pt_name_inst, OF at_name_inst])
|
|
184 |
then have "(pi\<bullet>x,T)\<in>set (pi\<bullet>\<Gamma>)" by (simp add: set_eqvt)
|
|
185 |
ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Var x)) T" using a1 by simp
|
|
186 |
next
|
|
187 |
case (t_App \<Gamma> e1 T1 T2 e2 pi c)
|
|
188 |
thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(App e1 e2)) T2" using a2 by (simp, blast intro: typing_eqvt)
|
|
189 |
next
|
|
190 |
case (t_Lam x \<Gamma> T1 t T2 pi c)
|
|
191 |
obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>\<Gamma>,pi\<bullet>t,c)" by (erule exists_fresh[OF fs_name1])
|
|
192 |
let ?sw = "[(pi\<bullet>x,y)]"
|
|
193 |
let ?pi' = "?sw@pi"
|
|
194 |
have f0: "x\<sharp>\<Gamma>" by fact
|
|
195 |
have f1: "(pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>)" using f0 by (simp add: fresh_bij)
|
|
196 |
have f2: "y\<sharp>?pi'\<bullet>\<Gamma>" by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
|
|
197 |
have pr1: "(x,T1)#\<Gamma> \<turnstile> t : T2" by fact
|
|
198 |
then have "(?pi'\<bullet>((x,T1)#\<Gamma>)) \<turnstile> (?pi'\<bullet>t) : T2" by (simp only: typing_eqvt)
|
|
199 |
moreover
|
|
200 |
have ih1: "\<And>c. P c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>t) T2" by fact
|
|
201 |
ultimately have "P c (?pi'\<bullet>\<Gamma>) (Lam [y].(?pi'\<bullet>t)) (T1\<rightarrow>T2)" using fs f2 a3
|
|
202 |
by (simp add: calc_atm)
|
|
203 |
then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) (T1\<rightarrow>T2)"
|
|
204 |
by (simp del: append_Cons add: calc_atm pt_name2)
|
|
205 |
moreover have "(?sw\<bullet>pi\<bullet>\<Gamma>) = (pi\<bullet>\<Gamma>)"
|
|
206 |
by (rule perm_fresh_fresh) (simp_all add: fs f1)
|
|
207 |
moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t))) = Lam [(pi\<bullet>x)].(pi\<bullet>t)"
|
|
208 |
by (rule perm_fresh_fresh) (simp_all add: fs f1 fresh_atm abs_fresh)
|
|
209 |
ultimately show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Lam [x].t)) (T1\<rightarrow>T2)"
|
|
210 |
by simp
|
|
211 |
next
|
|
212 |
case (t_Const \<Gamma> n pi c)
|
|
213 |
thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>(Const n)) (TBase)" using a4 by (simp, blast intro: valid_eqvt)
|
|
214 |
next
|
|
215 |
case (t_Unit \<Gamma> pi c)
|
|
216 |
thus "P c (pi\<bullet>\<Gamma>) (pi\<bullet>Unit) (TUnit)" using a5 by (simp, blast intro: valid_eqvt)
|
|
217 |
qed
|
|
218 |
then have "P c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>e) T" by blast
|
|
219 |
then show "P c \<Gamma> e T" by simp
|
|
220 |
qed
|
|
221 |
|
|
222 |
text {* capture-avoiding substitution *}
|
|
223 |
|
|
224 |
fun
|
|
225 |
lookup :: "(name\<times>trm) list \<Rightarrow> name \<Rightarrow> trm"
|
|
226 |
where
|
|
227 |
"lookup [] X = Var X"
|
|
228 |
"lookup ((Y,T)#\<theta>) X = (if X=Y then T else lookup \<theta> X)"
|
|
229 |
|
|
230 |
lemma lookup_eqvt:
|
|
231 |
fixes pi::"name prm"
|
|
232 |
and \<theta>::"(name\<times>trm) list"
|
|
233 |
and X::"name"
|
|
234 |
shows "pi\<bullet>(lookup \<theta> X) = lookup (pi\<bullet>\<theta>) (pi\<bullet>X)"
|
|
235 |
by (induct \<theta>)
|
|
236 |
(auto simp add: perm_bij)
|
|
237 |
|
|
238 |
lemma lookup_fresh:
|
|
239 |
fixes z::"name"
|
|
240 |
assumes "z\<sharp>\<theta>" "z\<sharp>x"
|
|
241 |
shows "z \<sharp>lookup \<theta> x"
|
|
242 |
using assms
|
|
243 |
by (induct rule: lookup.induct) (auto simp add: fresh_list_cons)
|
|
244 |
|
|
245 |
lemma lookup_fresh':
|
|
246 |
assumes "z\<sharp>\<theta>"
|
|
247 |
shows "lookup \<theta> z = Var z"
|
|
248 |
using assms
|
|
249 |
by (induct rule: lookup.induct)
|
|
250 |
(auto simp add: fresh_list_cons fresh_prod fresh_atm)
|
|
251 |
|
|
252 |
consts
|
|
253 |
psubst :: "(name\<times>trm) list \<Rightarrow> trm \<Rightarrow> trm" ("_<_>" [60,100] 100)
|
|
254 |
|
|
255 |
nominal_primrec
|
|
256 |
"\<theta><(Var x)> = (lookup \<theta> x)"
|
|
257 |
"\<theta><(App t1 t2)> = App (\<theta><t1>) (\<theta><t2>)"
|
|
258 |
"x\<sharp>\<theta> \<Longrightarrow> \<theta><(Lam [x].t)> = Lam [x].(\<theta><t>)"
|
|
259 |
"\<theta><(Const n)> = Const n"
|
|
260 |
"\<theta><(Unit)> = Unit"
|
|
261 |
apply(finite_guess add: fs_name1 lookup_eqvt)+
|
|
262 |
apply(perm_full_simp)
|
|
263 |
apply(simp add: fs_name1)
|
|
264 |
apply(rule TrueI)+
|
|
265 |
apply(simp add: abs_fresh)+
|
|
266 |
apply(fresh_guess add: fs_name1 lookup_eqvt)+
|
|
267 |
done
|
|
268 |
|
|
269 |
abbreviation
|
|
270 |
subst :: "trm \<Rightarrow> name \<Rightarrow> trm \<Rightarrow> trm" ("_[_::=_]" [100,100,100] 100)
|
|
271 |
where
|
|
272 |
"t[x::=t'] \<equiv> ([(x,t')])<t>"
|
|
273 |
|
|
274 |
lemma subst[simp]:
|
|
275 |
shows "(Var x)[y::=t'] = (if x=y then t' else (Var x))"
|
|
276 |
and "(App t1 t2)[y::=t'] = App (t1[y::=t']) (t2[y::=t'])"
|
|
277 |
and "x\<sharp>(y,t') \<Longrightarrow> (Lam [x].t)[y::=t'] = Lam [x].(t[y::=t'])"
|
|
278 |
and "Const n[y::=t'] = Const n"
|
|
279 |
and "Unit [y::=t'] = Unit"
|
|
280 |
by (simp_all add: fresh_list_cons fresh_list_nil)
|
|
281 |
|
|
282 |
lemma subst_eqvt:
|
|
283 |
fixes pi::"name prm"
|
|
284 |
and t::"trm"
|
|
285 |
shows "pi\<bullet>(t[x::=t']) = (pi\<bullet>t)[(pi\<bullet>x)::=(pi\<bullet>t')]"
|
|
286 |
by (nominal_induct t avoiding: x t' rule: trm.induct)
|
|
287 |
(perm_simp add: fresh_bij)+
|
|
288 |
|
|
289 |
lemma subst_rename:
|
|
290 |
fixes c::"name"
|
|
291 |
and t1::"trm"
|
|
292 |
assumes a: "c\<sharp>t1"
|
|
293 |
shows "t1[a::=t2] = ([(c,a)]\<bullet>t1)[c::=t2]"
|
|
294 |
using a
|
|
295 |
apply(nominal_induct t1 avoiding: a c t2 rule: trm.induct)
|
|
296 |
apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def)+
|
|
297 |
done
|
|
298 |
|
|
299 |
lemma fresh_psubst:
|
|
300 |
fixes z::"name"
|
|
301 |
and t::"trm"
|
|
302 |
assumes "z\<sharp>t" "z\<sharp>\<theta>"
|
|
303 |
shows "z\<sharp>(\<theta><t>)"
|
|
304 |
using assms
|
|
305 |
by (nominal_induct t avoiding: z \<theta> t rule: trm.induct)
|
|
306 |
(auto simp add: abs_fresh lookup_fresh)
|
|
307 |
|
|
308 |
lemma fresh_subst'':
|
|
309 |
fixes z::"name"
|
|
310 |
and t1::"trm"
|
|
311 |
and t2::"trm"
|
|
312 |
assumes "z\<sharp>t2"
|
|
313 |
shows "z\<sharp>t1[z::=t2]"
|
|
314 |
using assms
|
|
315 |
by (nominal_induct t1 avoiding: t2 z rule: trm.induct)
|
|
316 |
(auto simp add: abs_fresh fresh_nat fresh_atm)
|
|
317 |
|
|
318 |
lemma fresh_subst':
|
|
319 |
fixes z::"name"
|
|
320 |
and t1::"trm"
|
|
321 |
and t2::"trm"
|
|
322 |
assumes "z\<sharp>[y].t1" "z\<sharp>t2"
|
|
323 |
shows "z\<sharp>t1[y::=t2]"
|
|
324 |
using assms
|
|
325 |
by (nominal_induct t1 avoiding: y t2 z rule: trm.induct)
|
|
326 |
(auto simp add: abs_fresh fresh_nat fresh_atm)
|
|
327 |
|
|
328 |
lemma fresh_subst:
|
|
329 |
fixes z::"name"
|
|
330 |
and t1::"trm"
|
|
331 |
and t2::"trm"
|
|
332 |
assumes "z\<sharp>t1" "z\<sharp>t2"
|
|
333 |
shows "z\<sharp>t1[y::=t2]"
|
|
334 |
using assms fresh_subst'
|
|
335 |
by (auto simp add: abs_fresh)
|
|
336 |
|
|
337 |
lemma fresh_psubst_simpl:
|
|
338 |
assumes "x\<sharp>t"
|
|
339 |
shows "(x,u)#\<theta><t> = \<theta><t>"
|
|
340 |
using assms
|
|
341 |
proof (nominal_induct t avoiding: x u \<theta> rule: trm.induct)
|
|
342 |
case (Lam y t x u)
|
|
343 |
have fs: "y\<sharp>\<theta>" "y\<sharp>x" "y\<sharp>u" by fact
|
|
344 |
moreover have "x\<sharp> Lam [y].t" by fact
|
|
345 |
ultimately have "x\<sharp>t" by (simp add: abs_fresh fresh_atm)
|
|
346 |
moreover have ih:"\<And>n T. n\<sharp>t \<Longrightarrow> ((n,T)#\<theta>)<t> = \<theta><t>" by fact
|
|
347 |
ultimately have "(x,u)#\<theta><t> = \<theta><t>" by auto
|
|
348 |
moreover have "(x,u)#\<theta><Lam [y].t> = Lam [y]. ((x,u)#\<theta><t>)" using fs
|
|
349 |
by (simp add: fresh_list_cons fresh_prod)
|
|
350 |
moreover have " \<theta><Lam [y].t> = Lam [y]. (\<theta><t>)" using fs by simp
|
|
351 |
ultimately show "(x,u)#\<theta><Lam [y].t> = \<theta><Lam [y].t>" by auto
|
|
352 |
qed (auto simp add: fresh_atm abs_fresh)
|
|
353 |
|
|
354 |
text {* Equivalence (defined) *}
|
|
355 |
|
|
356 |
inductive2
|
|
357 |
equiv_def :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ == _ : _" [60,60] 60)
|
|
358 |
where
|
|
359 |
Q_Refl[intro]: "\<Gamma> \<turnstile> t : T \<Longrightarrow> \<Gamma> \<turnstile> t == t : T"
|
|
360 |
| Q_Symm[intro]: "\<Gamma> \<turnstile> t == s : T \<Longrightarrow> \<Gamma> \<turnstile> s == t : T"
|
|
361 |
| Q_Trans[intro]: "\<lbrakk>\<Gamma> \<turnstile> s == t : T; \<Gamma> \<turnstile> t == u : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s == u : T"
|
|
362 |
| Q_Abs[intro]: "\<lbrakk>x\<sharp>\<Gamma>; (x,T1)#\<Gamma> \<turnstile> s2 == t2 : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Lam [x]. s2 == Lam [x]. t2 : T1 \<rightarrow> T2"
|
|
363 |
| Q_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> s1 == t1 : T1 \<rightarrow> T2 ; \<Gamma> \<turnstile> s2 == t2 : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App s1 s2 == App t1 t2 : T2"
|
|
364 |
| Q_Beta[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s2,t2); (x,T1)#\<Gamma> \<turnstile> s12 == t12 : T2 ; \<Gamma> \<turnstile> s2 == t2 : T1\<rbrakk>
|
|
365 |
\<Longrightarrow> \<Gamma> \<turnstile> App (Lam [x]. s12) s2 == t12[x::=t2] : T2"
|
|
366 |
| Q_Ext[intro]: "\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T1)#\<Gamma> \<turnstile> App s (Var x) == App t (Var x) : T2\<rbrakk>
|
|
367 |
\<Longrightarrow> \<Gamma> \<turnstile> s == t : T1 \<rightarrow> T2"
|
|
368 |
|
|
369 |
lemma equiv_def_valid:
|
|
370 |
assumes "\<Gamma> \<turnstile> t == s : T"
|
|
371 |
shows "valid \<Gamma>"
|
|
372 |
using assms by (induct,auto elim:typing_valid)
|
|
373 |
|
|
374 |
lemma equiv_def_eqvt:
|
|
375 |
fixes pi::"name prm"
|
|
376 |
assumes a: "\<Gamma> \<turnstile> s == t : T"
|
|
377 |
shows "(pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) == (pi\<bullet>t) : T"
|
|
378 |
using a
|
|
379 |
apply(induct)
|
|
380 |
apply(auto intro: typing_eqvt valid_eqvt simp add: fresh_bij subst_eqvt)
|
|
381 |
apply(rule_tac x="pi\<bullet>x" in Q_Ext)
|
|
382 |
apply(simp add: fresh_bij)+
|
|
383 |
done
|
|
384 |
|
|
385 |
lemma equiv_def_strong_induct
|
|
386 |
[consumes 1, case_names Q_Refl Q_Symm Q_Trans Q_Abs Q_App Q_Beta Q_Ext]:
|
|
387 |
fixes c::"'a::fs_name"
|
|
388 |
assumes a0: "\<Gamma> \<turnstile> s == t : T"
|
|
389 |
and a1: "\<And>\<Gamma> t T c. \<Gamma> \<turnstile> t : T \<Longrightarrow> P c \<Gamma> t t T"
|
|
390 |
and a2: "\<And>\<Gamma> t s T c. \<lbrakk>\<Gamma> \<turnstile> t == s : T; \<And>d. P d \<Gamma> t s T\<rbrakk> \<Longrightarrow> P c \<Gamma> s t T"
|
|
391 |
and a3: "\<And>\<Gamma> s t T u c.
|
|
392 |
\<lbrakk>\<Gamma> \<turnstile> s == t : T; \<And>d. P d \<Gamma> s t T; \<Gamma> \<turnstile> t == u : T; \<And>d. P d \<Gamma> t u T\<rbrakk>
|
|
393 |
\<Longrightarrow> P c \<Gamma> s u T"
|
|
394 |
and a4: "\<And>x \<Gamma> T1 s2 t2 T2 c. \<lbrakk>x\<sharp>\<Gamma>; x\<sharp>c; (x,T1)#\<Gamma> \<turnstile> s2 == t2 : T2; \<And>d. P d ((x,T1)#\<Gamma>) s2 t2 T2\<rbrakk>
|
|
395 |
\<Longrightarrow> P c \<Gamma> (Lam [x].s2) (Lam [x].t2) (T1\<rightarrow>T2)"
|
|
396 |
and a5: "\<And>\<Gamma> s1 t1 T1 T2 s2 t2 c.
|
|
397 |
\<lbrakk>\<Gamma> \<turnstile> s1 == t1 : T1\<rightarrow>T2; \<And>d. P d \<Gamma> s1 t1 (T1\<rightarrow>T2);
|
|
398 |
\<Gamma> \<turnstile> s2 == t2 : T1; \<And>d. P d \<Gamma> s2 t2 T1\<rbrakk> \<Longrightarrow> P c \<Gamma> (App s1 s2) (App t1 t2) T2"
|
|
399 |
and a6: "\<And>x \<Gamma> T1 s12 t12 T2 s2 t2 c.
|
|
400 |
\<lbrakk>x\<sharp>(\<Gamma>,s2,t2); x\<sharp>c; (x,T1)#\<Gamma> \<turnstile> s12 == t12 : T2; \<And>d. P d ((x,T1)#\<Gamma>) s12 t12 T2;
|
|
401 |
\<Gamma> \<turnstile> s2 == t2 : T1; \<And>d. P d \<Gamma> s2 t2 T1\<rbrakk>
|
|
402 |
\<Longrightarrow> P c \<Gamma> (App (Lam [x].s12) s2) (t12[x::=t2]) T2"
|
|
403 |
and a7: "\<And>x \<Gamma> s t T1 T2 c.
|
|
404 |
\<lbrakk>x\<sharp>(\<Gamma>,s,t); (x,T1)#\<Gamma> \<turnstile> App s (Var x) == App t (Var x) : T2;
|
|
405 |
\<And>d. P d ((x,T1)#\<Gamma>) (App s (Var x)) (App t (Var x)) T2\<rbrakk>
|
|
406 |
\<Longrightarrow> P c \<Gamma> s t (T1\<rightarrow>T2)"
|
|
407 |
shows "P c \<Gamma> s t T"
|
|
408 |
proof -
|
|
409 |
from a0 have "\<And>(pi::name prm) c. P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) T"
|
|
410 |
proof(induct)
|
|
411 |
case (Q_Refl \<Gamma> t T pi c)
|
|
412 |
then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>t) (pi\<bullet>t) T" using a1 by (simp add: typing_eqvt)
|
|
413 |
next
|
|
414 |
case (Q_Symm \<Gamma> t s T pi c)
|
|
415 |
then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) T" using a2 by (simp only: equiv_def_eqvt)
|
|
416 |
next
|
|
417 |
case (Q_Trans \<Gamma> s t T u pi c)
|
|
418 |
then show " P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>u) T" using a3 by (blast intro: equiv_def_eqvt)
|
|
419 |
next
|
|
420 |
case (Q_App \<Gamma> s1 t1 T1 T2 s2 t2 pi c)
|
|
421 |
then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>App s1 s2) (pi\<bullet>App t1 t2) T2" using a5
|
|
422 |
by (simp, blast intro: equiv_def_eqvt)
|
|
423 |
next
|
|
424 |
case (Q_Ext x \<Gamma> s t T1 T2 pi c)
|
22082
|
425 |
then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (T1\<rightarrow>T2)"
|
|
426 |
apply(auto intro!: a7 simp add: fresh_bij)
|
|
427 |
apply(drule equiv_def_eqvt)
|
22073
|
428 |
apply(simp)
|
|
429 |
done
|
|
430 |
next
|
|
431 |
case (Q_Abs x \<Gamma> T1 s2 t2 T2 pi c)
|
|
432 |
obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>s2,pi\<bullet>t2,pi\<bullet>\<Gamma>,c)" by (rule exists_fresh[OF fs_name1])
|
|
433 |
let ?sw="[(pi\<bullet>x,y)]"
|
|
434 |
let ?pi'="?sw@pi"
|
|
435 |
have f1: "x\<sharp>\<Gamma>" by fact
|
|
436 |
have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>)" using f1 by (simp add: fresh_bij)
|
|
437 |
have f3: "y\<sharp>?pi'\<bullet>\<Gamma>" using f1 by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp)
|
|
438 |
have pr1: "(x,T1)#\<Gamma> \<turnstile> s2 == t2 : T2" by fact
|
|
439 |
then have "?pi'\<bullet>((x,T1)#\<Gamma>) \<turnstile> (?pi'\<bullet>s2) == (?pi'\<bullet>t2) : T2" by (rule equiv_def_eqvt)
|
|
440 |
then have "((y,T1)#(?pi'\<bullet>\<Gamma>)) \<turnstile> (?pi'\<bullet>s2) == (?pi'\<bullet>t2) : T2" by (simp add: calc_atm)
|
|
441 |
moreover
|
|
442 |
have ih1: "\<And>c. P c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>s2) (?pi'\<bullet>t2) T2" by fact
|
|
443 |
then have "\<And>c. P c ((y,T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>s2) (?pi'\<bullet>t2) T2" by (simp add: calc_atm)
|
|
444 |
ultimately have "P c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>Lam [x].s2) (?pi'\<bullet>Lam [x].t2) (T1\<rightarrow>T2)" using a4 f3 fs
|
|
445 |
by (simp add: calc_atm)
|
|
446 |
then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>Lam [(pi\<bullet>x)].(pi\<bullet>s2)) (?sw\<bullet>Lam [(pi\<bullet>x)].(pi\<bullet>t2)) (T1\<rightarrow>T2)"
|
|
447 |
by (simp del: append_Cons add: calc_atm pt_name2)
|
|
448 |
moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)"
|
|
449 |
by (rule perm_fresh_fresh) (simp_all add: fs f2)
|
|
450 |
moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>s2))) = Lam [(pi\<bullet>x)].(pi\<bullet>s2)"
|
|
451 |
by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
|
|
452 |
moreover have "(?sw\<bullet>(Lam [(pi\<bullet>x)].(pi\<bullet>t2))) = Lam [(pi\<bullet>x)].(pi\<bullet>t2)"
|
|
453 |
by (rule perm_fresh_fresh) (simp_all add: fs f2 abs_fresh)
|
|
454 |
ultimately have "P c (pi\<bullet>\<Gamma>) (Lam [(pi\<bullet>x)].(pi\<bullet>s2)) (Lam [(pi\<bullet>x)].(pi\<bullet>t2)) (T1\<rightarrow>T2)" by simp
|
|
455 |
then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>Lam [x].s2) (pi\<bullet>Lam [x].t2) (T1\<rightarrow>T2)" by simp
|
|
456 |
next
|
|
457 |
case (Q_Beta x \<Gamma> s2 t2 T1 s12 t12 T2 pi c)
|
|
458 |
obtain y::"name" where fs: "y\<sharp>(pi\<bullet>x,pi\<bullet>s12,pi\<bullet>t12,pi\<bullet>s2,pi\<bullet>t2,pi\<bullet>\<Gamma>,c)"
|
|
459 |
by (rule exists_fresh[OF fs_name1])
|
|
460 |
let ?sw="[(pi\<bullet>x,y)]"
|
|
461 |
let ?pi'="?sw@pi"
|
|
462 |
have f1: "x\<sharp>(\<Gamma>,s2,t2)" by fact
|
|
463 |
have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>(\<Gamma>,s2,t2))" using f1 by (simp add: fresh_bij)
|
|
464 |
have f3: "y\<sharp>(?pi'\<bullet>(\<Gamma>,s2,t2))" using f1
|
|
465 |
by (auto simp add: pt_name2 fresh_left calc_atm perm_pi_simp fresh_prod)
|
|
466 |
have pr1: "(x,T1)#\<Gamma> \<turnstile> s12 == t12 : T2" by fact
|
|
467 |
then have "?pi'\<bullet>((x,T1)#\<Gamma>) \<turnstile> (?pi'\<bullet>s12) == (?pi'\<bullet>t12) : T2" by (rule equiv_def_eqvt)
|
|
468 |
then have "((y,T1)#(?pi'\<bullet>\<Gamma>)) \<turnstile> (?pi'\<bullet>s12) == (?pi'\<bullet>t12) : T2" by (simp add: calc_atm)
|
|
469 |
moreover
|
|
470 |
have ih1: "\<And>c. P c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>s12) (?pi'\<bullet>t12) T2" by fact
|
|
471 |
then have "\<And>c. P c ((y,T1)#(?pi'\<bullet>\<Gamma>)) (?pi'\<bullet>s12) (?pi'\<bullet>t12) T2" by (simp add: calc_atm)
|
|
472 |
moreover
|
|
473 |
have pr2: "\<Gamma> \<turnstile> s2 == t2 : T1" by fact
|
|
474 |
then have "(?pi'\<bullet>\<Gamma>) \<turnstile> (?pi'\<bullet>s2) == (?pi'\<bullet>t2) : T1" by (rule equiv_def_eqvt)
|
|
475 |
moreover
|
|
476 |
have ih2: "\<And>c. P c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>s2) (?pi'\<bullet>t2) T1" by fact
|
|
477 |
ultimately have "P c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>(App (Lam [x].s12) s2)) (?pi'\<bullet>(t12[x::=t2])) T2"
|
|
478 |
using a6 f3 fs by (simp add: subst_eqvt calc_atm)
|
|
479 |
then have "P c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>(App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2)))
|
|
480 |
(?sw\<bullet>((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)])) T2"
|
|
481 |
by (simp del: append_Cons add: calc_atm pt_name2 subst_eqvt)
|
|
482 |
moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)"
|
|
483 |
by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
|
|
484 |
moreover have "(?sw\<bullet>(App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2))) = App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2)"
|
|
485 |
by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified] abs_fresh)
|
|
486 |
moreover have "(?sw\<bullet>((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)])) = ((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)])"
|
|
487 |
by (rule perm_fresh_fresh)
|
|
488 |
(simp_all add: fs[simplified] f2[simplified] fresh_subst fresh_subst'')
|
|
489 |
ultimately have "P c (pi\<bullet>\<Gamma>) (App (Lam [(pi\<bullet>x)].(pi\<bullet>s12)) (pi\<bullet>s2)) ((pi\<bullet>t12)[(pi\<bullet>x)::=(pi\<bullet>t2)]) T2"
|
|
490 |
by simp
|
|
491 |
then show "P c (pi\<bullet>\<Gamma>) (pi\<bullet>App (Lam [x].s12) s2) (pi\<bullet>t12[x::=t2]) T2" by (simp add: subst_eqvt)
|
|
492 |
qed
|
|
493 |
then have "P c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>s) (([]::name prm)\<bullet>t) T" by blast
|
|
494 |
then show "P c \<Gamma> s t T" by simp
|
|
495 |
qed
|
|
496 |
|
|
497 |
text {* Weak head reduction *}
|
|
498 |
|
|
499 |
inductive2
|
|
500 |
whr_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<leadsto> _" [80,80] 80)
|
|
501 |
where
|
|
502 |
QAR_Beta[intro]: "App (Lam [x]. t12) t2 \<leadsto> t12[x::=t2]"
|
|
503 |
| QAR_App[intro]: "t1 \<leadsto> t1' \<Longrightarrow> App t1 t2 \<leadsto> App t1' t2"
|
|
504 |
|
|
505 |
declare trm.inject [simp add]
|
|
506 |
declare ty.inject [simp add]
|
|
507 |
|
|
508 |
inductive_cases2 whr_Gen[elim]: "t \<leadsto> t'"
|
|
509 |
inductive_cases2 whr_Lam[elim]: "Lam [x].t \<leadsto> t'"
|
|
510 |
inductive_cases2 whr_App_Lam[elim]: "App (Lam [x].t12) t2 \<leadsto> t"
|
|
511 |
inductive_cases2 whr_Var[elim]: "Var x \<leadsto> t"
|
|
512 |
inductive_cases2 whr_Const[elim]: "Const n \<leadsto> t"
|
|
513 |
inductive_cases2 whr_App[elim]: "App p q \<leadsto> t"
|
|
514 |
inductive_cases2 whr_Const_Right[elim]: "t \<leadsto> Const n"
|
|
515 |
inductive_cases2 whr_Var_Right[elim]: "t \<leadsto> Var x"
|
|
516 |
inductive_cases2 whr_App_Right[elim]: "t \<leadsto> App p q"
|
|
517 |
|
|
518 |
declare trm.inject [simp del]
|
|
519 |
declare ty.inject [simp del]
|
|
520 |
|
|
521 |
text {* Weak head normalization *}
|
|
522 |
|
|
523 |
abbreviation
|
|
524 |
nf :: "trm \<Rightarrow> bool" ("_ \<leadsto>|" [100] 100)
|
|
525 |
where
|
|
526 |
"t\<leadsto>| \<equiv> \<not>(\<exists> u. t \<leadsto> u)"
|
|
527 |
|
|
528 |
inductive2
|
|
529 |
whn_def :: "trm\<Rightarrow>trm\<Rightarrow>bool" ("_ \<Down> _" [80,80] 80)
|
|
530 |
where
|
|
531 |
QAN_Reduce[intro]: "\<lbrakk>s \<leadsto> t; t \<Down> u\<rbrakk> \<Longrightarrow> s \<Down> u"
|
|
532 |
| QAN_Normal[intro]: "t\<leadsto>| \<Longrightarrow> t \<Down> t"
|
|
533 |
|
|
534 |
lemma whr_eqvt:
|
|
535 |
fixes pi::"name prm"
|
|
536 |
assumes a: "t \<leadsto> t'"
|
|
537 |
shows "(pi\<bullet>t) \<leadsto> (pi\<bullet>t')"
|
22082
|
538 |
using a
|
|
539 |
by (induct) (auto simp add: subst_eqvt fresh_bij)
|
22073
|
540 |
|
|
541 |
lemma whn_eqvt:
|
|
542 |
fixes pi::"name prm"
|
|
543 |
assumes a: "t \<Down> t'"
|
|
544 |
shows "(pi\<bullet>t) \<Down> (pi\<bullet>t')"
|
22082
|
545 |
using a
|
|
546 |
apply(induct)
|
|
547 |
apply(rule QAN_Reduce)
|
|
548 |
apply(rule whr_eqvt)
|
|
549 |
apply(assumption)+
|
|
550 |
apply(rule QAN_Normal)
|
|
551 |
apply(auto)
|
|
552 |
apply(drule_tac pi="rev pi" in whr_eqvt)
|
|
553 |
apply(perm_simp)
|
|
554 |
done
|
22073
|
555 |
|
|
556 |
text {* Algorithmic term equivalence and algorithmic path equivalence *}
|
|
557 |
|
|
558 |
inductive2
|
|
559 |
alg_equiv :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ <=> _ : _" [60,60] 60)
|
|
560 |
and
|
|
561 |
alg_path_equiv :: "(name\<times>ty) list\<Rightarrow>trm\<Rightarrow>trm\<Rightarrow>ty\<Rightarrow>bool" ("_ \<turnstile> _ \<leftrightarrow> _ : _" [60,60] 60)
|
|
562 |
where
|
|
563 |
QAT_Base[intro]: "\<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : TBase"
|
|
564 |
| QAT_Arrow[intro]: "\<lbrakk>x\<sharp>\<Gamma>; x\<sharp>s; x\<sharp>t; (x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2\<rbrakk>
|
|
565 |
\<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T1 \<rightarrow> T2"
|
|
566 |
| QAT_One[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : TUnit"
|
|
567 |
| QAP_Var[intro]: "\<lbrakk>valid \<Gamma>;(x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T"
|
|
568 |
| QAP_App[intro]: "\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T1 \<rightarrow> T2; \<Gamma> \<turnstile> s <=> t : T1 \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> App p s \<leftrightarrow> App q t : T2"
|
|
569 |
| QAP_Const[intro]: "valid \<Gamma> \<Longrightarrow> \<Gamma> \<turnstile> Const n \<leftrightarrow> Const n : TBase"
|
|
570 |
|
|
571 |
lemma alg_equiv_alg_path_equiv_eqvt:
|
|
572 |
fixes pi::"name prm"
|
|
573 |
shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) <=> (pi\<bullet>t) : T"
|
|
574 |
and "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>p) \<leftrightarrow> (pi\<bullet>q) : T"
|
|
575 |
apply(induct rule: alg_equiv_alg_path_equiv.inducts)
|
|
576 |
apply(auto intro: whn_eqvt simp add: fresh_bij valid_eqvt)
|
|
577 |
apply(rule_tac x="pi\<bullet>x" in QAT_Arrow)
|
|
578 |
apply(auto simp add: fresh_bij)
|
|
579 |
apply(rule QAP_Var)
|
|
580 |
apply(simp add: valid_eqvt)
|
|
581 |
apply(simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
|
|
582 |
apply(perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
|
|
583 |
done
|
|
584 |
|
|
585 |
lemma alg_equiv_alg_path_equiv_strong_induct
|
|
586 |
[case_names QAT_Base QAT_Arrow QAT_One QAP_Var QAP_App QAP_Const]:
|
|
587 |
fixes c::"'a::fs_name"
|
|
588 |
assumes a1: "\<And>s p t q \<Gamma> c. \<lbrakk>s \<Down> p; t \<Down> q; \<Gamma> \<turnstile> p \<leftrightarrow> q : TBase; \<And>d. P2 d \<Gamma> p q TBase\<rbrakk>
|
|
589 |
\<Longrightarrow> P1 c \<Gamma> s t TBase"
|
|
590 |
and a2: "\<And>x \<Gamma> s t T1 T2 c.
|
|
591 |
\<lbrakk>x\<sharp>\<Gamma>; x\<sharp>s; x\<sharp>t; x\<sharp>c; (x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2;
|
|
592 |
\<And>d. P1 d ((x,T1)#\<Gamma>) (App s (Var x)) (App t (Var x)) T2\<rbrakk>
|
|
593 |
\<Longrightarrow> P1 c \<Gamma> s t (T1\<rightarrow>T2)"
|
|
594 |
and a3: "\<And>\<Gamma> s t c. valid \<Gamma> \<Longrightarrow> P1 c \<Gamma> s t TUnit"
|
|
595 |
and a4: "\<And>\<Gamma> x T c. \<lbrakk>valid \<Gamma>; (x,T) \<in> set \<Gamma>\<rbrakk> \<Longrightarrow> P2 c \<Gamma> (Var x) (Var x) T"
|
|
596 |
and a5: "\<And>\<Gamma> p q T1 T2 s t c.
|
|
597 |
\<lbrakk>\<Gamma> \<turnstile> p \<leftrightarrow> q : T1\<rightarrow>T2; \<And>d. P2 d \<Gamma> p q (T1\<rightarrow>T2); \<Gamma> \<turnstile> s <=> t : T1; \<And>d. P1 d \<Gamma> s t T1\<rbrakk>
|
|
598 |
\<Longrightarrow> P2 c \<Gamma> (App p s) (App q t) T2"
|
|
599 |
and a6: "\<And>\<Gamma> n c. valid \<Gamma> \<Longrightarrow> P2 c \<Gamma> (Const n) (Const n) TBase"
|
|
600 |
shows "(\<Gamma> \<turnstile> s <=> t : T \<longrightarrow>P1 c \<Gamma> s t T) \<and> (\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> P2 c \<Gamma>' s' t' T')"
|
|
601 |
proof -
|
|
602 |
have "(\<Gamma> \<turnstile> s <=> t : T \<longrightarrow> (\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) T)) \<and>
|
|
603 |
(\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> (\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>') (pi\<bullet>s') (pi\<bullet>t') T'))"
|
|
604 |
proof(induct rule: alg_equiv_alg_path_equiv.induct)
|
|
605 |
case (QAT_Base s q t p \<Gamma>)
|
|
606 |
then show "\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) TBase"
|
|
607 |
apply(auto)
|
|
608 |
apply(rule_tac p="pi\<bullet>q" and q="pi\<bullet>p" in a1)
|
|
609 |
apply(simp_all add: whn_eqvt alg_equiv_alg_path_equiv_eqvt)
|
|
610 |
done
|
|
611 |
next
|
|
612 |
case (QAT_Arrow x \<Gamma> s t T1 T2)
|
|
613 |
show ?case
|
|
614 |
proof (rule allI, rule allI)
|
|
615 |
fix c::"'a::fs_name" and pi::"name prm"
|
|
616 |
obtain y::"name" where fs: "y\<sharp>(pi\<bullet>s,pi\<bullet>t,pi\<bullet>\<Gamma>,c)" by (rule exists_fresh[OF fs_name1])
|
|
617 |
let ?sw="[(pi\<bullet>x,y)]"
|
|
618 |
let ?pi'="?sw@pi"
|
|
619 |
have f0: "x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
|
|
620 |
then have f1: "x\<sharp>(\<Gamma>,s,t)" by simp
|
|
621 |
have f2: "(pi\<bullet>x)\<sharp>(pi\<bullet>(\<Gamma>,s,t))" using f1 by (simp add: fresh_bij)
|
|
622 |
have f3: "y\<sharp>?pi'\<bullet>(\<Gamma>,s,t)" using f1
|
|
623 |
by (simp only: pt_name2 fresh_left, auto simp add: perm_pi_simp calc_atm)
|
|
624 |
then have f3': "y\<sharp>?pi'\<bullet>\<Gamma>" "y\<sharp>?pi'\<bullet>s" "y\<sharp>?pi'\<bullet>t" by (auto simp add: fresh_prod)
|
|
625 |
have pr1: "(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2" by fact
|
|
626 |
then have "(?pi'\<bullet>((x,T1)#\<Gamma>)) \<turnstile> (?pi'\<bullet>(App s (Var x))) <=> (?pi'\<bullet>(App t (Var x))) : T2"
|
|
627 |
by (simp only: alg_equiv_alg_path_equiv_eqvt)
|
|
628 |
then have "(y,T1)#(?pi'\<bullet>\<Gamma>) \<turnstile> (App (?pi'\<bullet>s) (Var y)) <=> (App (?pi'\<bullet>t) (Var y)) : T2"
|
|
629 |
by (simp add: calc_atm)
|
|
630 |
moreover
|
|
631 |
have ih1: "\<forall>c (pi::name prm). P1 c (pi\<bullet>((x,T1)#\<Gamma>)) (pi\<bullet>App s (Var x)) (pi\<bullet>App t (Var x)) T2"
|
|
632 |
by fact
|
|
633 |
then have "\<And>c. P1 c (?pi'\<bullet>((x,T1)#\<Gamma>)) (?pi'\<bullet>App s (Var x)) (?pi'\<bullet>App t (Var x)) T2"
|
|
634 |
by auto
|
|
635 |
then have "\<And>c. P1 c ((y,T1)#(?pi'\<bullet>\<Gamma>)) (App (?pi'\<bullet>s) (Var y)) (App (?pi'\<bullet>t) (Var y)) T2"
|
|
636 |
by (simp add: calc_atm)
|
|
637 |
ultimately have "P1 c (?pi'\<bullet>\<Gamma>) (?pi'\<bullet>s) (?pi'\<bullet>t) (T1\<rightarrow>T2)" using a2 f3' fs by simp
|
|
638 |
then have "P1 c (?sw\<bullet>pi\<bullet>\<Gamma>) (?sw\<bullet>pi\<bullet>s) (?sw\<bullet>pi\<bullet>t) (T1\<rightarrow>T2)"
|
|
639 |
by (simp del: append_Cons add: calc_atm pt_name2)
|
|
640 |
moreover have "(?sw\<bullet>(pi\<bullet>\<Gamma>)) = (pi\<bullet>\<Gamma>)"
|
|
641 |
by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
|
|
642 |
moreover have "(?sw\<bullet>(pi\<bullet>s)) = (pi\<bullet>s)"
|
|
643 |
by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
|
|
644 |
moreover have "(?sw\<bullet>(pi\<bullet>t)) = (pi\<bullet>t)"
|
|
645 |
by (rule perm_fresh_fresh) (simp_all add: fs[simplified] f2[simplified])
|
|
646 |
ultimately show "P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) (T1\<rightarrow>T2)" by (simp)
|
|
647 |
qed
|
|
648 |
next
|
|
649 |
case (QAT_One \<Gamma> s t)
|
|
650 |
then show "\<forall>c (pi::name prm). P1 c (pi\<bullet>\<Gamma>) (pi\<bullet>s) (pi\<bullet>t) TUnit"
|
|
651 |
by (auto intro!: a3 simp add: valid_eqvt)
|
|
652 |
next
|
|
653 |
case (QAP_Var \<Gamma> x T)
|
|
654 |
then show "\<forall>c (pi::name prm). P2 c (pi \<bullet> \<Gamma>) (pi \<bullet> Var x) (pi \<bullet> Var x) T"
|
|
655 |
apply(auto intro!: a4 simp add: valid_eqvt)
|
|
656 |
apply(simp add: pt_list_set_pi[OF pt_name_inst, symmetric])
|
|
657 |
apply(perm_simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst])
|
|
658 |
done
|
|
659 |
next
|
|
660 |
case (QAP_App \<Gamma> p q T1 T2 s t)
|
|
661 |
then show "\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>) (pi\<bullet>App p s) (pi\<bullet>App q t) T2"
|
|
662 |
by (auto intro!: a5 simp add: alg_equiv_alg_path_equiv_eqvt)
|
|
663 |
next
|
|
664 |
case (QAP_Const \<Gamma> n)
|
|
665 |
then show "\<forall>c (pi::name prm). P2 c (pi\<bullet>\<Gamma>) (pi\<bullet>Const n) (pi\<bullet>Const n) TBase"
|
|
666 |
by (auto intro!: a6 simp add: valid_eqvt)
|
|
667 |
qed
|
|
668 |
then have "(\<Gamma> \<turnstile> s <=> t : T \<longrightarrow> P1 c (([]::name prm)\<bullet>\<Gamma>) (([]::name prm)\<bullet>s) (([]::name prm)\<bullet>t) T) \<and>
|
|
669 |
(\<Gamma>' \<turnstile> s' \<leftrightarrow> t' : T' \<longrightarrow> P2 c (([]::name prm)\<bullet>\<Gamma>') (([]::name prm)\<bullet>s') (([]::name prm)\<bullet>t') T')"
|
|
670 |
by blast
|
|
671 |
then show ?thesis by simp
|
|
672 |
qed
|
|
673 |
|
22082
|
674 |
(* post-processing of the strong induction principle proved above; *)
|
|
675 |
(* the code extracts the strong_inducts-version from strong_induct *)
|
22073
|
676 |
setup {*
|
|
677 |
PureThy.add_thmss
|
|
678 |
[(("alg_equiv_alg_path_equiv_strong_inducts",
|
|
679 |
ProjectRule.projects (ProofContext.init (the_context())) [1,2]
|
|
680 |
(thm "alg_equiv_alg_path_equiv_strong_induct")), [])] #> snd
|
|
681 |
*}
|
|
682 |
|
|
683 |
declare trm.inject [simp add]
|
|
684 |
declare ty.inject [simp add]
|
|
685 |
|
|
686 |
inductive_cases2 whn_inv_auto[elim]: "t \<Down> t'"
|
|
687 |
|
|
688 |
inductive_cases2 alg_equiv_Base_inv_auto[elim]: "\<Gamma> \<turnstile> s<=>t : TBase"
|
|
689 |
inductive_cases2 alg_equiv_Arrow_inv_auto[elim]: "\<Gamma> \<turnstile> s<=>t : T1 \<rightarrow> T2"
|
|
690 |
|
|
691 |
inductive_cases2 alg_path_equiv_Base_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : TBase"
|
|
692 |
inductive_cases2 alg_path_equiv_Unit_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : TUnit"
|
|
693 |
inductive_cases2 alg_path_equiv_Arrow_inv_auto[elim]: "\<Gamma> \<turnstile> s\<leftrightarrow>t : T1 \<rightarrow> T2"
|
|
694 |
|
|
695 |
inductive_cases2 alg_path_equiv_Var_left_inv_auto[elim]: "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T"
|
|
696 |
inductive_cases2 alg_path_equiv_Var_left_inv_auto'[elim]: "\<Gamma> \<turnstile> Var x \<leftrightarrow> t : T'"
|
|
697 |
inductive_cases2 alg_path_equiv_Var_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T"
|
|
698 |
inductive_cases2 alg_path_equiv_Var_right_inv_auto'[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Var x : T'"
|
|
699 |
inductive_cases2 alg_path_equiv_Const_left_inv_auto[elim]: "\<Gamma> \<turnstile> Const n \<leftrightarrow> t : T"
|
|
700 |
inductive_cases2 alg_path_equiv_Const_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> Const n : T"
|
|
701 |
inductive_cases2 alg_path_equiv_App_left_inv_auto[elim]: "\<Gamma> \<turnstile> App p s \<leftrightarrow> t : T"
|
|
702 |
inductive_cases2 alg_path_equiv_App_right_inv_auto[elim]: "\<Gamma> \<turnstile> s \<leftrightarrow> App q t : T"
|
|
703 |
inductive_cases2 alg_path_equiv_Lam_left_inv_auto[elim]: "\<Gamma> \<turnstile> Lam[x].s \<leftrightarrow> t : T"
|
|
704 |
inductive_cases2 alg_path_equiv_Lam_right_inv_auto[elim]: "\<Gamma> \<turnstile> t \<leftrightarrow> Lam[x].s : T"
|
|
705 |
|
|
706 |
declare trm.inject [simp del]
|
|
707 |
declare ty.inject [simp del]
|
|
708 |
|
|
709 |
text {* Subcontext. *}
|
|
710 |
|
|
711 |
abbreviation
|
|
712 |
"sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [55,55] 55)
|
|
713 |
where
|
|
714 |
"\<Gamma>1 \<lless> \<Gamma>2 \<equiv> (\<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow> (a,\<sigma>)\<in>set \<Gamma>2) \<and> valid \<Gamma>2"
|
|
715 |
|
|
716 |
lemma alg_equiv_valid:
|
|
717 |
shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> valid \<Gamma>" and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> valid \<Gamma>"
|
|
718 |
by (induct rule : alg_equiv_alg_path_equiv.inducts, auto)
|
|
719 |
|
|
720 |
lemma algorithmic_monotonicity:
|
|
721 |
fixes \<Gamma>':: "(name\<times>ty) list"
|
|
722 |
shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s <=> t : T"
|
|
723 |
and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma>\<lless>\<Gamma>' \<Longrightarrow> \<Gamma>' \<turnstile> s \<leftrightarrow> t : T"
|
|
724 |
proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: \<Gamma>' rule: alg_equiv_alg_path_equiv_strong_inducts)
|
|
725 |
case (QAT_Arrow x \<Gamma> s t T1 T2 \<Gamma>')
|
|
726 |
have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
|
|
727 |
have h1:"(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2" by fact
|
|
728 |
have h2:"\<Gamma>\<lless>\<Gamma>'" by fact
|
|
729 |
have ih:"\<And>\<Gamma>'. \<lbrakk>(x,T1)#\<Gamma> \<lless> \<Gamma>'\<rbrakk> \<Longrightarrow> \<Gamma>' \<turnstile> App s (Var x) <=> App t (Var x) : T2" by fact
|
|
730 |
have "x\<sharp>\<Gamma>'" by fact
|
|
731 |
then have sub:"(x,T1)#\<Gamma> \<lless> (x,T1)#\<Gamma>'" using h2 by auto
|
|
732 |
then have "(x,T1)#\<Gamma>' \<turnstile> App s (Var x) <=> App t (Var x) : T2" using ih by auto
|
|
733 |
then show "\<Gamma>' \<turnstile> s <=> t : T1\<rightarrow>T2" using sub fs by auto
|
|
734 |
qed (auto)
|
|
735 |
|
|
736 |
text {* Logical equivalence. *}
|
|
737 |
|
|
738 |
function log_equiv :: "((name\<times>ty) list \<Rightarrow> trm \<Rightarrow> trm \<Rightarrow> ty \<Rightarrow> bool)"
|
|
739 |
("_ \<turnstile> _ is _ : _" [60,60,60,60] 60)
|
|
740 |
where
|
|
741 |
"\<Gamma> \<turnstile> s is t : TUnit = valid \<Gamma>"
|
|
742 |
| "\<Gamma> \<turnstile> s is t : TBase = \<Gamma> \<turnstile> s <=> t : TBase"
|
22082
|
743 |
| "\<Gamma> \<turnstile> s is t : (T1 \<rightarrow> T2) =
|
|
744 |
(valid \<Gamma> \<and> (\<forall>\<Gamma>' s' t'. \<Gamma>\<lless>\<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T2)))"
|
22073
|
745 |
apply (auto simp add: ty.inject)
|
22082
|
746 |
apply (subgoal_tac "(\<exists>T1 T2. b=T1 \<rightarrow> T2) \<or> b=TUnit \<or> b=TBase" )
|
22073
|
747 |
apply (force)
|
|
748 |
apply (rule ty_cases)
|
|
749 |
done
|
|
750 |
|
|
751 |
termination
|
|
752 |
apply(relation "measure (\<lambda>(_,_,_,T). size T)")
|
|
753 |
apply(auto)
|
|
754 |
done
|
|
755 |
|
|
756 |
lemma log_equiv_valid:
|
|
757 |
assumes "\<Gamma> \<turnstile> s is t : T"
|
|
758 |
shows "valid \<Gamma>"
|
|
759 |
using assms
|
|
760 |
by (induct rule: log_equiv.induct,auto elim: alg_equiv_valid)
|
|
761 |
|
|
762 |
lemma logical_monotonicity :
|
|
763 |
fixes T::ty
|
|
764 |
assumes "\<Gamma> \<turnstile> s is t : T" "\<Gamma>\<lless>\<Gamma>'"
|
|
765 |
shows "\<Gamma>' \<turnstile> s is t : T"
|
|
766 |
using assms
|
|
767 |
proof (induct arbitrary: \<Gamma>' rule: log_equiv.induct)
|
|
768 |
case (2 \<Gamma> s t \<Gamma>')
|
|
769 |
then show "\<Gamma>' \<turnstile> s is t : TBase" using algorithmic_monotonicity by auto
|
|
770 |
next
|
|
771 |
case (3 \<Gamma> s t T1 T2 \<Gamma>')
|
|
772 |
have h1:"\<Gamma> \<turnstile> s is t : T1\<rightarrow>T2" by fact
|
|
773 |
have h2:"\<Gamma>\<lless>\<Gamma>'" by fact
|
|
774 |
{
|
|
775 |
fix \<Gamma>'' s' t'
|
|
776 |
assume "\<Gamma>'\<lless>\<Gamma>''" "\<Gamma>'' \<turnstile> s' is t' : T1"
|
|
777 |
then have "\<Gamma>'' \<turnstile> (App s s') is (App t t') : T2" using h1 h2 by auto
|
|
778 |
}
|
|
779 |
then show "\<Gamma>' \<turnstile> s is t : T1\<rightarrow>T2" using h2 by auto
|
|
780 |
qed (auto)
|
|
781 |
|
|
782 |
lemma forget:
|
|
783 |
fixes x::"name"
|
|
784 |
and L::"trm"
|
|
785 |
assumes a: "x\<sharp>L"
|
|
786 |
shows "L[x::=P] = L"
|
|
787 |
using a
|
22082
|
788 |
by (nominal_induct L avoiding: x P rule: trm.induct)
|
|
789 |
(auto simp add: fresh_atm abs_fresh)
|
22073
|
790 |
|
|
791 |
lemma subst_fun_eq:
|
|
792 |
fixes u::trm
|
|
793 |
assumes h:"[x].t1 = [y].t2"
|
|
794 |
shows "t1[x::=u] = t2[y::=u]"
|
|
795 |
proof -
|
|
796 |
{
|
|
797 |
assume "x=y" and "t1=t2"
|
|
798 |
then have ?thesis using h by simp
|
|
799 |
}
|
|
800 |
moreover
|
|
801 |
{
|
|
802 |
assume h1:"x \<noteq> y" and h2:"t1=[(x,y)] \<bullet> t2" and h3:"x \<sharp> t2"
|
|
803 |
then have "([(x,y)] \<bullet> t2)[x::=u] = t2[y::=u]" by (simp add: subst_rename)
|
|
804 |
then have ?thesis using h2 by simp
|
|
805 |
}
|
|
806 |
ultimately show ?thesis using alpha h by blast
|
|
807 |
qed
|
|
808 |
|
|
809 |
lemma psubst_empty[simp]:
|
|
810 |
shows "[]<t> = t"
|
|
811 |
by (nominal_induct t rule: trm.induct) (auto simp add: fresh_list_nil)
|
|
812 |
|
|
813 |
lemma psubst_subst_psubst:
|
|
814 |
assumes h:"c \<sharp> \<theta>"
|
|
815 |
shows "\<theta><t>[c::=s] = (c,s)#\<theta><t>"
|
|
816 |
using h
|
22082
|
817 |
by (nominal_induct t avoiding: \<theta> c s rule: trm.induct)
|
|
818 |
(auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst)
|
22073
|
819 |
|
|
820 |
lemma subst_fresh_simpl:
|
22082
|
821 |
assumes a: "x\<sharp>\<theta>"
|
22073
|
822 |
shows "\<theta><Var x> = Var x"
|
22082
|
823 |
using a
|
|
824 |
by (induct \<theta> arbitrary: x, auto simp add:fresh_list_cons fresh_prod fresh_atm)
|
22073
|
825 |
|
|
826 |
lemma psubst_subst_propagate:
|
|
827 |
assumes "x\<sharp>\<theta>"
|
|
828 |
shows "\<theta><t[x::=u]> = \<theta><t>[x::=\<theta><u>]"
|
|
829 |
using assms
|
|
830 |
proof (nominal_induct t avoiding: x u \<theta> rule: trm.induct)
|
|
831 |
case (Var n x u \<theta>)
|
|
832 |
{ assume "x=n"
|
|
833 |
moreover have "x\<sharp>\<theta>" by fact
|
|
834 |
ultimately have "\<theta><Var n[x::=u]> = \<theta><Var n>[x::=\<theta><u>]" using subst_fresh_simpl by auto
|
|
835 |
}
|
|
836 |
moreover
|
|
837 |
{ assume h:"x\<noteq>n"
|
|
838 |
then have "x\<sharp>Var n" by (auto simp add: fresh_atm)
|
|
839 |
moreover have "x\<sharp>\<theta>" by fact
|
|
840 |
ultimately have "x\<sharp>\<theta><Var n>" using fresh_psubst by blast
|
|
841 |
then have " \<theta><Var n>[x::=\<theta><u>] = \<theta><Var n>" using forget by auto
|
|
842 |
then have "\<theta><Var n[x::=u]> = \<theta><Var n>[x::=\<theta><u>]" using h by auto
|
|
843 |
}
|
|
844 |
ultimately show ?case by auto
|
|
845 |
next
|
|
846 |
case (Lam n t x u \<theta>)
|
|
847 |
have fs:"n\<sharp>x" "n\<sharp>u" "n\<sharp>\<theta>" "x\<sharp>\<theta>" by fact
|
|
848 |
have ih:"\<And> y s \<theta>. y\<sharp>\<theta> \<Longrightarrow> ((\<theta><(t[y::=s])>) = ((\<theta><t>)[y::=(\<theta><s>)]))" by fact
|
|
849 |
have "\<theta> <(Lam [n].t)[x::=u]> = \<theta><Lam [n]. (t [x::=u])>" using fs by auto
|
|
850 |
then have "\<theta> <(Lam [n].t)[x::=u]> = Lam [n]. \<theta><t [x::=u]>" using fs by auto
|
|
851 |
moreover have "\<theta><t[x::=u]> = \<theta><t>[x::=\<theta><u>]" using ih fs by blast
|
|
852 |
ultimately have "\<theta> <(Lam [n].t)[x::=u]> = Lam [n].(\<theta><t>[x::=\<theta><u>])" by auto
|
|
853 |
moreover have "Lam [n].(\<theta><t>[x::=\<theta><u>]) = (Lam [n].\<theta><t>)[x::=\<theta><u>]" using fs fresh_psubst by auto
|
|
854 |
ultimately have "\<theta><(Lam [n].t)[x::=u]> = (Lam [n].\<theta><t>)[x::=\<theta><u>]" using fs by auto
|
|
855 |
then show "\<theta><(Lam [n].t)[x::=u]> = \<theta><Lam [n].t>[x::=\<theta><u>]" using fs by auto
|
|
856 |
qed (auto)
|
|
857 |
|
|
858 |
lemma fresh_subst_fresh:
|
|
859 |
assumes "a\<sharp>e"
|
|
860 |
shows "a\<sharp>t[a::=e]"
|
|
861 |
using assms
|
|
862 |
by (nominal_induct t avoiding: a e rule: trm.induct)
|
|
863 |
(auto simp add: fresh_atm abs_fresh fresh_nat)
|
|
864 |
|
|
865 |
lemma path_equiv_implies_nf:
|
|
866 |
assumes "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
|
|
867 |
shows "s \<leadsto>|" and "t \<leadsto>|"
|
|
868 |
using assms
|
22082
|
869 |
by (induct rule: alg_equiv_alg_path_equiv.inducts(2)) (simp, auto)
|
22073
|
870 |
|
|
871 |
lemma path_equiv_implies_nf:
|
|
872 |
shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> True"
|
|
873 |
and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> s \<leadsto>|"
|
|
874 |
"\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> t \<leadsto>|"
|
22082
|
875 |
by (induct rule: alg_equiv_alg_path_equiv.inducts) (auto)
|
22073
|
876 |
|
|
877 |
lemma main_lemma:
|
|
878 |
shows "\<Gamma> \<turnstile> s is t : T \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T" and "\<Gamma> \<turnstile> p \<leftrightarrow> q : T \<Longrightarrow> \<Gamma> \<turnstile> p is q : T"
|
|
879 |
proof (nominal_induct T arbitrary: \<Gamma> s t p q rule:ty.induct)
|
|
880 |
case (Arrow T1 T2)
|
|
881 |
{
|
|
882 |
case (1 \<Gamma> s t)
|
|
883 |
have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s is t : T2 \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T2" by fact
|
|
884 |
have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T1 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T1" by fact
|
|
885 |
have h:"\<Gamma> \<turnstile> s is t : T1\<rightarrow>T2" by fact
|
|
886 |
obtain x::name where fs:"x\<sharp>(\<Gamma>,s,t)" by (erule exists_fresh[OF fs_name1])
|
|
887 |
have "valid \<Gamma>" using h log_equiv_valid by auto
|
|
888 |
then have v:"valid ((x,T1)#\<Gamma>)" using fs by auto
|
|
889 |
then have "(x,T1)#\<Gamma> \<turnstile> Var x \<leftrightarrow> Var x : T1" by auto
|
|
890 |
then have "(x,T1)#\<Gamma> \<turnstile> Var x is Var x : T1" using ih2 by auto
|
|
891 |
then have "(x,T1)#\<Gamma> \<turnstile> App s (Var x) is App t (Var x) : T2" using h v by auto
|
|
892 |
then have "(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App t (Var x) : T2" using ih1 by auto
|
|
893 |
then show "\<Gamma> \<turnstile> s <=> t : T1\<rightarrow>T2" using fs by (auto simp add: fresh_prod)
|
|
894 |
next
|
|
895 |
case (2 \<Gamma> p q)
|
|
896 |
have h:"\<Gamma> \<turnstile> p \<leftrightarrow> q : T1\<rightarrow>T2" by fact
|
|
897 |
have ih1:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s \<leftrightarrow> t : T2 \<Longrightarrow> \<Gamma> \<turnstile> s is t : T2" by fact
|
|
898 |
have ih2:"\<And>\<Gamma> s t. \<Gamma> \<turnstile> s is t : T1 \<Longrightarrow> \<Gamma> \<turnstile> s <=> t : T1" by fact
|
|
899 |
{
|
|
900 |
fix \<Gamma>' s t
|
|
901 |
assume "\<Gamma>\<lless>\<Gamma>'" and hl:"\<Gamma>' \<turnstile> s is t : T1"
|
|
902 |
then have "\<Gamma>' \<turnstile> p \<leftrightarrow> q : T1 \<rightarrow> T2" using h algorithmic_monotonicity by auto
|
|
903 |
moreover have "\<Gamma>' \<turnstile> s <=> t : T1" using ih2 hl by auto
|
|
904 |
ultimately have "\<Gamma>' \<turnstile> App p s \<leftrightarrow> App q t : T2" by auto
|
|
905 |
then have "\<Gamma>' \<turnstile> App p s is App q t : T2" using ih1 by auto
|
|
906 |
}
|
|
907 |
moreover have "valid \<Gamma>" using h alg_equiv_valid by auto
|
|
908 |
ultimately show "\<Gamma> \<turnstile> p is q : T1\<rightarrow>T2" by auto
|
|
909 |
}
|
|
910 |
next
|
|
911 |
case TBase
|
|
912 |
{ case 2
|
|
913 |
have h:"\<Gamma> \<turnstile> s \<leftrightarrow> t : TBase" by fact
|
|
914 |
then have "s \<leadsto>|" and "t \<leadsto>|" using path_equiv_implies_nf by auto
|
|
915 |
then have "s \<Down> s" and "t \<Down> t" by auto
|
|
916 |
then have "\<Gamma> \<turnstile> s <=> t : TBase" using h by auto
|
|
917 |
then show "\<Gamma> \<turnstile> s is t : TBase" by auto
|
|
918 |
}
|
|
919 |
qed (auto elim:alg_equiv_valid)
|
|
920 |
|
|
921 |
corollary corollary_main:
|
22082
|
922 |
assumes a: "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
|
22073
|
923 |
shows "\<Gamma> \<turnstile> s <=> t : T"
|
22082
|
924 |
using a main_lemma by blast
|
22073
|
925 |
|
|
926 |
lemma algorithmic_symmetry_aux:
|
22082
|
927 |
shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> \<Gamma> \<turnstile> t <=> s : T"
|
|
928 |
and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> \<Gamma> \<turnstile> t \<leftrightarrow> s : T"
|
|
929 |
by (induct rule: alg_equiv_alg_path_equiv.inducts) (auto)
|
22073
|
930 |
|
|
931 |
lemma algorithmic_symmetry:
|
22082
|
932 |
assumes a: "\<Gamma> \<turnstile> s <=> t : T"
|
22073
|
933 |
shows "\<Gamma> \<turnstile> t <=> s : T"
|
22082
|
934 |
using a by (simp add: algorithmic_symmetry_aux)
|
22073
|
935 |
|
|
936 |
lemma algorithmic_path_symmetry:
|
22082
|
937 |
assumes a: "\<Gamma> \<turnstile> s \<leftrightarrow> t : T"
|
22073
|
938 |
shows "\<Gamma> \<turnstile> t \<leftrightarrow> s : T"
|
22082
|
939 |
using a by (simp add: algorithmic_symmetry_aux)
|
22073
|
940 |
|
|
941 |
lemma red_unicity :
|
22082
|
942 |
assumes a: "x \<leadsto> a"
|
|
943 |
and b: "x \<leadsto> b"
|
22073
|
944 |
shows "a=b"
|
22082
|
945 |
using a b
|
22073
|
946 |
apply (induct arbitrary: b)
|
|
947 |
apply (erule whr_App_Lam)
|
|
948 |
apply (clarify)
|
|
949 |
apply (rule subst_fun_eq)
|
22082
|
950 |
apply (simp)
|
22073
|
951 |
apply (force)
|
|
952 |
apply (erule whr_App)
|
|
953 |
apply (blast)+
|
|
954 |
done
|
|
955 |
|
|
956 |
lemma nf_unicity :
|
|
957 |
assumes "x \<Down> a" "x \<Down> b"
|
|
958 |
shows "a=b"
|
|
959 |
using assms
|
|
960 |
proof (induct arbitrary: b)
|
|
961 |
case (QAN_Reduce x t a b)
|
|
962 |
have h:"x \<leadsto> t" "t \<Down> a" by fact
|
|
963 |
have ih:"\<And>b. t \<Down> b \<Longrightarrow> a = b" by fact
|
|
964 |
have "x \<Down> b" by fact
|
|
965 |
then obtain t' where "x \<leadsto> t'" and hl:"t' \<Down> b" using h by auto
|
|
966 |
then have "t=t'" using h red_unicity by auto
|
|
967 |
then show "a=b" using ih hl by auto
|
|
968 |
qed (auto)
|
|
969 |
|
|
970 |
lemma Q_eqvt :
|
22082
|
971 |
fixes pi:: "name prm"
|
22073
|
972 |
shows "\<Gamma> \<turnstile> s <=> t : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) <=> (pi\<bullet>t) : T"
|
|
973 |
and "\<Gamma> \<turnstile> s \<leftrightarrow> t : T \<Longrightarrow> (pi\<bullet>\<Gamma>) \<turnstile> (pi\<bullet>s) \<leftrightarrow> (pi\<bullet>t) : T"
|
|
974 |
using assms
|
|
975 |
proof (induct rule:alg_equiv_alg_path_equiv.inducts)
|
|
976 |
case (QAP_Var \<Gamma> x T)
|
|
977 |
then have "pi\<bullet>(x,T) \<in> (pi\<bullet>(set \<Gamma>))" using in_eqvt by blast
|
|
978 |
then have "(pi\<bullet>x,T) \<in> (set (pi\<bullet>\<Gamma>))" using set_eqvt perm_ty by auto
|
|
979 |
moreover have "valid \<Gamma>" by fact
|
|
980 |
ultimately show ?case using valid_eqvt by auto
|
|
981 |
next
|
|
982 |
case (QAT_Arrow x \<Gamma> s t T1 T2)
|
|
983 |
then have h:"((pi\<bullet>x)\<sharp>(pi\<bullet>\<Gamma>))" "((pi\<bullet>x)\<sharp>(pi\<bullet>s))" "((pi\<bullet>x)\<sharp>(pi\<bullet>t))" using fresh_bij by auto
|
|
984 |
have "(pi\<bullet>((x,T1)#\<Gamma>)) \<turnstile> pi \<bullet> (App s (Var x)) <=> pi \<bullet> (App t (Var x)) : T2" by fact
|
|
985 |
then have "(pi\<bullet>((x,T1)#\<Gamma>)) \<turnstile> (App (pi\<bullet>s) (Var (pi\<bullet>x))) <=> (App (pi\<bullet>t) (Var (pi\<bullet>x))) : T2" by auto
|
|
986 |
moreover have "pi\<bullet>((x,T1)#\<Gamma>) = (pi\<bullet>x,pi\<bullet>T1)#(pi\<bullet>\<Gamma>)" by auto
|
|
987 |
ultimately have "((pi\<bullet>x,T1)#(pi\<bullet>\<Gamma>)) \<turnstile> (App (pi\<bullet>s) (Var (pi\<bullet>x))) <=> (App (pi\<bullet>t) (Var (pi\<bullet>x))) : T2" using perm_ty by auto
|
|
988 |
then show ?case using h by auto
|
|
989 |
qed (auto elim:whn_eqvt valid_eqvt)
|
|
990 |
|
|
991 |
(* FIXME this lemma should not be here I am reinventing the wheel I guess *)
|
|
992 |
|
|
993 |
lemma perm_basic[simp]:
|
22082
|
994 |
fixes x y::"name"
|
22073
|
995 |
shows "[(x,y)]\<bullet>y = x"
|
|
996 |
using assms using calc_atm by perm_simp
|
|
997 |
|
|
998 |
lemma Q_Arrow_strong_inversion:
|
|
999 |
assumes fs:"x\<sharp>\<Gamma>" "x\<sharp>t" "x\<sharp>u" and h:"\<Gamma> \<turnstile> t <=> u : T1\<rightarrow>T2"
|
|
1000 |
shows "(x,T1)#\<Gamma> \<turnstile> App t (Var x) <=> App u (Var x) : T2"
|
|
1001 |
proof -
|
22082
|
1002 |
obtain y where fs2:"y\<sharp>\<Gamma>" "y\<sharp>u" "y\<sharp>t" and "(y,T1)#\<Gamma> \<turnstile> App t (Var y) <=> App u (Var y) : T2"
|
|
1003 |
using h by auto
|
|
1004 |
then have "([(x,y)]\<bullet>((y,T1)#\<Gamma>)) \<turnstile> [(x,y)]\<bullet> App t (Var y) <=> [(x,y)]\<bullet> App u (Var y) : T2"
|
|
1005 |
using Q_eqvt by blast
|
|
1006 |
then show ?thesis using fs fs2 by (perm_simp)
|
22073
|
1007 |
qed
|
|
1008 |
|
22082
|
1009 |
lemma fresh_context:
|
|
1010 |
fixes \<Gamma> :: "(name\<times>ty) list"
|
22073
|
1011 |
and a :: "name"
|
|
1012 |
assumes "a\<sharp>\<Gamma>"
|
|
1013 |
shows "\<not>(\<exists>\<tau>::ty. (a,\<tau>)\<in>set \<Gamma>)"
|
|
1014 |
using assms by(induct \<Gamma>, auto simp add: fresh_prod fresh_list_cons fresh_atm)
|
|
1015 |
|
|
1016 |
lemma type_unicity_in_context:
|
|
1017 |
fixes \<Gamma> :: "(name\<times>ty)list"
|
|
1018 |
and pi:: "name prm"
|
|
1019 |
and a :: "name"
|
|
1020 |
and \<tau> :: "ty"
|
|
1021 |
assumes "valid \<Gamma>" and "(x,T1) \<in> set \<Gamma>" and "(x,T2) \<in> set \<Gamma>"
|
|
1022 |
shows "T1=T2"
|
|
1023 |
using assms by (induct \<Gamma>, auto dest!: fresh_context)
|
|
1024 |
|
|
1025 |
(*
|
|
1026 |
|
|
1027 |
Warning: This lemma is false.
|
|
1028 |
|
|
1029 |
lemma algorithmic_type_unicity:
|
|
1030 |
shows "\<lbrakk> \<Gamma> \<turnstile> s <=> t : T ; \<Gamma> \<turnstile> s <=> u : T' \<rbrakk> \<Longrightarrow> T = T'"
|
|
1031 |
and "\<lbrakk> \<Gamma> \<turnstile> s \<leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"
|
|
1032 |
|
|
1033 |
Here is the counter example :
|
|
1034 |
\<Gamma> \<turnstile> Const n <=> Const n : Tbase and \<Gamma> \<turnstile> Const n <=> Const n : TUnit
|
|
1035 |
|
|
1036 |
*)
|
|
1037 |
|
|
1038 |
lemma algorithmic_path_type_unicity:
|
|
1039 |
shows "\<lbrakk> \<Gamma> \<turnstile> s \<leftrightarrow> t : T ; \<Gamma> \<turnstile> s \<leftrightarrow> u : T' \<rbrakk> \<Longrightarrow> T = T'"
|
22082
|
1040 |
proof (induct arbitrary: u T'
|
|
1041 |
rule: alg_equiv_alg_path_equiv.inducts(2) [of _ _ _ _ _ "%a b c d . True" ])
|
22073
|
1042 |
case (QAP_Var \<Gamma> x T u T')
|
|
1043 |
have "\<Gamma> \<turnstile> Var x \<leftrightarrow> u : T'" by fact
|
|
1044 |
then have "u=Var x" and "(x,T') \<in> set \<Gamma>" by auto
|
|
1045 |
moreover have "valid \<Gamma>" "(x,T) \<in> set \<Gamma>" by fact
|
|
1046 |
ultimately show "T=T'" using type_unicity_in_context by auto
|
|
1047 |
next
|
|
1048 |
case (QAP_App \<Gamma> p q T1 T2 s t u T2')
|
|
1049 |
have ih:"\<And>u T. \<Gamma> \<turnstile> p \<leftrightarrow> u : T \<Longrightarrow> T1\<rightarrow>T2 = T" by fact
|
|
1050 |
have "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T2'" by fact
|
|
1051 |
then obtain r t T1' where "u = App r t" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T1' \<rightarrow> T2'" by auto
|
|
1052 |
then have "T1\<rightarrow>T2 = T1' \<rightarrow> T2'" by auto
|
|
1053 |
then show "T2=T2'" using ty.inject by auto
|
|
1054 |
qed (auto)
|
|
1055 |
|
|
1056 |
lemma algorithmic_transitivity:
|
|
1057 |
shows "\<lbrakk> \<Gamma> \<turnstile> s <=> t : T ; \<Gamma> \<turnstile> t <=> u : T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s <=> u : T"
|
|
1058 |
and "\<lbrakk> \<Gamma> \<turnstile> s \<leftrightarrow> t : T ; \<Gamma> \<turnstile> t \<leftrightarrow> u : T \<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s \<leftrightarrow> u : T"
|
|
1059 |
proof (nominal_induct \<Gamma> s t T and \<Gamma> s t T avoiding: u rule: alg_equiv_alg_path_equiv_strong_inducts)
|
|
1060 |
case (QAT_Base s p t q \<Gamma> u)
|
|
1061 |
have h:"s \<Down> p" "t \<Down> q" by fact
|
|
1062 |
have ih:"\<And>u. \<Gamma> \<turnstile> q \<leftrightarrow> u : TBase \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> u : TBase" by fact
|
|
1063 |
have "\<Gamma> \<turnstile> t <=> u : TBase" by fact
|
|
1064 |
then obtain r q' where hl:"t \<Down> q'" "u \<Down> r" "\<Gamma> \<turnstile> q' \<leftrightarrow> r : TBase" by auto
|
|
1065 |
moreover have eq:"q=q'" using nf_unicity hl h by auto
|
|
1066 |
ultimately have "\<Gamma> \<turnstile> p \<leftrightarrow> r : TBase" using ih by auto
|
|
1067 |
then show "\<Gamma> \<turnstile> s <=> u : TBase" using hl h by auto
|
|
1068 |
next
|
|
1069 |
case (QAT_Arrow x \<Gamma> s t T1 T2 u)
|
|
1070 |
have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
|
|
1071 |
moreover have h:"\<Gamma> \<turnstile> t <=> u : T1\<rightarrow>T2" by fact
|
22082
|
1072 |
moreover then obtain y where "y\<sharp>\<Gamma>" "y\<sharp>t" "y\<sharp>u" and "(y,T1)#\<Gamma> \<turnstile> App t (Var y) <=> App u (Var y) : T2"
|
|
1073 |
by auto
|
22073
|
1074 |
moreover have fs2:"x\<sharp>u" by fact
|
|
1075 |
ultimately have "(x,T1)#\<Gamma> \<turnstile> App t (Var x) <=> App u (Var x) : T2" using Q_Arrow_strong_inversion by blast
|
22082
|
1076 |
moreover have ih:"\<And> v. (x,T1)#\<Gamma> \<turnstile> App t (Var x) <=> v : T2 \<Longrightarrow> (x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> v : T2"
|
|
1077 |
by fact
|
22073
|
1078 |
ultimately have "(x,T1)#\<Gamma> \<turnstile> App s (Var x) <=> App u (Var x) : T2" by auto
|
|
1079 |
then show "\<Gamma> \<turnstile> s <=> u : T1\<rightarrow>T2" using fs fs2 by auto
|
|
1080 |
next
|
|
1081 |
case (QAP_App \<Gamma> p q T1 T2 s t u)
|
|
1082 |
have h1:"\<Gamma> \<turnstile> p \<leftrightarrow> q : T1\<rightarrow>T2" by fact
|
|
1083 |
have ih1:"\<And>u. \<Gamma> \<turnstile> q \<leftrightarrow> u : T1\<rightarrow>T2 \<Longrightarrow> \<Gamma> \<turnstile> p \<leftrightarrow> u : T1\<rightarrow>T2" by fact
|
|
1084 |
have "\<Gamma> \<turnstile> s <=> t : T1" by fact
|
|
1085 |
have ih2:"\<And>u. \<Gamma> \<turnstile> t <=> u : T1 \<Longrightarrow> \<Gamma> \<turnstile> s <=> u : T1" by fact
|
|
1086 |
have "\<Gamma> \<turnstile> App q t \<leftrightarrow> u : T2" by fact
|
22082
|
1087 |
then obtain r T1' v where ha:"\<Gamma> \<turnstile> q \<leftrightarrow> r : T1'\<rightarrow>T2" and hb:"\<Gamma> \<turnstile> t <=> v : T1'" and eq:"u = App r v"
|
|
1088 |
by auto
|
22073
|
1089 |
moreover have "\<Gamma> \<turnstile> q \<leftrightarrow> p : T1\<rightarrow>T2" using h1 algorithmic_path_symmetry by auto
|
|
1090 |
ultimately have "T1'\<rightarrow>T2 = T1\<rightarrow>T2" using algorithmic_path_type_unicity by auto
|
|
1091 |
then have "T1' = T1" using ty.inject by auto
|
|
1092 |
then have "\<Gamma> \<turnstile> s <=> v : T1" "\<Gamma> \<turnstile> p \<leftrightarrow> r : T1\<rightarrow>T2" using ih1 ih2 ha hb by auto
|
|
1093 |
then show "\<Gamma> \<turnstile> App p s \<leftrightarrow> u : T2" using eq by auto
|
|
1094 |
qed (auto)
|
|
1095 |
|
|
1096 |
lemma algorithmic_weak_head_closure:
|
|
1097 |
shows "\<lbrakk>\<Gamma> \<turnstile> s <=> t : T ; s' \<leadsto> s; t' \<leadsto> t\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' <=> t' : T"
|
|
1098 |
by (nominal_induct \<Gamma> s t T avoiding: s' t'
|
|
1099 |
rule: alg_equiv_alg_path_equiv_strong_inducts(1) [of _ _ _ _ "%a b c d e. True"])
|
|
1100 |
(auto)
|
|
1101 |
|
|
1102 |
lemma logical_symmetry:
|
22082
|
1103 |
assumes a: "\<Gamma> \<turnstile> s is t : T"
|
22073
|
1104 |
shows "\<Gamma> \<turnstile> t is s : T"
|
22082
|
1105 |
using a
|
|
1106 |
by (nominal_induct arbitrary: \<Gamma> s t rule:ty.induct) (auto simp add: algorithmic_symmetry)
|
22073
|
1107 |
|
|
1108 |
lemma logical_transitivity:
|
|
1109 |
assumes "\<Gamma> \<turnstile> s is t : T" "\<Gamma> \<turnstile> t is u : T"
|
|
1110 |
shows "\<Gamma> \<turnstile> s is u : T"
|
|
1111 |
using assms
|
|
1112 |
proof (nominal_induct arbitrary: \<Gamma> s t u rule:ty.induct)
|
|
1113 |
case TBase
|
|
1114 |
then show "\<Gamma> \<turnstile> s is u : TBase" by (auto elim: algorithmic_transitivity)
|
|
1115 |
next
|
|
1116 |
case (Arrow T1 T2 \<Gamma> s t u)
|
|
1117 |
have h1:"\<Gamma> \<turnstile> s is t : T1 \<rightarrow> T2" by fact
|
|
1118 |
have h2:"\<Gamma> \<turnstile> t is u : T1 \<rightarrow> T2" by fact
|
|
1119 |
have ih1:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T1; \<Gamma> \<turnstile> t is u : T1\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T1" by fact
|
|
1120 |
have ih2:"\<And>\<Gamma> s t u. \<lbrakk>\<Gamma> \<turnstile> s is t : T2; \<Gamma> \<turnstile> t is u : T2\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s is u : T2" by fact
|
|
1121 |
{
|
|
1122 |
fix \<Gamma>' s' u'
|
|
1123 |
assume hsub:"\<Gamma>\<lless>\<Gamma>'" and hl:"\<Gamma>' \<turnstile> s' is u' : T1"
|
|
1124 |
then have "\<Gamma>' \<turnstile> u' is s' : T1" using logical_symmetry by blast
|
|
1125 |
then have "\<Gamma>' \<turnstile> u' is u' : T1" using ih1 hl by blast
|
|
1126 |
then have "\<Gamma>' \<turnstile> App t u' is App u u' : T2" using h2 hsub by auto
|
|
1127 |
moreover have "\<Gamma>' \<turnstile> App s s' is App t u' : T2" using h1 hsub hl by auto
|
|
1128 |
ultimately have "\<Gamma>' \<turnstile> App s s' is App u u' : T2" using ih2 by blast
|
|
1129 |
}
|
|
1130 |
moreover have "valid \<Gamma>" using h1 alg_equiv_valid by auto
|
|
1131 |
ultimately show "\<Gamma> \<turnstile> s is u : T1 \<rightarrow> T2" by auto
|
|
1132 |
qed (auto)
|
|
1133 |
|
|
1134 |
lemma logical_weak_head_closure:
|
22082
|
1135 |
assumes a: "\<Gamma> \<turnstile> s is t : T" "s' \<leadsto> s" "t' \<leadsto> t"
|
22073
|
1136 |
shows "\<Gamma> \<turnstile> s' is t' : T"
|
22082
|
1137 |
using a
|
|
1138 |
apply(nominal_induct arbitrary: \<Gamma> s t s' t' rule:ty.induct)
|
|
1139 |
apply(auto simp add: algorithmic_weak_head_closure)
|
|
1140 |
apply(blast)
|
22073
|
1141 |
done
|
|
1142 |
|
|
1143 |
lemma logical_weak_head_closure':
|
|
1144 |
assumes "\<Gamma> \<turnstile> s is t : T" "s' \<leadsto> s"
|
|
1145 |
shows "\<Gamma> \<turnstile> s' is t : T"
|
|
1146 |
using assms
|
22082
|
1147 |
proof (nominal_induct arbitrary: \<Gamma> s t s' rule: ty.induct)
|
22073
|
1148 |
case (TBase \<Gamma> s t s')
|
|
1149 |
then show ?case by force
|
|
1150 |
next
|
|
1151 |
case (TUnit \<Gamma> s t s')
|
|
1152 |
then show ?case by auto
|
|
1153 |
next
|
|
1154 |
case (Arrow T1 T2 \<Gamma> s t s')
|
|
1155 |
have h1:"s' \<leadsto> s" by fact
|
|
1156 |
have ih:"\<And>\<Gamma> s t s'. \<lbrakk>\<Gamma> \<turnstile> s is t : T2; s' \<leadsto> s\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s' is t : T2" by fact
|
|
1157 |
have h2:"\<Gamma> \<turnstile> s is t : T1\<rightarrow>T2" by fact
|
|
1158 |
then have hb:"\<forall>\<Gamma>' s' t'. \<Gamma>\<lless>\<Gamma>' \<longrightarrow> \<Gamma>' \<turnstile> s' is t' : T1 \<longrightarrow> (\<Gamma>' \<turnstile> (App s s') is (App t t') : T2)" by auto
|
|
1159 |
{
|
|
1160 |
fix \<Gamma>' s2 t2
|
|
1161 |
assume "\<Gamma>\<lless>\<Gamma>'" and "\<Gamma>' \<turnstile> s2 is t2 : T1"
|
|
1162 |
then have "\<Gamma>' \<turnstile> (App s s2) is (App t t2) : T2" using hb by auto
|
|
1163 |
moreover have "(App s' s2) \<leadsto> (App s s2)" using h1 by auto
|
|
1164 |
ultimately have "\<Gamma>' \<turnstile> App s' s2 is App t t2 : T2" using ih by auto
|
|
1165 |
}
|
|
1166 |
moreover have "valid \<Gamma>" using h2 log_equiv_valid by auto
|
|
1167 |
ultimately show "\<Gamma> \<turnstile> s' is t : T1\<rightarrow>T2" by auto
|
|
1168 |
qed
|
|
1169 |
|
|
1170 |
abbreviation
|
|
1171 |
log_equiv_subst :: "(name\<times>ty) list \<Rightarrow> (name\<times>trm) list \<Rightarrow> (name\<times>trm) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool"
|
|
1172 |
("_ \<turnstile> _ is _ over _" [60,60] 60)
|
|
1173 |
where
|
|
1174 |
"\<Gamma>' \<turnstile> \<gamma> is \<delta> over \<Gamma> \<equiv> valid \<Gamma>' \<and> (\<forall> x T. (x,T) \<in> set \<Gamma> \<longrightarrow> \<Gamma>' \<turnstile> \<gamma><Var x> is \<delta><Var x> : T)"
|
|
1175 |
|
|
1176 |
lemma logical_pseudo_reflexivity:
|
|
1177 |
assumes "\<Gamma>' \<turnstile> t is s over \<Gamma>"
|
|
1178 |
shows "\<Gamma>' \<turnstile> s is s over \<Gamma>"
|
|
1179 |
proof -
|
|
1180 |
have "\<Gamma>' \<turnstile> t is s over \<Gamma>" by fact
|
|
1181 |
moreover then have "\<Gamma>' \<turnstile> s is t over \<Gamma>" using logical_symmetry by blast
|
|
1182 |
ultimately show "\<Gamma>' \<turnstile> s is s over \<Gamma>" using logical_transitivity by blast
|
|
1183 |
qed
|
|
1184 |
|
|
1185 |
lemma logical_subst_monotonicity :
|
|
1186 |
fixes \<Gamma>::"(name\<times>ty) list"
|
|
1187 |
and \<Gamma>'::"(name\<times>ty) list"
|
|
1188 |
and \<Gamma>''::"(name\<times>ty) list"
|
|
1189 |
assumes "\<Gamma>' \<turnstile> s is t over \<Gamma>" "\<Gamma>'\<lless>\<Gamma>''"
|
|
1190 |
shows "\<Gamma>'' \<turnstile> s is t over \<Gamma>"
|
|
1191 |
using assms logical_monotonicity by blast
|
|
1192 |
|
|
1193 |
lemma equiv_subst_ext :
|
|
1194 |
assumes h1:"\<Gamma>' \<turnstile> \<gamma> is \<delta> over \<Gamma>" and h2:"\<Gamma>' \<turnstile> s is t : T" and fs:"x\<sharp>\<Gamma>"
|
|
1195 |
shows "\<Gamma>' \<turnstile> (x,s)#\<gamma> is (x,t)#\<delta> over (x,T)#\<Gamma>"
|
|
1196 |
using assms
|
|
1197 |
proof -
|
|
1198 |
{
|
|
1199 |
fix y U
|
|
1200 |
assume "(y,U) \<in> set ((x,T)#\<Gamma>)"
|
|
1201 |
moreover
|
|
1202 |
{
|
|
1203 |
assume "(y,U) \<in> set [(x,T)]"
|
|
1204 |
then have "\<Gamma>' \<turnstile> (x,s)#\<gamma><Var y> is (x,t)#\<delta><Var y> : U" by auto
|
|
1205 |
}
|
|
1206 |
moreover
|
|
1207 |
{
|
|
1208 |
assume hl:"(y,U) \<in> set \<Gamma>"
|
|
1209 |
then have "\<not> y\<sharp>\<Gamma>" by (induct \<Gamma>) (auto simp add: fresh_list_cons fresh_atm fresh_prod)
|
|
1210 |
then have hf:"x\<sharp> Var y" using fs by (auto simp add: fresh_atm)
|
|
1211 |
then have "(x,s)#\<gamma><Var y> = \<gamma><Var y>" "(x,t)#\<delta><Var y> = \<delta><Var y>" using fresh_psubst_simpl by blast+
|
|
1212 |
moreover have "\<Gamma>' \<turnstile> \<gamma><Var y> is \<delta><Var y> : U" using h1 hl by auto
|
|
1213 |
ultimately have "\<Gamma>' \<turnstile> (x,s)#\<gamma><Var y> is (x,t)#\<delta><Var y> : U" by auto
|
|
1214 |
}
|
|
1215 |
ultimately have "\<Gamma>' \<turnstile> (x,s)#\<gamma><Var y> is (x,t)#\<delta><Var y> : U" by auto
|
|
1216 |
}
|
|
1217 |
moreover have "valid \<Gamma>'" using h2 log_equiv_valid by blast
|
|
1218 |
ultimately show "\<Gamma>' \<turnstile> (x,s)#\<gamma> is (x,t)#\<delta> over (x,T)#\<Gamma>" by auto
|
|
1219 |
qed
|
|
1220 |
|
|
1221 |
theorem fundamental_theorem:
|
|
1222 |
assumes "\<Gamma> \<turnstile> t : T" "\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>"
|
|
1223 |
shows "\<Gamma>' \<turnstile> \<gamma><t> is \<theta><t> : T"
|
|
1224 |
using assms
|
|
1225 |
proof (nominal_induct avoiding:\<Gamma>' \<gamma> \<theta> rule: typing_induct_strong)
|
|
1226 |
case (t_Lam x \<Gamma> T1 t2 T2 \<Gamma>' \<gamma> \<theta>)
|
|
1227 |
have fs:"x\<sharp>\<gamma>" "x\<sharp>\<theta>" "x\<sharp>\<Gamma>" by fact
|
|
1228 |
have h:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1229 |
have ih:"\<And> \<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><t2> is \<theta><t2> : T2" by fact
|
|
1230 |
{
|
|
1231 |
fix \<Gamma>'' s' t'
|
|
1232 |
assume "\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T1"
|
|
1233 |
then have "\<Gamma>'' \<turnstile> \<gamma> is \<theta> over \<Gamma>" using logical_subst_monotonicity h by blast
|
|
1234 |
then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext hl fs by blast
|
|
1235 |
then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><t2> is (x,t')#\<theta><t2> : T2" using ih by auto
|
|
1236 |
then have "\<Gamma>''\<turnstile>\<gamma><t2>[x::=s'] is \<theta><t2>[x::=t'] : T2" using psubst_subst_psubst fs by simp
|
|
1237 |
moreover have "App (Lam [x].\<gamma><t2>) s' \<leadsto> \<gamma><t2>[x::=s']" by auto
|
|
1238 |
moreover have "App (Lam [x].\<theta><t2>) t' \<leadsto> \<theta><t2>[x::=t']" by auto
|
|
1239 |
ultimately have "\<Gamma>''\<turnstile> App (Lam [x].\<gamma><t2>) s' is App (Lam [x].\<theta><t2>) t' : T2"
|
|
1240 |
using logical_weak_head_closure by auto
|
|
1241 |
}
|
|
1242 |
moreover have "valid \<Gamma>'" using h by auto
|
|
1243 |
ultimately show "\<Gamma>' \<turnstile> \<gamma><Lam [x].t2> is \<theta><Lam [x].t2> : T1\<rightarrow>T2" using fs by auto
|
|
1244 |
qed (auto)
|
|
1245 |
|
|
1246 |
theorem fundamental_theorem_2:
|
|
1247 |
assumes h1: "\<Gamma> \<turnstile> s == t : T"
|
|
1248 |
and h2: "\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>"
|
|
1249 |
shows "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T"
|
|
1250 |
using h1 h2
|
|
1251 |
proof (nominal_induct \<Gamma> s t T avoiding: \<Gamma>' \<gamma> \<theta> rule: equiv_def_strong_induct)
|
|
1252 |
case (Q_Refl \<Gamma> t T \<Gamma>' \<gamma> \<theta>)
|
|
1253 |
have "\<Gamma> \<turnstile> t : T" by fact
|
|
1254 |
moreover have "\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1255 |
ultimately show "\<Gamma>' \<turnstile> \<gamma><t> is \<theta><t> : T" using fundamental_theorem by blast
|
|
1256 |
next
|
|
1257 |
case (Q_Symm \<Gamma> t s T \<Gamma>' \<gamma> \<theta>)
|
|
1258 |
have "\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1259 |
moreover have ih:"\<And> \<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><t> is \<theta><s> : T" by fact
|
|
1260 |
ultimately show "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T" using logical_symmetry by blast
|
|
1261 |
next
|
|
1262 |
case (Q_Trans \<Gamma> s t T u \<Gamma>' \<gamma> \<theta>)
|
|
1263 |
have ih1:"\<And> \<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T" by fact
|
|
1264 |
have ih2:"\<And> \<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><t> is \<theta><u> : T" by fact
|
|
1265 |
have h:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1266 |
then have "\<Gamma>' \<turnstile> \<theta> is \<theta> over \<Gamma>" using logical_pseudo_reflexivity by auto
|
|
1267 |
then have "\<Gamma>' \<turnstile> \<theta><t> is \<theta><u> : T" using ih2 by auto
|
|
1268 |
moreover have "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T" using ih1 h by auto
|
|
1269 |
ultimately show "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><u> : T" using logical_transitivity by blast
|
|
1270 |
next
|
|
1271 |
case (Q_Abs x \<Gamma> T1 s2 t2 T2 \<Gamma>' \<gamma> \<theta>)
|
|
1272 |
have fs:"x\<sharp>\<Gamma>" by fact
|
|
1273 |
have fs2: "x\<sharp>\<gamma>" "x\<sharp>\<theta>" by fact
|
|
1274 |
have h1:"(x,T1)#\<Gamma> \<turnstile> s2 == t2 : T2" by fact
|
|
1275 |
have h2:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1276 |
have ih:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s2> is \<theta><t2> : T2" by fact
|
|
1277 |
{
|
|
1278 |
fix \<Gamma>'' s' t'
|
|
1279 |
assume "\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T1"
|
|
1280 |
then have "\<Gamma>'' \<turnstile> \<gamma> is \<theta> over \<Gamma>" using h2 logical_subst_monotonicity by blast
|
|
1281 |
then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext hl fs by blast
|
|
1282 |
then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><s2> is (x,t')#\<theta><t2> : T2" using ih by blast
|
|
1283 |
then have "\<Gamma>''\<turnstile> \<gamma><s2>[x::=s'] is \<theta><t2>[x::=t'] : T2" using fs2 psubst_subst_psubst by auto
|
|
1284 |
moreover have "App (Lam [x]. \<gamma><s2>) s' \<leadsto> \<gamma><s2>[x::=s']"
|
|
1285 |
and "App (Lam [x].\<theta><t2>) t' \<leadsto> \<theta><t2>[x::=t']" by auto
|
|
1286 |
ultimately have "\<Gamma>'' \<turnstile> App (Lam [x]. \<gamma><s2>) s' is App (Lam [x].\<theta><t2>) t' : T2"
|
|
1287 |
using logical_weak_head_closure by auto
|
|
1288 |
}
|
|
1289 |
moreover have "valid \<Gamma>'" using h2 by auto
|
|
1290 |
ultimately have "\<Gamma>' \<turnstile> Lam [x].\<gamma><s2> is Lam [x].\<theta><t2> : T1\<rightarrow>T2" by auto
|
|
1291 |
then show "\<Gamma>' \<turnstile> \<gamma><Lam [x].s2> is \<theta><Lam [x].t2> : T1\<rightarrow>T2" using fs2 by auto
|
|
1292 |
next
|
|
1293 |
case (Q_App \<Gamma> s1 t1 T1 T2 s2 t2 \<Gamma>' \<gamma> \<theta>)
|
|
1294 |
then show "\<Gamma>' \<turnstile> \<gamma><App s1 s2> is \<theta><App t1 t2> : T2" by auto
|
|
1295 |
next
|
|
1296 |
case (Q_Beta x \<Gamma> T1 s12 t12 T2 s2 t2 \<Gamma>' \<gamma> \<theta>)
|
|
1297 |
have h:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1298 |
have fs:"x\<sharp>\<Gamma>" by fact
|
|
1299 |
have fs2:"x\<sharp>\<gamma>" "x\<sharp>\<theta>" by fact
|
|
1300 |
have ih1:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s2> is \<theta><t2> : T1" by fact
|
|
1301 |
have ih2:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><s12> is \<theta><t12> : T2" by fact
|
|
1302 |
have "\<Gamma>' \<turnstile> \<gamma><s2> is \<theta><t2> : T1" using ih1 h by auto
|
|
1303 |
then have "\<Gamma>' \<turnstile> (x,\<gamma><s2>)#\<gamma> is (x,\<theta><t2>)#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext h fs by blast
|
|
1304 |
then have "\<Gamma>' \<turnstile> (x,\<gamma><s2>)#\<gamma><s12> is (x,\<theta><t2>)#\<theta><t12> : T2" using ih2 by auto
|
|
1305 |
then have "\<Gamma>' \<turnstile> \<gamma><s12>[x::=\<gamma><s2>] is \<theta><t12>[x::=\<theta><t2>] : T2" using fs2 psubst_subst_psubst by auto
|
|
1306 |
then have "\<Gamma>' \<turnstile> \<gamma><s12>[x::=\<gamma><s2>] is \<theta><t12[x::=t2]> : T2" using fs2 psubst_subst_propagate by auto
|
|
1307 |
moreover have "App (Lam [x].\<gamma><s12>) (\<gamma><s2>) \<leadsto> \<gamma><s12>[x::=\<gamma><s2>]" by auto
|
|
1308 |
ultimately have "\<Gamma>' \<turnstile> App (Lam [x].\<gamma><s12>) (\<gamma><s2>) is \<theta><t12[x::=t2]> : T2"
|
|
1309 |
using logical_weak_head_closure' by auto
|
|
1310 |
then show "\<Gamma>' \<turnstile> \<gamma><App (Lam [x].s12) s2> is \<theta><t12[x::=t2]> : T2" using fs2 by simp
|
|
1311 |
next
|
|
1312 |
case (Q_Ext x \<Gamma> s t T1 T2 \<Gamma>' \<gamma> \<theta>)
|
|
1313 |
have h2:"\<Gamma>' \<turnstile> \<gamma> is \<theta> over \<Gamma>" by fact
|
|
1314 |
have fs:"x\<sharp>\<Gamma>" "x\<sharp>s" "x\<sharp>t" by fact
|
|
1315 |
have ih:"\<And>\<Gamma>' \<gamma> \<theta>. \<Gamma>' \<turnstile> \<gamma> is \<theta> over (x,T1)#\<Gamma> \<Longrightarrow> \<Gamma>' \<turnstile> \<gamma><App s (Var x)> is \<theta><App t (Var x)> : T2"
|
|
1316 |
by fact
|
|
1317 |
{
|
|
1318 |
fix \<Gamma>'' s' t'
|
|
1319 |
assume hsub:"\<Gamma>'\<lless>\<Gamma>''" and hl:"\<Gamma>''\<turnstile> s' is t' : T1"
|
|
1320 |
then have "\<Gamma>'' \<turnstile> \<gamma> is \<theta> over \<Gamma>" using h2 logical_subst_monotonicity by blast
|
|
1321 |
then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma> is (x,t')#\<theta> over (x,T1)#\<Gamma>" using equiv_subst_ext hl fs by blast
|
|
1322 |
then have "\<Gamma>'' \<turnstile> (x,s')#\<gamma><App s (Var x)> is (x,t')#\<theta><App t (Var x)> : T2" using ih by blast
|
22082
|
1323 |
then have "\<Gamma>'' \<turnstile> App (((x,s')#\<gamma>)<s>) (((x,s')#\<gamma>)<(Var x)>) is App ((x,t')#\<theta><t>) ((x,t')#\<theta><(Var x)>) : T2"
|
|
1324 |
by auto
|
22073
|
1325 |
then have "\<Gamma>'' \<turnstile> App ((x,s')#\<gamma><s>) s' is App ((x,t')#\<theta><t>) t' : T2" by auto
|
|
1326 |
then have "\<Gamma>'' \<turnstile> App (\<gamma><s>) s' is App (\<theta><t>) t' : T2" using fs fresh_psubst_simpl by auto
|
|
1327 |
}
|
|
1328 |
moreover have "valid \<Gamma>'" using h2 by auto
|
|
1329 |
ultimately show "\<Gamma>' \<turnstile> \<gamma><s> is \<theta><t> : T1\<rightarrow>T2" by auto
|
|
1330 |
qed
|
|
1331 |
|
|
1332 |
theorem completeness:
|
|
1333 |
assumes "\<Gamma> \<turnstile> s == t : T"
|
|
1334 |
shows "\<Gamma> \<turnstile> s <=> t : T"
|
|
1335 |
using assms
|
|
1336 |
proof -
|
|
1337 |
{
|
|
1338 |
fix x T
|
|
1339 |
assume "(x,T) \<in> set \<Gamma>" "valid \<Gamma>"
|
|
1340 |
then have "\<Gamma> \<turnstile> Var x is Var x : T" using main_lemma by blast
|
|
1341 |
}
|
|
1342 |
moreover have "valid \<Gamma>" using equiv_def_valid assms by auto
|
|
1343 |
ultimately have "\<Gamma> \<turnstile> [] is [] over \<Gamma>" by auto
|
|
1344 |
then have "\<Gamma> \<turnstile> []<s> is []<t> : T" using fundamental_theorem_2 assms by blast
|
|
1345 |
then have "\<Gamma> \<turnstile> s is t : T" by simp
|
|
1346 |
then show "\<Gamma> \<turnstile> s <=> t : T" using main_lemma by simp
|
|
1347 |
qed
|
|
1348 |
|
22082
|
1349 |
(* Soundness is left as an exercise - like in the book - for the avid formalist
|
22073
|
1350 |
|
|
1351 |
theorem soundness:
|
22082
|
1352 |
shows "\<lbrakk>\<Gamma> \<turnstile> s <=> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s == t : T"
|
|
1353 |
and "\<lbrakk>\<Gamma> \<turnstile> s \<leftrightarrow> t : T; \<Gamma> \<turnstile> t : T; \<Gamma> \<turnstile> s : T\<rbrakk> \<Longrightarrow> \<Gamma> \<turnstile> s == t : T"
|
|
1354 |
|
22073
|
1355 |
*)
|
|
1356 |
|
|
1357 |
end
|
|
1358 |
|