| author | blanchet | 
| Thu, 13 Feb 2014 13:16:16 +0100 | |
| changeset 55451 | ea1d9408a233 | 
| parent 53015 | a1119cf551e8 | 
| child 57512 | cc97b347b301 | 
| permissions | -rw-r--r-- | 
| 27623 | 1  | 
(* Title: HOL/Nominal/Examples/Standardization.thy  | 
2  | 
Author: Stefan Berghofer and Tobias Nipkow  | 
|
3  | 
Copyright 2005, 2008 TU Muenchen  | 
|
4  | 
*)  | 
|
5  | 
||
6  | 
header {* Standardization *}
 | 
|
7  | 
||
8  | 
theory Standardization  | 
|
| 45966 | 9  | 
imports "../Nominal"  | 
| 27623 | 10  | 
begin  | 
11  | 
||
12  | 
text {*
 | 
|
13  | 
The proof of the standardization theorem, as well as most of the theorems about  | 
|
14  | 
lambda calculus in the following sections, are taken from @{text "HOL/Lambda"}.
 | 
|
15  | 
*}  | 
|
16  | 
||
17  | 
subsection {* Lambda terms *}
 | 
|
18  | 
||
19  | 
atom_decl name  | 
|
20  | 
||
21  | 
nominal_datatype lam =  | 
|
22  | 
Var "name"  | 
|
23  | 
| App "lam" "lam" (infixl "\<degree>" 200)  | 
|
24  | 
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [0, 10] 10)
 | 
|
25  | 
||
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
26  | 
instantiation lam :: size  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
27  | 
begin  | 
| 27623 | 28  | 
|
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
29  | 
nominal_primrec size_lam  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
30  | 
where  | 
| 27623 | 31  | 
"size (Var n) = 0"  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
32  | 
| "size (t \<degree> u) = size t + size u + 1"  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
33  | 
| "size (Lam [x].t) = size t + 1"  | 
| 27623 | 34  | 
apply finite_guess+  | 
35  | 
apply (rule TrueI)+  | 
|
36  | 
apply (simp add: fresh_nat)  | 
|
37  | 
apply fresh_guess+  | 
|
38  | 
done  | 
|
39  | 
||
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
40  | 
instance ..  | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
41  | 
|
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
42  | 
end  | 
| 27623 | 43  | 
|
44  | 
nominal_primrec  | 
|
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
45  | 
  subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam"  ("_[_::=_]" [300, 0, 0] 300)
 | 
| 
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
46  | 
where  | 
| 27623 | 47  | 
subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47397 
diff
changeset
 | 
48  | 
| subst_App: "(t\<^sub>1 \<degree> t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \<degree> t\<^sub>2[y::=s]"  | 
| 
29097
 
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
 
berghofe 
parents: 
27623 
diff
changeset
 | 
49  | 
| subst_Lam: "x \<sharp> (y, s) \<Longrightarrow> (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))"  | 
| 27623 | 50  | 
apply(finite_guess)+  | 
51  | 
apply(rule TrueI)+  | 
|
52  | 
apply(simp add: abs_fresh)  | 
|
53  | 
apply(fresh_guess)+  | 
|
54  | 
done  | 
|
55  | 
||
56  | 
lemma subst_eqvt [eqvt]:  | 
|
57  | 
"(pi::name prm) \<bullet> (t[x::=u]) = (pi \<bullet> t)[(pi \<bullet> x)::=(pi \<bullet> u)]"  | 
|
58  | 
by (nominal_induct t avoiding: x u rule: lam.strong_induct)  | 
|
59  | 
(perm_simp add: fresh_bij)+  | 
|
60  | 
||
61  | 
lemma subst_rename:  | 
|
62  | 
"y \<sharp> t \<Longrightarrow> ([(y, x)] \<bullet> t)[y::=u] = t[x::=u]"  | 
|
63  | 
by (nominal_induct t avoiding: x y u rule: lam.strong_induct)  | 
|
64  | 
(simp_all add: fresh_atm calc_atm abs_fresh)  | 
|
65  | 
||
66  | 
lemma fresh_subst:  | 
|
67  | 
"(x::name) \<sharp> t \<Longrightarrow> x \<sharp> u \<Longrightarrow> x \<sharp> t[y::=u]"  | 
|
68  | 
by (nominal_induct t avoiding: x y u rule: lam.strong_induct)  | 
|
69  | 
(auto simp add: abs_fresh fresh_atm)  | 
|
70  | 
||
71  | 
lemma fresh_subst':  | 
|
72  | 
"(x::name) \<sharp> u \<Longrightarrow> x \<sharp> t[x::=u]"  | 
|
73  | 
by (nominal_induct t avoiding: x u rule: lam.strong_induct)  | 
|
74  | 
(auto simp add: abs_fresh fresh_atm)  | 
|
75  | 
||
76  | 
lemma subst_forget: "(x::name) \<sharp> t \<Longrightarrow> t[x::=u] = t"  | 
|
77  | 
by (nominal_induct t avoiding: x u rule: lam.strong_induct)  | 
|
78  | 
(auto simp add: abs_fresh fresh_atm)  | 
|
79  | 
||
80  | 
lemma subst_subst:  | 
|
81  | 
"x \<noteq> y \<Longrightarrow> x \<sharp> v \<Longrightarrow> t[y::=v][x::=u[y::=v]] = t[x::=u][y::=v]"  | 
|
82  | 
by (nominal_induct t avoiding: x y u v rule: lam.strong_induct)  | 
|
83  | 
(auto simp add: fresh_subst subst_forget)  | 
|
84  | 
||
85  | 
declare subst_Var [simp del]  | 
|
86  | 
||
87  | 
lemma subst_eq [simp]: "(Var x)[x::=u] = u"  | 
|
88  | 
by (simp add: subst_Var)  | 
|
89  | 
||
90  | 
lemma subst_neq [simp]: "x \<noteq> y \<Longrightarrow> (Var x)[y::=u] = Var x"  | 
|
91  | 
by (simp add: subst_Var)  | 
|
92  | 
||
93  | 
inductive beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)  | 
|
94  | 
where  | 
|
95  | 
beta: "x \<sharp> t \<Longrightarrow> (Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]"  | 
|
96  | 
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"  | 
|
97  | 
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"  | 
|
98  | 
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> (Lam [x].s) \<rightarrow>\<^sub>\<beta> (Lam [x].t)"  | 
|
99  | 
||
100  | 
equivariance beta  | 
|
101  | 
nominal_inductive beta  | 
|
102  | 
by (simp_all add: abs_fresh fresh_subst')  | 
|
103  | 
||
104  | 
lemma better_beta [simp, intro!]: "(Lam [x].s) \<degree> t \<rightarrow>\<^sub>\<beta> s[x::=t]"  | 
|
105  | 
proof -  | 
|
106  | 
obtain y::name where y: "y \<sharp> (x, s, t)"  | 
|
107  | 
by (rule exists_fresh) (rule fin_supp)  | 
|
108  | 
then have "y \<sharp> t" by simp  | 
|
109  | 
then have "(Lam [y]. [(y, x)] \<bullet> s) \<degree> t \<rightarrow>\<^sub>\<beta> ([(y, x)] \<bullet> s)[y::=t]"  | 
|
110  | 
by (rule beta)  | 
|
111  | 
moreover from y have "(Lam [x].s) = (Lam [y]. [(y, x)] \<bullet> s)"  | 
|
112  | 
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)  | 
|
113  | 
ultimately show ?thesis using y by (simp add: subst_rename)  | 
|
114  | 
qed  | 
|
115  | 
||
116  | 
abbreviation  | 
|
117  | 
beta_reds :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50) where  | 
|
118  | 
"s \<rightarrow>\<^sub>\<beta>\<^sup>* t \<equiv> beta\<^sup>*\<^sup>* s t"  | 
|
119  | 
||
120  | 
||
121  | 
subsection {* Application of a term to a list of terms *}
 | 
|
122  | 
||
123  | 
abbreviation  | 
|
124  | 
list_application :: "lam \<Rightarrow> lam list \<Rightarrow> lam" (infixl "\<degree>\<degree>" 150) where  | 
|
125  | 
"t \<degree>\<degree> ts \<equiv> foldl (op \<degree>) t ts"  | 
|
126  | 
||
127  | 
lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)"  | 
|
128  | 
by (induct ts rule: rev_induct) (auto simp add: lam.inject)  | 
|
129  | 
||
130  | 
lemma Var_eq_apps_conv [iff]: "(Var m = s \<degree>\<degree> ss) = (Var m = s \<and> ss = [])"  | 
|
131  | 
by (induct ss arbitrary: s) auto  | 
|
132  | 
||
133  | 
lemma Var_apps_eq_Var_apps_conv [iff]:  | 
|
134  | 
"(Var m \<degree>\<degree> rs = Var n \<degree>\<degree> ss) = (m = n \<and> rs = ss)"  | 
|
135  | 
apply (induct rs arbitrary: ss rule: rev_induct)  | 
|
136  | 
apply (simp add: lam.inject)  | 
|
137  | 
apply blast  | 
|
138  | 
apply (induct_tac ss rule: rev_induct)  | 
|
139  | 
apply (auto simp add: lam.inject)  | 
|
140  | 
done  | 
|
141  | 
||
142  | 
lemma App_eq_foldl_conv:  | 
|
143  | 
"(r \<degree> s = t \<degree>\<degree> ts) =  | 
|
144  | 
(if ts = [] then r \<degree> s = t  | 
|
145  | 
else (\<exists>ss. ts = ss @ [s] \<and> r = t \<degree>\<degree> ss))"  | 
|
146  | 
apply (rule_tac xs = ts in rev_exhaust)  | 
|
147  | 
apply (auto simp add: lam.inject)  | 
|
148  | 
done  | 
|
149  | 
||
150  | 
lemma Abs_eq_apps_conv [iff]:  | 
|
151  | 
"((Lam [x].r) = s \<degree>\<degree> ss) = ((Lam [x].r) = s \<and> ss = [])"  | 
|
152  | 
by (induct ss rule: rev_induct) auto  | 
|
153  | 
||
154  | 
lemma apps_eq_Abs_conv [iff]: "(s \<degree>\<degree> ss = (Lam [x].r)) = (s = (Lam [x].r) \<and> ss = [])"  | 
|
155  | 
by (induct ss rule: rev_induct) auto  | 
|
156  | 
||
157  | 
lemma Abs_App_neq_Var_apps [iff]:  | 
|
158  | 
"(Lam [x].s) \<degree> t \<noteq> Var n \<degree>\<degree> ss"  | 
|
159  | 
by (induct ss arbitrary: s t rule: rev_induct) (auto simp add: lam.inject)  | 
|
160  | 
||
161  | 
lemma Var_apps_neq_Abs_apps [iff]:  | 
|
162  | 
"Var n \<degree>\<degree> ts \<noteq> (Lam [x].r) \<degree>\<degree> ss"  | 
|
163  | 
apply (induct ss arbitrary: ts rule: rev_induct)  | 
|
164  | 
apply simp  | 
|
165  | 
apply (induct_tac ts rule: rev_induct)  | 
|
166  | 
apply (auto simp add: lam.inject)  | 
|
167  | 
done  | 
|
168  | 
||
169  | 
lemma ex_head_tail:  | 
|
170  | 
"\<exists>ts h. t = h \<degree>\<degree> ts \<and> ((\<exists>n. h = Var n) \<or> (\<exists>x u. h = (Lam [x].u)))"  | 
|
171  | 
apply (induct t rule: lam.induct)  | 
|
| 39459 | 172  | 
apply (metis foldl_Nil)  | 
173  | 
apply (metis foldl_Cons foldl_Nil foldl_append)  | 
|
174  | 
apply (metis foldl_Nil)  | 
|
| 27623 | 175  | 
done  | 
176  | 
||
177  | 
lemma size_apps [simp]:  | 
|
178  | 
"size (r \<degree>\<degree> rs) = size r + foldl (op +) 0 (map size rs) + length rs"  | 
|
179  | 
by (induct rs rule: rev_induct) auto  | 
|
180  | 
||
181  | 
lemma lem0: "(0::nat) < k \<Longrightarrow> m \<le> n \<Longrightarrow> m < n + k"  | 
|
182  | 
by simp  | 
|
183  | 
||
184  | 
lemma subst_map [simp]:  | 
|
185  | 
"(t \<degree>\<degree> ts)[x::=u] = t[x::=u] \<degree>\<degree> map (\<lambda>t. t[x::=u]) ts"  | 
|
186  | 
by (induct ts arbitrary: t) simp_all  | 
|
187  | 
||
188  | 
lemma app_last: "(t \<degree>\<degree> ts) \<degree> u = t \<degree>\<degree> (ts @ [u])"  | 
|
189  | 
by simp  | 
|
190  | 
||
191  | 
lemma perm_apps [eqvt]:  | 
|
192  | 
"(pi::name prm) \<bullet> (t \<degree>\<degree> ts) = ((pi \<bullet> t) \<degree>\<degree> (pi \<bullet> ts))"  | 
|
193  | 
by (induct ts rule: rev_induct) (auto simp add: append_eqvt)  | 
|
194  | 
||
195  | 
lemma fresh_apps [simp]: "(x::name) \<sharp> (t \<degree>\<degree> ts) = (x \<sharp> t \<and> x \<sharp> ts)"  | 
|
196  | 
by (induct ts rule: rev_induct)  | 
|
197  | 
(auto simp add: fresh_list_append fresh_list_nil fresh_list_cons)  | 
|
198  | 
||
199  | 
text {* A customized induction schema for @{text "\<degree>\<degree>"}. *}
 | 
|
200  | 
||
201  | 
lemma lem:  | 
|
202  | 
assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)"  | 
|
203  | 
and "\<And>x u ts z. x \<sharp> z \<Longrightarrow> (\<And>z. P z u) \<Longrightarrow> (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z ((Lam [x].u) \<degree>\<degree> ts)"  | 
|
204  | 
shows "size t = n \<Longrightarrow> P z t"  | 
|
205  | 
apply (induct n arbitrary: t z rule: nat_less_induct)  | 
|
206  | 
apply (cut_tac t = t in ex_head_tail)  | 
|
207  | 
apply clarify  | 
|
208  | 
apply (erule disjE)  | 
|
209  | 
apply clarify  | 
|
210  | 
apply (rule assms)  | 
|
211  | 
apply clarify  | 
|
212  | 
apply (erule allE, erule impE)  | 
|
213  | 
prefer 2  | 
|
214  | 
apply (erule allE, erule impE, rule refl, erule spec)  | 
|
| 46129 | 215  | 
apply simp  | 
| 
47397
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
216  | 
apply (simp only: foldl_conv_fold add_commute fold_plus_listsum_rev)  | 
| 
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
217  | 
apply (fastforce simp add: listsum_map_remove1)  | 
| 27623 | 218  | 
apply clarify  | 
219  | 
apply (subgoal_tac "\<exists>y::name. y \<sharp> (x, u, z)")  | 
|
| 39459 | 220  | 
prefer 2  | 
221  | 
apply (blast intro: exists_fresh' fin_supp)  | 
|
| 27623 | 222  | 
apply (erule exE)  | 
223  | 
apply (subgoal_tac "(Lam [x].u) = (Lam [y].([(y, x)] \<bullet> u))")  | 
|
224  | 
prefer 2  | 
|
225  | 
apply (auto simp add: lam.inject alpha' fresh_prod fresh_atm)[]  | 
|
226  | 
apply (simp (no_asm_simp))  | 
|
227  | 
apply (rule assms)  | 
|
228  | 
apply (simp add: fresh_prod)  | 
|
229  | 
apply (erule allE, erule impE)  | 
|
230  | 
prefer 2  | 
|
231  | 
apply (erule allE, erule impE, rule refl, erule spec)  | 
|
232  | 
apply simp  | 
|
233  | 
apply clarify  | 
|
234  | 
apply (erule allE, erule impE)  | 
|
235  | 
prefer 2  | 
|
| 
47397
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
236  | 
apply blast  | 
| 
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
237  | 
apply simp  | 
| 
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
238  | 
apply (simp only: foldl_conv_fold add_commute fold_plus_listsum_rev)  | 
| 
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
239  | 
apply (fastforce simp add: listsum_map_remove1)  | 
| 27623 | 240  | 
done  | 
241  | 
||
242  | 
theorem Apps_lam_induct:  | 
|
243  | 
assumes "\<And>n ts (z::'a::fs_name). (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z (Var n \<degree>\<degree> ts)"  | 
|
244  | 
and "\<And>x u ts z. x \<sharp> z \<Longrightarrow> (\<And>z. P z u) \<Longrightarrow> (\<And>z. \<forall>t \<in> set ts. P z t) \<Longrightarrow> P z ((Lam [x].u) \<degree>\<degree> ts)"  | 
|
245  | 
shows "P z t"  | 
|
246  | 
apply (rule_tac t = t and z = z in lem)  | 
|
247  | 
prefer 3  | 
|
248  | 
apply (rule refl)  | 
|
249  | 
using assms apply blast+  | 
|
250  | 
done  | 
|
251  | 
||
252  | 
||
253  | 
subsection {* Congruence rules *}
 | 
|
254  | 
||
255  | 
lemma apps_preserves_beta [simp]:  | 
|
256  | 
"r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> r \<degree>\<degree> ss \<rightarrow>\<^sub>\<beta> s \<degree>\<degree> ss"  | 
|
257  | 
by (induct ss rule: rev_induct) auto  | 
|
258  | 
||
259  | 
lemma rtrancl_beta_Abs [intro!]:  | 
|
260  | 
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> (Lam [x].s) \<rightarrow>\<^sub>\<beta>\<^sup>* (Lam [x].s')"  | 
|
261  | 
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+  | 
|
262  | 
||
263  | 
lemma rtrancl_beta_AppL:  | 
|
264  | 
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t"  | 
|
265  | 
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+  | 
|
266  | 
||
267  | 
lemma rtrancl_beta_AppR:  | 
|
268  | 
"t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s \<degree> t'"  | 
|
269  | 
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+  | 
|
270  | 
||
271  | 
lemma rtrancl_beta_App [intro]:  | 
|
272  | 
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'"  | 
|
273  | 
by (blast intro!: rtrancl_beta_AppL rtrancl_beta_AppR intro: rtranclp_trans)  | 
|
274  | 
||
275  | 
||
276  | 
subsection {* Lifting an order to lists of elements *}
 | 
|
277  | 
||
278  | 
definition  | 
|
279  | 
  step1 :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
 | 
|
280  | 
"step1 r =  | 
|
281  | 
(\<lambda>ys xs. \<exists>us z z' vs. xs = us @ z # vs \<and> r z' z \<and> ys =  | 
|
282  | 
us @ z' # vs)"  | 
|
283  | 
||
284  | 
lemma not_Nil_step1 [iff]: "\<not> step1 r [] xs"  | 
|
285  | 
apply (unfold step1_def)  | 
|
286  | 
apply blast  | 
|
287  | 
done  | 
|
288  | 
||
289  | 
lemma not_step1_Nil [iff]: "\<not> step1 r xs []"  | 
|
290  | 
apply (unfold step1_def)  | 
|
291  | 
apply blast  | 
|
292  | 
done  | 
|
293  | 
||
294  | 
lemma Cons_step1_Cons [iff]:  | 
|
295  | 
"(step1 r (y # ys) (x # xs)) =  | 
|
296  | 
(r y x \<and> xs = ys \<or> x = y \<and> step1 r ys xs)"  | 
|
297  | 
apply (unfold step1_def)  | 
|
298  | 
apply (rule iffI)  | 
|
299  | 
apply (erule exE)  | 
|
300  | 
apply (rename_tac ts)  | 
|
301  | 
apply (case_tac ts)  | 
|
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
39779 
diff
changeset
 | 
302  | 
apply fastforce  | 
| 27623 | 303  | 
apply force  | 
304  | 
apply (erule disjE)  | 
|
305  | 
apply blast  | 
|
306  | 
apply (blast intro: Cons_eq_appendI)  | 
|
307  | 
done  | 
|
308  | 
||
309  | 
lemma append_step1I:  | 
|
310  | 
"step1 r ys xs \<and> vs = us \<or> ys = xs \<and> step1 r vs us  | 
|
311  | 
\<Longrightarrow> step1 r (ys @ vs) (xs @ us)"  | 
|
312  | 
apply (unfold step1_def)  | 
|
313  | 
apply auto  | 
|
314  | 
apply blast  | 
|
315  | 
apply (blast intro: append_eq_appendI)  | 
|
316  | 
done  | 
|
317  | 
||
318  | 
lemma Cons_step1E [elim!]:  | 
|
319  | 
assumes "step1 r ys (x # xs)"  | 
|
320  | 
and "\<And>y. ys = y # xs \<Longrightarrow> r y x \<Longrightarrow> R"  | 
|
321  | 
and "\<And>zs. ys = x # zs \<Longrightarrow> step1 r zs xs \<Longrightarrow> R"  | 
|
322  | 
shows R  | 
|
323  | 
using assms  | 
|
324  | 
apply (cases ys)  | 
|
325  | 
apply (simp add: step1_def)  | 
|
326  | 
apply blast  | 
|
327  | 
done  | 
|
328  | 
||
329  | 
lemma Snoc_step1_SnocD:  | 
|
330  | 
"step1 r (ys @ [y]) (xs @ [x])  | 
|
331  | 
\<Longrightarrow> (step1 r ys xs \<and> y = x \<or> ys = xs \<and> r y x)"  | 
|
332  | 
apply (unfold step1_def)  | 
|
333  | 
apply (clarify del: disjCI)  | 
|
334  | 
apply (rename_tac vs)  | 
|
335  | 
apply (rule_tac xs = vs in rev_exhaust)  | 
|
336  | 
apply force  | 
|
337  | 
apply simp  | 
|
338  | 
apply blast  | 
|
339  | 
done  | 
|
340  | 
||
341  | 
||
342  | 
subsection {* Lifting beta-reduction to lists *}
 | 
|
343  | 
||
344  | 
abbreviation  | 
|
345  | 
list_beta :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>\<beta>]\<^sub>1" 50) where  | 
|
346  | 
"rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<equiv> step1 beta rs ss"  | 
|
347  | 
||
348  | 
lemma head_Var_reduction:  | 
|
349  | 
"Var n \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> v \<Longrightarrow> \<exists>ss. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<and> v = Var n \<degree>\<degree> ss"  | 
|
350  | 
apply (induct u \<equiv> "Var n \<degree>\<degree> rs" v arbitrary: rs set: beta)  | 
|
351  | 
apply simp  | 
|
352  | 
apply (rule_tac xs = rs in rev_exhaust)  | 
|
353  | 
apply simp  | 
|
354  | 
apply (atomize, force intro: append_step1I iff: lam.inject)  | 
|
355  | 
apply (rule_tac xs = rs in rev_exhaust)  | 
|
356  | 
apply simp  | 
|
357  | 
apply (auto 0 3 intro: disjI2 [THEN append_step1I] simp add: lam.inject)  | 
|
358  | 
done  | 
|
359  | 
||
360  | 
lemma apps_betasE [case_names appL appR beta, consumes 1]:  | 
|
361  | 
assumes major: "r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> s"  | 
|
362  | 
and cases: "\<And>r'. r \<rightarrow>\<^sub>\<beta> r' \<Longrightarrow> s = r' \<degree>\<degree> rs \<Longrightarrow> R"  | 
|
363  | 
"\<And>rs'. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs' \<Longrightarrow> s = r \<degree>\<degree> rs' \<Longrightarrow> R"  | 
|
364  | 
"\<And>t u us. (x \<sharp> r \<Longrightarrow> r = (Lam [x].t) \<and> rs = u # us \<and> s = t[x::=u] \<degree>\<degree> us) \<Longrightarrow> R"  | 
|
365  | 
shows R  | 
|
366  | 
proof -  | 
|
367  | 
from major have  | 
|
368  | 
"(\<exists>r'. r \<rightarrow>\<^sub>\<beta> r' \<and> s = r' \<degree>\<degree> rs) \<or>  | 
|
369  | 
(\<exists>rs'. rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs' \<and> s = r \<degree>\<degree> rs') \<or>  | 
|
370  | 
(\<exists>t u us. x \<sharp> r \<longrightarrow> r = (Lam [x].t) \<and> rs = u # us \<and> s = t[x::=u] \<degree>\<degree> us)"  | 
|
371  | 
apply (nominal_induct u \<equiv> "r \<degree>\<degree> rs" s avoiding: x r rs rule: beta.strong_induct)  | 
|
372  | 
apply (simp add: App_eq_foldl_conv)  | 
|
373  | 
apply (split split_if_asm)  | 
|
374  | 
apply simp  | 
|
375  | 
apply blast  | 
|
376  | 
apply simp  | 
|
377  | 
apply (rule impI)+  | 
|
378  | 
apply (rule disjI2)  | 
|
379  | 
apply (rule disjI2)  | 
|
380  | 
apply (subgoal_tac "r = [(xa, x)] \<bullet> (Lam [x].s)")  | 
|
381  | 
prefer 2  | 
|
382  | 
apply (simp add: perm_fresh_fresh)  | 
|
383  | 
apply (drule conjunct1)  | 
|
384  | 
apply (subgoal_tac "r = (Lam [xa]. [(xa, x)] \<bullet> s)")  | 
|
385  | 
prefer 2  | 
|
386  | 
apply (simp add: calc_atm)  | 
|
387  | 
apply (thin_tac "r = ?t")  | 
|
388  | 
apply simp  | 
|
389  | 
apply (rule exI)  | 
|
390  | 
apply (rule conjI)  | 
|
391  | 
apply (rule refl)  | 
|
392  | 
apply (simp add: abs_fresh fresh_atm fresh_left calc_atm subst_rename)  | 
|
393  | 
apply (drule App_eq_foldl_conv [THEN iffD1])  | 
|
394  | 
apply (split split_if_asm)  | 
|
395  | 
apply simp  | 
|
396  | 
apply blast  | 
|
397  | 
apply (force intro!: disjI1 [THEN append_step1I] simp add: fresh_list_append)  | 
|
398  | 
apply (drule App_eq_foldl_conv [THEN iffD1])  | 
|
399  | 
apply (split split_if_asm)  | 
|
400  | 
apply simp  | 
|
401  | 
apply blast  | 
|
402  | 
apply (clarify, auto 0 3 intro!: exI intro: append_step1I)  | 
|
403  | 
done  | 
|
404  | 
with cases show ?thesis by blast  | 
|
405  | 
qed  | 
|
406  | 
||
407  | 
lemma apps_preserves_betas [simp]:  | 
|
408  | 
"rs [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<Longrightarrow> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> r \<degree>\<degree> ss"  | 
|
409  | 
apply (induct rs arbitrary: ss rule: rev_induct)  | 
|
410  | 
apply simp  | 
|
411  | 
apply simp  | 
|
412  | 
apply (rule_tac xs = ss in rev_exhaust)  | 
|
413  | 
apply simp  | 
|
414  | 
apply simp  | 
|
415  | 
apply (drule Snoc_step1_SnocD)  | 
|
416  | 
apply blast  | 
|
417  | 
done  | 
|
418  | 
||
419  | 
||
420  | 
subsection {* Standard reduction relation *}
 | 
|
421  | 
||
422  | 
text {*
 | 
|
423  | 
Based on lecture notes by Ralph Matthes,  | 
|
424  | 
original proof idea due to Ralph Loader.  | 
|
425  | 
*}  | 
|
426  | 
||
427  | 
declare listrel_mono [mono_set]  | 
|
428  | 
||
429  | 
lemma listrelp_eqvt [eqvt]:  | 
|
| 46317 | 430  | 
fixes f :: "'a::pt_name \<Rightarrow> 'b::pt_name \<Rightarrow> bool"  | 
| 27623 | 431  | 
assumes xy: "listrelp f (x::'a::pt_name list) y"  | 
432  | 
shows "listrelp ((pi::name prm) \<bullet> f) (pi \<bullet> x) (pi \<bullet> y)" using xy  | 
|
| 45966 | 433  | 
by induct (simp_all add: listrelp.intros perm_app [symmetric])  | 
| 27623 | 434  | 
|
435  | 
inductive  | 
|
436  | 
sred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>s" 50)  | 
|
437  | 
and sredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>s]" 50)  | 
|
438  | 
where  | 
|
439  | 
"s [\<rightarrow>\<^sub>s] t \<equiv> listrelp op \<rightarrow>\<^sub>s s t"  | 
|
440  | 
| Var: "rs [\<rightarrow>\<^sub>s] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> rs'"  | 
|
441  | 
| Abs: "x \<sharp> (ss, ss') \<Longrightarrow> r \<rightarrow>\<^sub>s r' \<Longrightarrow> ss [\<rightarrow>\<^sub>s] ss' \<Longrightarrow> (Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> ss'"  | 
|
442  | 
| Beta: "x \<sharp> (s, ss, t) \<Longrightarrow> r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t \<Longrightarrow> (Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t"  | 
|
443  | 
||
444  | 
equivariance sred  | 
|
445  | 
nominal_inductive sred  | 
|
446  | 
by (simp add: abs_fresh)+  | 
|
447  | 
||
448  | 
lemma better_sred_Abs:  | 
|
449  | 
assumes H1: "r \<rightarrow>\<^sub>s r'"  | 
|
450  | 
and H2: "ss [\<rightarrow>\<^sub>s] ss'"  | 
|
451  | 
shows "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> ss'"  | 
|
452  | 
proof -  | 
|
453  | 
obtain y::name where y: "y \<sharp> (x, r, r', ss, ss')"  | 
|
454  | 
by (rule exists_fresh) (rule fin_supp)  | 
|
455  | 
then have "y \<sharp> (ss, ss')" by simp  | 
|
456  | 
moreover from H1 have "[(y, x)] \<bullet> (r \<rightarrow>\<^sub>s r')" by (rule perm_boolI)  | 
|
457  | 
then have "([(y, x)] \<bullet> r) \<rightarrow>\<^sub>s ([(y, x)] \<bullet> r')" by (simp add: eqvts)  | 
|
458  | 
ultimately have "(Lam [y]. [(y, x)] \<bullet> r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [y]. [(y, x)] \<bullet> r') \<degree>\<degree> ss'" using H2  | 
|
459  | 
by (rule sred.Abs)  | 
|
460  | 
moreover from y have "(Lam [x].r) = (Lam [y]. [(y, x)] \<bullet> r)"  | 
|
461  | 
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)  | 
|
462  | 
moreover from y have "(Lam [x].r') = (Lam [y]. [(y, x)] \<bullet> r')"  | 
|
463  | 
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)  | 
|
464  | 
ultimately show ?thesis by simp  | 
|
465  | 
qed  | 
|
466  | 
||
467  | 
lemma better_sred_Beta:  | 
|
468  | 
assumes H: "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t"  | 
|
469  | 
shows "(Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t"  | 
|
470  | 
proof -  | 
|
471  | 
obtain y::name where y: "y \<sharp> (x, r, s, ss, t)"  | 
|
472  | 
by (rule exists_fresh) (rule fin_supp)  | 
|
473  | 
then have "y \<sharp> (s, ss, t)" by simp  | 
|
474  | 
moreover from y H have "([(y, x)] \<bullet> r)[y::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t"  | 
|
475  | 
by (simp add: subst_rename)  | 
|
476  | 
ultimately have "(Lam [y].[(y, x)] \<bullet> r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>s t"  | 
|
477  | 
by (rule sred.Beta)  | 
|
478  | 
moreover from y have "(Lam [x].r) = (Lam [y]. [(y, x)] \<bullet> r)"  | 
|
479  | 
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm)  | 
|
480  | 
ultimately show ?thesis by simp  | 
|
481  | 
qed  | 
|
482  | 
||
483  | 
lemmas better_sred_intros = sred.Var better_sred_Abs better_sred_Beta  | 
|
484  | 
||
485  | 
lemma refl_listrelp: "\<forall>x\<in>set xs. R x x \<Longrightarrow> listrelp R xs xs"  | 
|
486  | 
by (induct xs) (auto intro: listrelp.intros)  | 
|
487  | 
||
488  | 
lemma refl_sred: "t \<rightarrow>\<^sub>s t"  | 
|
489  | 
by (nominal_induct t rule: Apps_lam_induct) (auto intro: refl_listrelp better_sred_intros)  | 
|
490  | 
||
491  | 
lemma listrelp_conj1: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp R x y"  | 
|
492  | 
by (erule listrelp.induct) (auto intro: listrelp.intros)  | 
|
493  | 
||
494  | 
lemma listrelp_conj2: "listrelp (\<lambda>x y. R x y \<and> S x y) x y \<Longrightarrow> listrelp S x y"  | 
|
495  | 
by (erule listrelp.induct) (auto intro: listrelp.intros)  | 
|
496  | 
||
497  | 
lemma listrelp_app:  | 
|
498  | 
assumes xsys: "listrelp R xs ys"  | 
|
499  | 
shows "listrelp R xs' ys' \<Longrightarrow> listrelp R (xs @ xs') (ys @ ys')" using xsys  | 
|
500  | 
by (induct arbitrary: xs' ys') (auto intro: listrelp.intros)  | 
|
501  | 
||
502  | 
lemma lemma1:  | 
|
503  | 
assumes r: "r \<rightarrow>\<^sub>s r'" and s: "s \<rightarrow>\<^sub>s s'"  | 
|
504  | 
shows "r \<degree> s \<rightarrow>\<^sub>s r' \<degree> s'" using r  | 
|
505  | 
proof induct  | 
|
506  | 
case (Var rs rs' x)  | 
|
507  | 
then have "rs [\<rightarrow>\<^sub>s] rs'" by (rule listrelp_conj1)  | 
|
508  | 
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros)  | 
|
509  | 
ultimately have "rs @ [s] [\<rightarrow>\<^sub>s] rs' @ [s']" by (rule listrelp_app)  | 
|
510  | 
hence "Var x \<degree>\<degree> (rs @ [s]) \<rightarrow>\<^sub>s Var x \<degree>\<degree> (rs' @ [s'])" by (rule sred.Var)  | 
|
511  | 
thus ?case by (simp only: app_last)  | 
|
512  | 
next  | 
|
513  | 
case (Abs x ss ss' r r')  | 
|
514  | 
from Abs(4) have "ss [\<rightarrow>\<^sub>s] ss'" by (rule listrelp_conj1)  | 
|
515  | 
moreover have "[s] [\<rightarrow>\<^sub>s] [s']" by (iprover intro: s listrelp.intros)  | 
|
516  | 
ultimately have "ss @ [s] [\<rightarrow>\<^sub>s] ss' @ [s']" by (rule listrelp_app)  | 
|
517  | 
with `r \<rightarrow>\<^sub>s r'` have "(Lam [x].r) \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> (ss' @ [s'])"  | 
|
518  | 
by (rule better_sred_Abs)  | 
|
519  | 
thus ?case by (simp only: app_last)  | 
|
520  | 
next  | 
|
521  | 
case (Beta x u ss t r)  | 
|
522  | 
hence "r[x::=u] \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (simp only: app_last)  | 
|
523  | 
hence "(Lam [x].r) \<degree> u \<degree>\<degree> (ss @ [s]) \<rightarrow>\<^sub>s t \<degree> s'" by (rule better_sred_Beta)  | 
|
524  | 
thus ?case by (simp only: app_last)  | 
|
525  | 
qed  | 
|
526  | 
||
527  | 
lemma lemma1':  | 
|
528  | 
assumes ts: "ts [\<rightarrow>\<^sub>s] ts'"  | 
|
529  | 
shows "r \<rightarrow>\<^sub>s r' \<Longrightarrow> r \<degree>\<degree> ts \<rightarrow>\<^sub>s r' \<degree>\<degree> ts'" using ts  | 
|
530  | 
by (induct arbitrary: r r') (auto intro: lemma1)  | 
|
531  | 
||
532  | 
lemma listrelp_betas:  | 
|
533  | 
assumes ts: "listrelp op \<rightarrow>\<^sub>\<beta>\<^sup>* ts ts'"  | 
|
534  | 
shows "\<And>t t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<degree>\<degree> ts'" using ts  | 
|
535  | 
by induct auto  | 
|
536  | 
||
537  | 
lemma lemma2:  | 
|
538  | 
assumes t: "t \<rightarrow>\<^sub>s u"  | 
|
539  | 
shows "t \<rightarrow>\<^sub>\<beta>\<^sup>* u" using t  | 
|
540  | 
by induct (auto dest: listrelp_conj2  | 
|
541  | 
intro: listrelp_betas apps_preserves_beta converse_rtranclp_into_rtranclp)  | 
|
542  | 
||
543  | 
lemma lemma3:  | 
|
544  | 
assumes r: "r \<rightarrow>\<^sub>s r'"  | 
|
545  | 
shows "s \<rightarrow>\<^sub>s s' \<Longrightarrow> r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" using r  | 
|
546  | 
proof (nominal_induct avoiding: x s s' rule: sred.strong_induct)  | 
|
547  | 
case (Var rs rs' y)  | 
|
548  | 
hence "map (\<lambda>t. t[x::=s]) rs [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) rs'"  | 
|
549  | 
by induct (auto intro: listrelp.intros Var)  | 
|
550  | 
moreover have "Var y[x::=s] \<rightarrow>\<^sub>s Var y[x::=s']"  | 
|
551  | 
by (cases "y = x") (auto simp add: Var intro: refl_sred)  | 
|
552  | 
ultimately show ?case by simp (rule lemma1')  | 
|
553  | 
next  | 
|
554  | 
case (Abs y ss ss' r r')  | 
|
555  | 
then have "r[x::=s] \<rightarrow>\<^sub>s r'[x::=s']" by fast  | 
|
556  | 
moreover from Abs(8) `s \<rightarrow>\<^sub>s s'` have "map (\<lambda>t. t[x::=s]) ss [\<rightarrow>\<^sub>s] map (\<lambda>t. t[x::=s']) ss'"  | 
|
557  | 
by induct (auto intro: listrelp.intros Abs)  | 
|
558  | 
ultimately show ?case using Abs(6) `y \<sharp> x` `y \<sharp> s` `y \<sharp> s'`  | 
|
559  | 
by simp (rule better_sred_Abs)  | 
|
560  | 
next  | 
|
561  | 
case (Beta y u ss t r)  | 
|
562  | 
thus ?case by (auto simp add: subst_subst fresh_atm intro: better_sred_Beta)  | 
|
563  | 
qed  | 
|
564  | 
||
565  | 
lemma lemma4_aux:  | 
|
566  | 
assumes rs: "listrelp (\<lambda>t u. t \<rightarrow>\<^sub>s u \<and> (\<forall>r. u \<rightarrow>\<^sub>\<beta> r \<longrightarrow> t \<rightarrow>\<^sub>s r)) rs rs'"  | 
|
567  | 
shows "rs' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss \<Longrightarrow> rs [\<rightarrow>\<^sub>s] ss" using rs  | 
|
568  | 
proof (induct arbitrary: ss)  | 
|
569  | 
case Nil  | 
|
570  | 
thus ?case by cases (auto intro: listrelp.Nil)  | 
|
571  | 
next  | 
|
572  | 
case (Cons x y xs ys)  | 
|
573  | 
note Cons' = Cons  | 
|
574  | 
show ?case  | 
|
575  | 
proof (cases ss)  | 
|
576  | 
case Nil with Cons show ?thesis by simp  | 
|
577  | 
next  | 
|
578  | 
case (Cons y' ys')  | 
|
579  | 
hence ss: "ss = y' # ys'" by simp  | 
|
580  | 
from Cons Cons' have "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys \<or> y' = y \<and> ys [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ys'" by simp  | 
|
581  | 
hence "x # xs [\<rightarrow>\<^sub>s] y' # ys'"  | 
|
582  | 
proof  | 
|
583  | 
assume H: "y \<rightarrow>\<^sub>\<beta> y' \<and> ys' = ys"  | 
|
584  | 
with Cons' have "x \<rightarrow>\<^sub>s y'" by blast  | 
|
585  | 
moreover from Cons' have "xs [\<rightarrow>\<^sub>s] ys" by (iprover dest: listrelp_conj1)  | 
|
586  | 
ultimately have "x # xs [\<rightarrow>\<^sub>s] y' # ys" by (rule listrelp.Cons)  | 
|
587  | 
with H show ?thesis by simp  | 
|
588  | 
next  | 
|
589  | 
assume H: "y' = y \<and> ys [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ys'"  | 
|
590  | 
with Cons' have "x \<rightarrow>\<^sub>s y'" by blast  | 
|
591  | 
moreover from H have "xs [\<rightarrow>\<^sub>s] ys'" by (blast intro: Cons')  | 
|
592  | 
ultimately show ?thesis by (rule listrelp.Cons)  | 
|
593  | 
qed  | 
|
594  | 
with ss show ?thesis by simp  | 
|
595  | 
qed  | 
|
596  | 
qed  | 
|
597  | 
||
598  | 
lemma lemma4:  | 
|
599  | 
assumes r: "r \<rightarrow>\<^sub>s r'"  | 
|
600  | 
shows "r' \<rightarrow>\<^sub>\<beta> r'' \<Longrightarrow> r \<rightarrow>\<^sub>s r''" using r  | 
|
601  | 
proof (nominal_induct avoiding: r'' rule: sred.strong_induct)  | 
|
602  | 
case (Var rs rs' x)  | 
|
603  | 
then obtain ss where rs: "rs' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ss" and r'': "r'' = Var x \<degree>\<degree> ss"  | 
|
604  | 
by (blast dest: head_Var_reduction)  | 
|
605  | 
from Var(1) [simplified] rs have "rs [\<rightarrow>\<^sub>s] ss" by (rule lemma4_aux)  | 
|
606  | 
hence "Var x \<degree>\<degree> rs \<rightarrow>\<^sub>s Var x \<degree>\<degree> ss" by (rule sred.Var)  | 
|
607  | 
with r'' show ?case by simp  | 
|
608  | 
next  | 
|
609  | 
case (Abs x ss ss' r r')  | 
|
610  | 
from `(Lam [x].r') \<degree>\<degree> ss' \<rightarrow>\<^sub>\<beta> r''` show ?case  | 
|
611  | 
proof (cases rule: apps_betasE [where x=x])  | 
|
612  | 
case (appL s)  | 
|
613  | 
then obtain r''' where s: "s = (Lam [x].r''')" and r''': "r' \<rightarrow>\<^sub>\<beta> r'''" using `x \<sharp> r''`  | 
|
614  | 
by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)  | 
|
615  | 
from r''' have "r \<rightarrow>\<^sub>s r'''" by (blast intro: Abs)  | 
|
616  | 
moreover from Abs have "ss [\<rightarrow>\<^sub>s] ss'" by (iprover dest: listrelp_conj1)  | 
|
617  | 
ultimately have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r''') \<degree>\<degree> ss'" by (rule better_sred_Abs)  | 
|
618  | 
with appL s show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp  | 
|
619  | 
next  | 
|
620  | 
case (appR rs')  | 
|
621  | 
from Abs(6) [simplified] `ss' [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs'`  | 
|
622  | 
have "ss [\<rightarrow>\<^sub>s] rs'" by (rule lemma4_aux)  | 
|
623  | 
with `r \<rightarrow>\<^sub>s r'` have "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> rs'" by (rule better_sred_Abs)  | 
|
624  | 
with appR show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by simp  | 
|
625  | 
next  | 
|
626  | 
case (beta t u' us')  | 
|
627  | 
then have Lam_eq: "(Lam [x].r') = (Lam [x].t)" and ss': "ss' = u' # us'"  | 
|
628  | 
and r'': "r'' = t[x::=u'] \<degree>\<degree> us'"  | 
|
629  | 
by (simp_all add: abs_fresh)  | 
|
630  | 
from Abs(6) ss' obtain u us where  | 
|
631  | 
ss: "ss = u # us" and u: "u \<rightarrow>\<^sub>s u'" and us: "us [\<rightarrow>\<^sub>s] us'"  | 
|
632  | 
by cases (auto dest!: listrelp_conj1)  | 
|
633  | 
have "r[x::=u] \<rightarrow>\<^sub>s r'[x::=u']" using `r \<rightarrow>\<^sub>s r'` and u by (rule lemma3)  | 
|
634  | 
with us have "r[x::=u] \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule lemma1')  | 
|
635  | 
hence "(Lam [x].r) \<degree> u \<degree>\<degree> us \<rightarrow>\<^sub>s r'[x::=u'] \<degree>\<degree> us'" by (rule better_sred_Beta)  | 
|
636  | 
with ss r'' Lam_eq show "(Lam [x].r) \<degree>\<degree> ss \<rightarrow>\<^sub>s r''" by (simp add: lam.inject alpha)  | 
|
637  | 
qed  | 
|
638  | 
next  | 
|
639  | 
case (Beta x s ss t r)  | 
|
640  | 
show ?case  | 
|
641  | 
by (rule better_sred_Beta) (rule Beta)+  | 
|
642  | 
qed  | 
|
643  | 
||
644  | 
lemma rtrancl_beta_sred:  | 
|
645  | 
assumes r: "r \<rightarrow>\<^sub>\<beta>\<^sup>* r'"  | 
|
646  | 
shows "r \<rightarrow>\<^sub>s r'" using r  | 
|
647  | 
by induct (iprover intro: refl_sred lemma4)+  | 
|
648  | 
||
649  | 
||
650  | 
subsection {* Terms in normal form *}
 | 
|
651  | 
||
652  | 
lemma listsp_eqvt [eqvt]:  | 
|
653  | 
assumes xs: "listsp p (xs::'a::pt_name list)"  | 
|
654  | 
shows "listsp ((pi::name prm) \<bullet> p) (pi \<bullet> xs)" using xs  | 
|
655  | 
apply induct  | 
|
656  | 
apply simp  | 
|
657  | 
apply simp  | 
|
658  | 
apply (rule listsp.intros)  | 
|
659  | 
apply (drule_tac pi=pi in perm_boolI)  | 
|
660  | 
apply perm_simp  | 
|
661  | 
apply assumption  | 
|
662  | 
done  | 
|
663  | 
||
664  | 
inductive NF :: "lam \<Rightarrow> bool"  | 
|
665  | 
where  | 
|
666  | 
App: "listsp NF ts \<Longrightarrow> NF (Var x \<degree>\<degree> ts)"  | 
|
667  | 
| Abs: "NF t \<Longrightarrow> NF (Lam [x].t)"  | 
|
668  | 
||
669  | 
equivariance NF  | 
|
670  | 
nominal_inductive NF  | 
|
671  | 
by (simp add: abs_fresh)  | 
|
672  | 
||
673  | 
lemma Abs_NF:  | 
|
674  | 
assumes NF: "NF ((Lam [x].t) \<degree>\<degree> ts)"  | 
|
675  | 
shows "ts = []" using NF  | 
|
676  | 
proof cases  | 
|
677  | 
case (App us i)  | 
|
678  | 
thus ?thesis by (simp add: Var_apps_neq_Abs_apps [THEN not_sym])  | 
|
679  | 
next  | 
|
680  | 
case (Abs u)  | 
|
681  | 
thus ?thesis by simp  | 
|
682  | 
qed  | 
|
683  | 
||
684  | 
text {*
 | 
|
685  | 
@{term NF} characterizes exactly the terms that are in normal form.
 | 
|
686  | 
*}  | 
|
687  | 
||
688  | 
lemma NF_eq: "NF t = (\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t')"  | 
|
689  | 
proof  | 
|
690  | 
assume H: "NF t"  | 
|
691  | 
show "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'"  | 
|
692  | 
proof  | 
|
693  | 
fix t'  | 
|
694  | 
from H show "\<not> t \<rightarrow>\<^sub>\<beta> t'"  | 
|
695  | 
proof (nominal_induct avoiding: t' rule: NF.strong_induct)  | 
|
696  | 
case (App ts t)  | 
|
697  | 
show ?case  | 
|
698  | 
proof  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
699  | 
assume "Var t \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> t'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
700  | 
then obtain rs where "ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 rs"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
701  | 
by (iprover dest: head_Var_reduction)  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
702  | 
with App show False  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
703  | 
by (induct rs arbitrary: ts) (auto del: in_listspD)  | 
| 27623 | 704  | 
qed  | 
705  | 
next  | 
|
706  | 
case (Abs t x)  | 
|
707  | 
show ?case  | 
|
708  | 
proof  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
709  | 
assume "(Lam [x].t) \<rightarrow>\<^sub>\<beta> t'"  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
710  | 
then show False using Abs  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
711  | 
by (cases rule: beta.strong_cases) (auto simp add: abs_fresh lam.inject alpha)  | 
| 27623 | 712  | 
qed  | 
713  | 
qed  | 
|
714  | 
qed  | 
|
715  | 
next  | 
|
716  | 
assume H: "\<forall>t'. \<not> t \<rightarrow>\<^sub>\<beta> t'"  | 
|
717  | 
then show "NF t"  | 
|
718  | 
proof (nominal_induct t rule: Apps_lam_induct)  | 
|
719  | 
case (1 n ts)  | 
|
720  | 
then have "\<forall>ts'. \<not> ts [\<rightarrow>\<^sub>\<beta>]\<^sub>1 ts'"  | 
|
721  | 
by (iprover intro: apps_preserves_betas)  | 
|
722  | 
with 1(1) have "listsp NF ts"  | 
|
723  | 
by (induct ts) (auto simp add: in_listsp_conv_set)  | 
|
724  | 
then show ?case by (rule NF.App)  | 
|
725  | 
next  | 
|
726  | 
case (2 x u ts)  | 
|
727  | 
show ?case  | 
|
728  | 
proof (cases ts)  | 
|
| 39459 | 729  | 
case Nil thus ?thesis by (metis 2 NF.Abs abs foldl_Nil)  | 
| 27623 | 730  | 
next  | 
731  | 
case (Cons r rs)  | 
|
732  | 
have "(Lam [x].u) \<degree> r \<rightarrow>\<^sub>\<beta> u[x::=r]" ..  | 
|
733  | 
then have "(Lam [x].u) \<degree> r \<degree>\<degree> rs \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs"  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
734  | 
by (rule apps_preserves_beta)  | 
| 27623 | 735  | 
with Cons have "(Lam [x].u) \<degree>\<degree> ts \<rightarrow>\<^sub>\<beta> u[x::=r] \<degree>\<degree> rs"  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
29097 
diff
changeset
 | 
736  | 
by simp  | 
| 27623 | 737  | 
with 2 show ?thesis by iprover  | 
738  | 
qed  | 
|
739  | 
qed  | 
|
740  | 
qed  | 
|
741  | 
||
742  | 
||
743  | 
subsection {* Leftmost reduction and weakly normalizing terms *}
 | 
|
744  | 
||
745  | 
inductive  | 
|
746  | 
lred :: "lam \<Rightarrow> lam \<Rightarrow> bool" (infixl "\<rightarrow>\<^sub>l" 50)  | 
|
747  | 
and lredlist :: "lam list \<Rightarrow> lam list \<Rightarrow> bool" (infixl "[\<rightarrow>\<^sub>l]" 50)  | 
|
748  | 
where  | 
|
749  | 
"s [\<rightarrow>\<^sub>l] t \<equiv> listrelp op \<rightarrow>\<^sub>l s t"  | 
|
750  | 
| Var: "rs [\<rightarrow>\<^sub>l] rs' \<Longrightarrow> Var x \<degree>\<degree> rs \<rightarrow>\<^sub>l Var x \<degree>\<degree> rs'"  | 
|
751  | 
| Abs: "r \<rightarrow>\<^sub>l r' \<Longrightarrow> (Lam [x].r) \<rightarrow>\<^sub>l (Lam [x].r')"  | 
|
752  | 
| Beta: "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>l t \<Longrightarrow> (Lam [x].r) \<degree> s \<degree>\<degree> ss \<rightarrow>\<^sub>l t"  | 
|
753  | 
||
754  | 
lemma lred_imp_sred:  | 
|
755  | 
assumes lred: "s \<rightarrow>\<^sub>l t"  | 
|
756  | 
shows "s \<rightarrow>\<^sub>s t" using lred  | 
|
757  | 
proof induct  | 
|
758  | 
case (Var rs rs' x)  | 
|
759  | 
then have "rs [\<rightarrow>\<^sub>s] rs'"  | 
|
760  | 
by induct (iprover intro: listrelp.intros)+  | 
|
761  | 
then show ?case by (rule sred.Var)  | 
|
762  | 
next  | 
|
763  | 
case (Abs r r' x)  | 
|
764  | 
from `r \<rightarrow>\<^sub>s r'`  | 
|
765  | 
have "(Lam [x].r) \<degree>\<degree> [] \<rightarrow>\<^sub>s (Lam [x].r') \<degree>\<degree> []" using listrelp.Nil  | 
|
766  | 
by (rule better_sred_Abs)  | 
|
767  | 
then show ?case by simp  | 
|
768  | 
next  | 
|
769  | 
case (Beta r x s ss t)  | 
|
770  | 
from `r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>s t`  | 
|
771  | 
show ?case by (rule better_sred_Beta)  | 
|
772  | 
qed  | 
|
773  | 
||
774  | 
inductive WN :: "lam \<Rightarrow> bool"  | 
|
775  | 
where  | 
|
776  | 
Var: "listsp WN rs \<Longrightarrow> WN (Var n \<degree>\<degree> rs)"  | 
|
777  | 
| Lambda: "WN r \<Longrightarrow> WN (Lam [x].r)"  | 
|
778  | 
| Beta: "WN ((r[x::=s]) \<degree>\<degree> ss) \<Longrightarrow> WN (((Lam [x].r) \<degree> s) \<degree>\<degree> ss)"  | 
|
779  | 
||
780  | 
lemma listrelp_imp_listsp1:  | 
|
781  | 
assumes H: "listrelp (\<lambda>x y. P x) xs ys"  | 
|
782  | 
shows "listsp P xs" using H  | 
|
783  | 
by induct auto  | 
|
784  | 
||
785  | 
lemma listrelp_imp_listsp2:  | 
|
786  | 
assumes H: "listrelp (\<lambda>x y. P y) xs ys"  | 
|
787  | 
shows "listsp P ys" using H  | 
|
788  | 
by induct auto  | 
|
789  | 
||
790  | 
lemma lemma5:  | 
|
791  | 
assumes lred: "r \<rightarrow>\<^sub>l r'"  | 
|
792  | 
shows "WN r" and "NF r'" using lred  | 
|
793  | 
by induct  | 
|
794  | 
(iprover dest: listrelp_conj1 listrelp_conj2  | 
|
795  | 
listrelp_imp_listsp1 listrelp_imp_listsp2 intro: WN.intros  | 
|
796  | 
NF.intros)+  | 
|
797  | 
||
798  | 
lemma lemma6:  | 
|
799  | 
assumes wn: "WN r"  | 
|
800  | 
shows "\<exists>r'. r \<rightarrow>\<^sub>l r'" using wn  | 
|
801  | 
proof induct  | 
|
802  | 
case (Var rs n)  | 
|
803  | 
then have "\<exists>rs'. rs [\<rightarrow>\<^sub>l] rs'"  | 
|
804  | 
by induct (iprover intro: listrelp.intros)+  | 
|
805  | 
then show ?case by (iprover intro: lred.Var)  | 
|
806  | 
qed (iprover intro: lred.intros)+  | 
|
807  | 
||
808  | 
lemma lemma7:  | 
|
809  | 
assumes r: "r \<rightarrow>\<^sub>s r'"  | 
|
810  | 
shows "NF r' \<Longrightarrow> r \<rightarrow>\<^sub>l r'" using r  | 
|
811  | 
proof induct  | 
|
812  | 
case (Var rs rs' x)  | 
|
813  | 
from `NF (Var x \<degree>\<degree> rs')` have "listsp NF rs'"  | 
|
814  | 
by cases simp_all  | 
|
815  | 
with Var(1) have "rs [\<rightarrow>\<^sub>l] rs'"  | 
|
816  | 
proof induct  | 
|
817  | 
case Nil  | 
|
818  | 
show ?case by (rule listrelp.Nil)  | 
|
819  | 
next  | 
|
| 39459 | 820  | 
case (Cons x y xs ys)  | 
| 27623 | 821  | 
hence "x \<rightarrow>\<^sub>l y" and "xs [\<rightarrow>\<^sub>l] ys" by (auto del: in_listspD)  | 
822  | 
thus ?case by (rule listrelp.Cons)  | 
|
823  | 
qed  | 
|
824  | 
thus ?case by (rule lred.Var)  | 
|
825  | 
next  | 
|
826  | 
case (Abs x ss ss' r r')  | 
|
827  | 
from `NF ((Lam [x].r') \<degree>\<degree> ss')`  | 
|
828  | 
have ss': "ss' = []" by (rule Abs_NF)  | 
|
829  | 
from Abs(4) have ss: "ss = []" using ss'  | 
|
830  | 
by cases simp_all  | 
|
831  | 
from ss' Abs have "NF (Lam [x].r')" by simp  | 
|
832  | 
hence "NF r'" by (cases rule: NF.strong_cases) (auto simp add: abs_fresh lam.inject alpha)  | 
|
833  | 
with Abs have "r \<rightarrow>\<^sub>l r'" by simp  | 
|
834  | 
hence "(Lam [x].r) \<rightarrow>\<^sub>l (Lam [x].r')" by (rule lred.Abs)  | 
|
835  | 
with ss ss' show ?case by simp  | 
|
836  | 
next  | 
|
837  | 
case (Beta x s ss t r)  | 
|
838  | 
hence "r[x::=s] \<degree>\<degree> ss \<rightarrow>\<^sub>l t" by simp  | 
|
839  | 
thus ?case by (rule lred.Beta)  | 
|
840  | 
qed  | 
|
841  | 
||
842  | 
lemma WN_eq: "WN t = (\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t')"  | 
|
843  | 
proof  | 
|
844  | 
assume "WN t"  | 
|
845  | 
then have "\<exists>t'. t \<rightarrow>\<^sub>l t'" by (rule lemma6)  | 
|
846  | 
then obtain t' where t': "t \<rightarrow>\<^sub>l t'" ..  | 
|
847  | 
then have NF: "NF t'" by (rule lemma5)  | 
|
848  | 
from t' have "t \<rightarrow>\<^sub>s t'" by (rule lred_imp_sred)  | 
|
849  | 
then have "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" by (rule lemma2)  | 
|
850  | 
with NF show "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'" by iprover  | 
|
851  | 
next  | 
|
852  | 
assume "\<exists>t'. t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<and> NF t'"  | 
|
853  | 
then obtain t' where t': "t \<rightarrow>\<^sub>\<beta>\<^sup>* t'" and NF: "NF t'"  | 
|
854  | 
by iprover  | 
|
855  | 
from t' have "t \<rightarrow>\<^sub>s t'" by (rule rtrancl_beta_sred)  | 
|
856  | 
then have "t \<rightarrow>\<^sub>l t'" using NF by (rule lemma7)  | 
|
857  | 
then show "WN t" by (rule lemma5)  | 
|
858  | 
qed  | 
|
859  | 
||
860  | 
end  | 
|
| 
47397
 
d654c73e4b12
no preference wrt. fold(l/r); prefer fold rather than foldr for iterating over lists in generated code
 
haftmann 
parents: 
46317 
diff
changeset
 | 
861  |