|
1 (* CR_Takahashi from Nominal1 ported to Nominal2 *) |
|
2 theory CR imports Nominal2 begin |
|
3 |
|
4 atom_decl name |
|
5 |
|
6 nominal_datatype lam = |
|
7 Var "name" |
|
8 | App "lam" "lam" |
|
9 | Lam x::"name" l::"lam" binds x in l ("Lam [_]. _" [100, 100] 100) |
|
10 |
|
11 nominal_primrec |
|
12 subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_ [_ ::= _]" [90, 90, 90] 90) |
|
13 where |
|
14 "(Var x)[y ::= s] = (if x = y then s else (Var x))" |
|
15 | "(App t1 t2)[y ::= s] = App (t1[y ::= s]) (t2[y ::= s])" |
|
16 | "atom x \<sharp> (y, s) \<Longrightarrow> (Lam [x]. t)[y ::= s] = Lam [x].(t[y ::= s])" |
|
17 unfolding eqvt_def subst_graph_def |
|
18 apply (rule, perm_simp, rule) |
|
19 apply(rule TrueI) |
|
20 apply(auto simp add: lam.distinct lam.eq_iff) |
|
21 apply(rule_tac y="a" and c="(aa, b)" in lam.strong_exhaust) |
|
22 apply(blast)+ |
|
23 apply(simp_all add: fresh_star_def fresh_Pair_elim) |
|
24 apply (erule_tac c="(ya,sa)" in Abs_lst1_fcb2) |
|
25 apply(simp_all add: Abs_fresh_iff) |
|
26 apply(simp add: fresh_star_def fresh_Pair) |
|
27 apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) |
|
28 apply(simp add: eqvt_at_def atom_eqvt fresh_star_Pair perm_supp_eq) |
|
29 done |
|
30 |
|
31 termination (eqvt) |
|
32 by lexicographic_order |
|
33 |
|
34 lemma forget: |
|
35 shows "atom x \<sharp> t \<Longrightarrow> t[x ::= s] = t" |
|
36 by (nominal_induct t avoiding: x s rule: lam.strong_induct) |
|
37 (auto simp add: lam.fresh fresh_at_base) |
|
38 |
|
39 lemma fresh_fact: |
|
40 fixes z::"name" |
|
41 assumes a: "atom z \<sharp> s" |
|
42 and b: "z = y \<or> atom z \<sharp> t" |
|
43 shows "atom z \<sharp> t[y ::= s]" |
|
44 using a b |
|
45 by (nominal_induct t avoiding: z y s rule: lam.strong_induct) |
|
46 (auto simp add: lam.fresh fresh_at_base) |
|
47 |
|
48 lemma substitution_lemma: |
|
49 assumes a: "x \<noteq> y" "atom x \<sharp> u" |
|
50 shows "t[x ::= s][y ::= u] = t[y ::= u][x ::= s[y ::= u]]" |
|
51 using a |
|
52 by (nominal_induct t avoiding: x y s u rule: lam.strong_induct) |
|
53 (auto simp add: fresh_fact forget) |
|
54 |
|
55 lemma subst_rename: |
|
56 assumes a: "atom y \<sharp> t" |
|
57 shows "t[x ::= s] = ((y \<leftrightarrow> x) \<bullet>t)[y ::= s]" |
|
58 using a |
|
59 by (nominal_induct t avoiding: x y s rule: lam.strong_induct) |
|
60 (auto simp add: lam.fresh fresh_at_base) |
|
61 |
|
62 lemma supp_subst: |
|
63 shows "supp (t[x ::= s]) \<subseteq> (supp t - {atom x}) \<union> supp s" |
|
64 by (induct t x s rule: subst.induct) (auto simp add: lam.supp supp_at_base) |
|
65 |
|
66 inductive |
|
67 beta :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>b _" [80,80] 80) |
|
68 where |
|
69 b1[intro]: "t1 \<longrightarrow>b t2 \<Longrightarrow> App t1 s \<longrightarrow>b App t2 s" |
|
70 | b2[intro]: "s1 \<longrightarrow>b s2 \<Longrightarrow> App t s1 \<longrightarrow>b App t s2" |
|
71 | b3[intro]: "t1 \<longrightarrow>b t2 \<Longrightarrow> Lam [x]. t1 \<longrightarrow>b Lam [x]. t2" |
|
72 | b4[intro]: "atom x \<sharp> s \<Longrightarrow> App (Lam [x]. t) s \<longrightarrow>b t[x ::= s]" |
|
73 |
|
74 equivariance beta |
|
75 |
|
76 nominal_inductive beta |
|
77 avoids b3: x |
|
78 | b4: x |
|
79 by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) |
|
80 |
|
81 section {* Transitive Closure of Beta *} |
|
82 |
|
83 inductive |
|
84 beta_star :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>b* _" [80,80] 80) |
|
85 where |
|
86 bs1[intro, simp]: "M \<longrightarrow>b* M" |
|
87 | bs2[intro]: "\<lbrakk>M1\<longrightarrow>b* M2; M2 \<longrightarrow>b M3\<rbrakk> \<Longrightarrow> M1 \<longrightarrow>b* M3" |
|
88 |
|
89 lemma bs3[intro, trans]: |
|
90 assumes "A \<longrightarrow>b* B" |
|
91 and "B \<longrightarrow>b* C" |
|
92 shows "A \<longrightarrow>b* C" |
|
93 using assms(2) assms(1) |
|
94 by induct auto |
|
95 |
|
96 section {* One-Reduction *} |
|
97 |
|
98 inductive |
|
99 One :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>1 _" [80,80] 80) |
|
100 where |
|
101 o1[intro]: "Var x \<longrightarrow>1 Var x" |
|
102 | o2[intro]: "\<lbrakk>t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>1 App t2 s2" |
|
103 | o3[intro]: "t1 \<longrightarrow>1 t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>1 Lam [x].t2" |
|
104 | o4[intro]: "\<lbrakk>atom x \<sharp> (s1, s2); t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>1 t2[x ::= s2]" |
|
105 |
|
106 equivariance One |
|
107 |
|
108 nominal_inductive One |
|
109 avoids o3: "x" |
|
110 | o4: "x" |
|
111 by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) |
|
112 |
|
113 lemma One_refl: |
|
114 shows "t \<longrightarrow>1 t" |
|
115 by (nominal_induct t rule: lam.strong_induct) (auto) |
|
116 |
|
117 lemma One_subst: |
|
118 assumes a: "t1 \<longrightarrow>1 t2" "s1 \<longrightarrow>1 s2" |
|
119 shows "t1[x ::= s1] \<longrightarrow>1 t2[x ::= s2]" |
|
120 using a |
|
121 by (nominal_induct t1 t2 avoiding: s1 s2 x rule: One.strong_induct) |
|
122 (auto simp add: substitution_lemma fresh_at_base fresh_fact fresh_Pair) |
|
123 |
|
124 lemma better_o4_intro: |
|
125 assumes a: "t1 \<longrightarrow>1 t2" "s1 \<longrightarrow>1 s2" |
|
126 shows "App (Lam [x]. t1) s1 \<longrightarrow>1 t2[ x ::= s2]" |
|
127 proof - |
|
128 obtain y::"name" where fs: "atom y \<sharp> (x, t1, s1, t2, s2)" by (rule obtain_fresh) |
|
129 have "App (Lam [x]. t1) s1 = App (Lam [y]. ((y \<leftrightarrow> x) \<bullet> t1)) s1" using fs |
|
130 by (auto simp add: lam.eq_iff Abs1_eq_iff' flip_def fresh_Pair fresh_at_base) |
|
131 also have "\<dots> \<longrightarrow>1 ((y \<leftrightarrow> x) \<bullet> t2)[y ::= s2]" using fs a by (auto simp add: One.eqvt) |
|
132 also have "\<dots> = t2[x ::= s2]" using fs by (simp add: subst_rename[symmetric]) |
|
133 finally show "App (Lam [x].t1) s1 \<longrightarrow>1 t2[x ::= s2]" by simp |
|
134 qed |
|
135 |
|
136 lemma One_Var: |
|
137 assumes a: "Var x \<longrightarrow>1 M" |
|
138 shows "M = Var x" |
|
139 using a by (cases rule: One.cases) (simp_all) |
|
140 |
|
141 lemma One_Lam: |
|
142 assumes a: "Lam [x].t \<longrightarrow>1 s" "atom x \<sharp> s" |
|
143 shows "\<exists>t'. s = Lam [x].t' \<and> t \<longrightarrow>1 t'" |
|
144 using a |
|
145 apply (cases rule: One.cases) |
|
146 apply (auto simp add: Abs1_eq_iff) |
|
147 apply (rule_tac x="(atom xa \<rightleftharpoons> atom x) \<bullet> t2" in exI) |
|
148 apply (auto simp add: fresh_permute_left lam.fresh) |
|
149 by (metis swap_commute One.eqvt) |
|
150 |
|
151 lemma One_App: |
|
152 assumes a: "App t s \<longrightarrow>1 r" |
|
153 shows "(\<exists>t' s'. r = App t' s' \<and> t \<longrightarrow>1 t' \<and> s \<longrightarrow>1 s') \<or> |
|
154 (\<exists>x p p' s'. r = p'[x ::= s'] \<and> t = Lam [x].p \<and> p \<longrightarrow>1 p' \<and> s \<longrightarrow>1 s' \<and> atom x \<sharp> (s,s'))" |
|
155 using a by (cases rule: One.cases) auto |
|
156 |
|
157 lemma One_preserves_fresh: |
|
158 fixes x::"name" |
|
159 assumes a: "M \<longrightarrow>1 N" |
|
160 shows "atom x \<sharp> M \<Longrightarrow> atom x \<sharp> N" |
|
161 using a |
|
162 by (induct, auto simp add: lam.fresh) |
|
163 (metis fresh_fact)+ |
|
164 |
|
165 (* TODO *) |
|
166 lemma One_strong_cases[consumes 1]: |
|
167 "\<lbrakk> a1 \<longrightarrow>1 a2; \<And>x. \<lbrakk>a1 = Var x; a2 = Var x\<rbrakk> \<Longrightarrow> P; |
|
168 \<And>t1 t2 s1 s2. \<lbrakk>a1 = App t1 s1; a2 = App t2 s2; t1 \<longrightarrow>1 t2; s1 \<longrightarrow>1 s2\<rbrakk> \<Longrightarrow> P; |
|
169 \<And>t1 t2. (\<lbrakk>atom xa \<sharp> a1; atom xa \<sharp> a2\<rbrakk> \<Longrightarrow> a1 = Lam [xa].t1 \<and> a2 = Lam [xa].t2 \<and> t1 \<longrightarrow>1 t2) \<Longrightarrow> P; |
|
170 \<And>s1 s2 t1 t2. |
|
171 (\<lbrakk>atom xaa \<sharp> a1; atom xaa \<sharp> a2\<rbrakk> |
|
172 \<Longrightarrow> a1 = App (Lam [xaa].t1) s1 \<and> a2 = t2[xaa::=s2] \<and> atom xaa \<sharp> (s1, s2) \<and> t1 \<longrightarrow>1 t2 \<and> s1 \<longrightarrow>1 s2) \<Longrightarrow> |
|
173 P\<rbrakk> |
|
174 \<Longrightarrow> P" |
|
175 apply (nominal_induct avoiding: a1 a2 rule: One.strong_induct) |
|
176 apply blast |
|
177 apply blast |
|
178 apply (simp add: fresh_Pair_elim Abs1_eq_iff lam.fresh) |
|
179 apply (case_tac "xa = x") |
|
180 apply (simp_all)[2] |
|
181 apply blast |
|
182 apply (rotate_tac 6) |
|
183 apply (drule_tac x="(atom x \<rightleftharpoons> atom xa) \<bullet> t1" in meta_spec) |
|
184 apply (rotate_tac -1) |
|
185 apply (drule_tac x="(atom x \<rightleftharpoons> atom xa) \<bullet> t2" in meta_spec) |
|
186 apply (simp add: One.eqvt fresh_permute_left) |
|
187 apply (simp add: fresh_Pair_elim Abs1_eq_iff lam.fresh) |
|
188 apply (case_tac "xaa = x") |
|
189 apply (simp_all add: fresh_Pair)[2] |
|
190 apply blast |
|
191 apply (rotate_tac -2) |
|
192 apply (drule_tac x="s1" in meta_spec) |
|
193 apply (rotate_tac -1) |
|
194 apply (drule_tac x="s2" in meta_spec) |
|
195 apply (rotate_tac -1) |
|
196 apply (drule_tac x="(atom x \<rightleftharpoons> atom xaa) \<bullet> t1" in meta_spec) |
|
197 apply (rotate_tac -1) |
|
198 apply (drule_tac x="(atom x \<rightleftharpoons> atom xaa) \<bullet> t2" in meta_spec) |
|
199 apply (rotate_tac -1) |
|
200 apply (simp add: One_preserves_fresh fresh_permute_left One.eqvt) |
|
201 by (metis Nominal2_Base.swap_commute One_preserves_fresh flip_def subst_rename) |
|
202 |
|
203 lemma One_Redex: |
|
204 assumes a: "App (Lam [x].t) s \<longrightarrow>1 r" "atom x \<sharp> (s,r)" |
|
205 shows "(\<exists>t' s'. r = App (Lam [x].t') s' \<and> t \<longrightarrow>1 t' \<and> s \<longrightarrow>1 s') \<or> |
|
206 (\<exists>t' s'. r = t'[x ::= s'] \<and> t \<longrightarrow>1 t' \<and> s \<longrightarrow>1 s')" |
|
207 using a |
|
208 by (cases rule: One_strong_cases) |
|
209 (auto dest!: One_Lam simp add: fresh_Pair lam.fresh Abs1_eq_iff) |
|
210 |
|
211 inductive |
|
212 One_star :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>1* _" [80,80] 80) |
|
213 where |
|
214 os1[intro]: "t \<longrightarrow>1* t" |
|
215 | os2[intro]: "t \<longrightarrow>1 s \<Longrightarrow> t \<longrightarrow>1* s" |
|
216 | os3[intro, trans]: "t1 \<longrightarrow>1* t2 \<Longrightarrow> t2 \<longrightarrow>1* t3 \<Longrightarrow> t1 \<longrightarrow>1* t3" |
|
217 |
|
218 section {* Complete Development Reduction *} |
|
219 |
|
220 inductive |
|
221 Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>d _" [80,80] 80) |
|
222 where |
|
223 d1[intro]: "Var x \<longrightarrow>d Var x" |
|
224 | d2[intro]: "t \<longrightarrow>d s \<Longrightarrow> Lam [x].t \<longrightarrow>d Lam[x].s" |
|
225 | d3[intro]: "\<lbrakk>\<not>(\<exists>y t'. t1 = Lam [y].t'); t1 \<longrightarrow>d t2; s1 \<longrightarrow>d s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>d App t2 s2" |
|
226 | d4[intro]: "\<lbrakk>atom x \<sharp> (s1,s2); t1 \<longrightarrow>d t2; s1 \<longrightarrow>d s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>d t2[x::=s2]" |
|
227 |
|
228 equivariance Dev |
|
229 nominal_inductive Dev |
|
230 avoids d2: "x" |
|
231 | d4: "x" |
|
232 by (simp_all add: fresh_star_def fresh_Pair lam.fresh fresh_fact) |
|
233 |
|
234 lemma better_d4_intro: |
|
235 assumes a: "t1 \<longrightarrow>d t2" "s1 \<longrightarrow>d s2" |
|
236 shows "App (Lam [x].t1) s1 \<longrightarrow>d t2[x::=s2]" |
|
237 proof - |
|
238 obtain y::"name" where fs: "atom y\<sharp>(x,t1,s1,t2,s2)" by (rule obtain_fresh) |
|
239 have "App (Lam [x].t1) s1 = App (Lam [y].((y \<leftrightarrow> x)\<bullet>t1)) s1" using fs |
|
240 by (auto simp add: Abs1_eq_iff' flip_def fresh_Pair fresh_at_base) |
|
241 also have "\<dots> \<longrightarrow>d ((y \<leftrightarrow> x) \<bullet> t2)[y ::= s2]" using fs a by (auto simp add: Dev.eqvt) |
|
242 also have "\<dots> = t2[x::=s2]" using fs by (simp add: subst_rename[symmetric]) |
|
243 finally show "App (Lam [x].t1) s1 \<longrightarrow>d t2[x::=s2]" by simp |
|
244 qed |
|
245 |
|
246 lemma Dev_preserves_fresh: |
|
247 fixes x::"name" |
|
248 assumes a: "M\<longrightarrow>d N" |
|
249 shows "atom x\<sharp>M \<Longrightarrow> atom x\<sharp>N" |
|
250 using a |
|
251 by (induct, auto simp add: lam.fresh) |
|
252 (metis fresh_fact)+ |
|
253 |
|
254 lemma Dev_Lam: |
|
255 assumes a: "Lam [x].M \<longrightarrow>d N" |
|
256 shows "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>d N'" |
|
257 proof - |
|
258 from a have "atom x \<sharp> Lam [x].M" by (simp add: lam.fresh) |
|
259 with a have "atom x \<sharp> N" by (simp add: Dev_preserves_fresh) |
|
260 with a show "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>d N'" |
|
261 apply (cases rule: Dev.cases) |
|
262 apply (auto simp add: Abs1_eq_iff lam.fresh) |
|
263 apply (rule_tac x="(atom xa \<rightleftharpoons> atom x) \<bullet> s" in exI) |
|
264 apply (auto simp add: fresh_permute_left lam.fresh) |
|
265 by (metis swap_commute Dev.eqvt) |
|
266 qed |
|
267 |
|
268 lemma Development_existence: |
|
269 shows "\<exists>M'. M \<longrightarrow>d M'" |
|
270 by (nominal_induct M rule: lam.strong_induct) |
|
271 (auto dest!: Dev_Lam intro: better_d4_intro) |
|
272 |
|
273 lemma Triangle: |
|
274 assumes a: "t \<longrightarrow>d t1" "t \<longrightarrow>1 t2" |
|
275 shows "t2 \<longrightarrow>1 t1" |
|
276 using a |
|
277 proof(nominal_induct avoiding: t2 rule: Dev.strong_induct) |
|
278 case (d4 x s1 s2 t1 t1' t2) |
|
279 have fc: "atom x\<sharp>t2" "atom x\<sharp>s1" by fact+ |
|
280 have "App (Lam [x].t1) s1 \<longrightarrow>1 t2" by fact |
|
281 then obtain t' s' where reds: |
|
282 "(t2 = App (Lam [x].t') s' \<and> t1 \<longrightarrow>1 t' \<and> s1 \<longrightarrow>1 s') \<or> |
|
283 (t2 = t'[x::=s'] \<and> t1 \<longrightarrow>1 t' \<and> s1 \<longrightarrow>1 s')" |
|
284 using fc by (auto dest!: One_Redex) |
|
285 have ih1: "t1 \<longrightarrow>1 t' \<Longrightarrow> t' \<longrightarrow>1 t1'" by fact |
|
286 have ih2: "s1 \<longrightarrow>1 s' \<Longrightarrow> s' \<longrightarrow>1 s2" by fact |
|
287 { assume "t1 \<longrightarrow>1 t'" "s1 \<longrightarrow>1 s'" |
|
288 then have "App (Lam [x].t') s' \<longrightarrow>1 t1'[x::=s2]" |
|
289 using ih1 ih2 by (auto intro: better_o4_intro) |
|
290 } |
|
291 moreover |
|
292 { assume "t1 \<longrightarrow>1 t'" "s1 \<longrightarrow>1 s'" |
|
293 then have "t'[x::=s'] \<longrightarrow>1 t1'[x::=s2]" |
|
294 using ih1 ih2 by (auto intro: One_subst) |
|
295 } |
|
296 ultimately show "t2 \<longrightarrow>1 t1'[x::=s2]" using reds by auto |
|
297 qed (auto dest!: One_Lam One_Var One_App) |
|
298 |
|
299 lemma Diamond_for_One: |
|
300 assumes a: "t \<longrightarrow>1 t1" "t \<longrightarrow>1 t2" |
|
301 shows "\<exists>t3. t2 \<longrightarrow>1 t3 \<and> t1 \<longrightarrow>1 t3" |
|
302 proof - |
|
303 obtain tc where "t \<longrightarrow>d tc" using Development_existence by blast |
|
304 with a have "t2 \<longrightarrow>1 tc" and "t1 \<longrightarrow>1 tc" by (simp_all add: Triangle) |
|
305 then show "\<exists>t3. t2 \<longrightarrow>1 t3 \<and> t1 \<longrightarrow>1 t3" by blast |
|
306 qed |
|
307 |
|
308 lemma Rectangle_for_One: |
|
309 assumes a: "t \<longrightarrow>1* t1" "t \<longrightarrow>1 t2" |
|
310 shows "\<exists>t3. t1 \<longrightarrow>1 t3 \<and> t2 \<longrightarrow>1* t3" |
|
311 using a Diamond_for_One by (induct arbitrary: t2) (blast)+ |
|
312 |
|
313 lemma CR_for_One_star: |
|
314 assumes a: "t \<longrightarrow>1* t1" "t \<longrightarrow>1* t2" |
|
315 shows "\<exists>t3. t2 \<longrightarrow>1* t3 \<and> t1 \<longrightarrow>1* t3" |
|
316 using a Rectangle_for_One by (induct arbitrary: t2) (blast)+ |
|
317 |
|
318 section {* Establishing the Equivalence of Beta-star and One-star *} |
|
319 |
|
320 lemma Beta_Lam_cong: |
|
321 assumes a: "t1 \<longrightarrow>b* t2" |
|
322 shows "Lam [x].t1 \<longrightarrow>b* Lam [x].t2" |
|
323 using a by (induct) (blast)+ |
|
324 |
|
325 lemma Beta_App_cong_aux: |
|
326 assumes a: "t1 \<longrightarrow>b* t2" |
|
327 shows "App t1 s\<longrightarrow>b* App t2 s" |
|
328 and "App s t1 \<longrightarrow>b* App s t2" |
|
329 using a by (induct) (blast)+ |
|
330 |
|
331 lemma Beta_App_cong: |
|
332 assumes a: "t1 \<longrightarrow>b* t2" "s1 \<longrightarrow>b* s2" |
|
333 shows "App t1 s1 \<longrightarrow>b* App t2 s2" |
|
334 using a by (blast intro: Beta_App_cong_aux) |
|
335 |
|
336 lemmas Beta_congs = Beta_Lam_cong Beta_App_cong |
|
337 |
|
338 lemma One_implies_Beta_star: |
|
339 assumes a: "t \<longrightarrow>1 s" |
|
340 shows "t \<longrightarrow>b* s" |
|
341 using a by (induct, auto intro!: Beta_congs) |
|
342 (metis (hide_lams, no_types) Beta_App_cong_aux(1) Beta_App_cong_aux(2) Beta_Lam_cong b4 bs2 bs3 fresh_PairD(2)) |
|
343 |
|
344 lemma One_congs: |
|
345 assumes a: "t1 \<longrightarrow>1* t2" |
|
346 shows "Lam [x].t1 \<longrightarrow>1* Lam [x].t2" |
|
347 and "App t1 s \<longrightarrow>1* App t2 s" |
|
348 and "App s t1 \<longrightarrow>1* App s t2" |
|
349 using a by (induct) (auto intro: One_refl) |
|
350 |
|
351 lemma Beta_implies_One_star: |
|
352 assumes a: "t1 \<longrightarrow>b t2" |
|
353 shows "t1 \<longrightarrow>1* t2" |
|
354 using a by (induct) (auto intro: One_refl One_congs better_o4_intro) |
|
355 |
|
356 lemma Beta_star_equals_One_star: |
|
357 shows "t1 \<longrightarrow>1* t2 = t1 \<longrightarrow>b* t2" |
|
358 proof |
|
359 assume "t1 \<longrightarrow>1* t2" |
|
360 then show "t1 \<longrightarrow>b* t2" by (induct) (auto intro: One_implies_Beta_star) |
|
361 next |
|
362 assume "t1 \<longrightarrow>b* t2" |
|
363 then show "t1 \<longrightarrow>1* t2" by (induct) (auto intro: Beta_implies_One_star) |
|
364 qed |
|
365 |
|
366 section {* The Church-Rosser Theorem *} |
|
367 |
|
368 theorem CR_for_Beta_star: |
|
369 assumes a: "t \<longrightarrow>b* t1" "t\<longrightarrow>b* t2" |
|
370 shows "\<exists>t3. t1 \<longrightarrow>b* t3 \<and> t2 \<longrightarrow>b* t3" |
|
371 proof - |
|
372 from a have "t \<longrightarrow>1* t1" and "t\<longrightarrow>1* t2" by (simp_all add: Beta_star_equals_One_star) |
|
373 then have "\<exists>t3. t1 \<longrightarrow>1* t3 \<and> t2 \<longrightarrow>1* t3" by (simp add: CR_for_One_star) |
|
374 then show "\<exists>t3. t1 \<longrightarrow>b* t3 \<and> t2 \<longrightarrow>b* t3" by (simp add: Beta_star_equals_One_star) |
|
375 qed |
|
376 |
|
377 end |