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