author | haftmann |
Sun, 06 Sep 2015 22:14:52 +0200 | |
changeset 61128 | 8e5072cba671 |
parent 53015 | a1119cf551e8 |
child 63167 | 0909deb8059b |
permissions | -rw-r--r-- |
23159 | 1 |
(* Authors: Christian Urban and Mathilde Arnaud *) |
2 |
(* *) |
|
3 |
(* A formalisation of the Church-Rosser proof by Masako Takahashi.*) |
|
4 |
(* This formalisation follows with some very slight exceptions *) |
|
5 |
(* the version of this proof given by Randy Pollack in the paper: *) |
|
6 |
(* *) |
|
40945 | 7 |
(* Polishing Up the Tait-Martin Löf Proof of the *) |
23159 | 8 |
(* Church-Rosser Theorem (1995). *) |
9 |
||
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
10 |
theory CR_Takahashi |
25867 | 11 |
imports "../Nominal" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
12 |
begin |
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
13 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
14 |
atom_decl name |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
15 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
16 |
nominal_datatype lam = |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
17 |
Var "name" |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
18 |
| App "lam" "lam" |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
19 |
| Lam "\<guillemotleft>name\<guillemotright>lam" ("Lam [_]._" [100,100] 100) |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
20 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
21 |
nominal_primrec |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
26966
diff
changeset
|
22 |
subst :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 100) |
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
26966
diff
changeset
|
23 |
where |
23159 | 24 |
"(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:
40945
diff
changeset
|
25 |
| "(App t\<^sub>1 t\<^sub>2)[y::=s] = App (t\<^sub>1[y::=s]) (t\<^sub>2[y::=s])" |
29097
68245155eb58
Modified nominal_primrec to make it work with local theories, unified syntax
berghofe
parents:
26966
diff
changeset
|
26 |
| "x\<sharp>(y,s) \<Longrightarrow> (Lam [x].t)[y::=s] = Lam [x].(t[y::=s])" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
27 |
apply(finite_guess)+ |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
28 |
apply(rule TrueI)+ |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
29 |
apply(simp add: abs_fresh) |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
30 |
apply(fresh_guess)+ |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
31 |
done |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
32 |
|
22829
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
33 |
section {* Lemmas about Capture-Avoiding Substitution *} |
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
34 |
|
23159 | 35 |
lemma subst_eqvt[eqvt]: |
36 |
fixes pi::"name prm" |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
37 |
shows "pi\<bullet>(t1[x::=t2]) = (pi\<bullet>t1)[(pi\<bullet>x)::=(pi\<bullet>t2)]" |
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
38 |
by (nominal_induct t1 avoiding: x t2 rule: lam.strong_induct) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
39 |
(auto simp add: perm_bij fresh_atm fresh_bij) |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
40 |
|
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
41 |
lemma forget: |
23159 | 42 |
shows "x\<sharp>t \<Longrightarrow> t[x::=s] = t" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
43 |
by (nominal_induct t avoiding: x s rule: lam.strong_induct) |
23159 | 44 |
(auto simp add: abs_fresh fresh_atm) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
45 |
|
26458 | 46 |
lemma fresh_fact: |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
47 |
fixes z::"name" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
48 |
shows "\<lbrakk>z\<sharp>s; (z=y \<or> z\<sharp>t)\<rbrakk> \<Longrightarrow> z\<sharp>t[y::=s]" |
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
49 |
by (nominal_induct t avoiding: z y s rule: lam.strong_induct) |
23159 | 50 |
(auto simp add: abs_fresh fresh_prod fresh_atm) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
51 |
|
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
52 |
lemma substitution_lemma: |
23159 | 53 |
assumes a: "x\<noteq>y" "x\<sharp>u" |
54 |
shows "t[x::=s][y::=u] = t[y::=u][x::=s[y::=u]]" |
|
29103 | 55 |
using a |
56 |
by (nominal_induct t avoiding: x y s u rule: lam.strong_induct) |
|
57 |
(auto simp add: fresh_fact forget) |
|
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
58 |
|
22829
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
59 |
lemma subst_rename: |
23159 | 60 |
assumes a: "y\<sharp>t" |
61 |
shows "t[x::=s] = ([(y,x)]\<bullet>t)[y::=s]" |
|
29103 | 62 |
using a |
63 |
by (nominal_induct t avoiding: x y s rule: lam.strong_induct) |
|
64 |
(auto simp add: swap_simps fresh_atm abs_fresh) |
|
22829
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
65 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
66 |
section {* Beta-Reduction *} |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
67 |
|
23760 | 68 |
inductive |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
69 |
"Beta" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta> _" [80,80] 80) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
70 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
71 |
b1[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> App t1 s \<longrightarrow>\<^sub>\<beta> App t2 s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
72 |
| b2[intro]: "s1 \<longrightarrow>\<^sub>\<beta> s2 \<Longrightarrow> App t s1 \<longrightarrow>\<^sub>\<beta> App t s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
73 |
| b3[intro]: "t1 \<longrightarrow>\<^sub>\<beta> t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^sub>\<beta> Lam [x].t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
74 |
| b4[intro]: "App (Lam [x].t) s \<longrightarrow>\<^sub>\<beta> t[x::=s]" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
75 |
|
22829
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
76 |
section {* Transitive Closure of Beta *} |
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
77 |
|
23760 | 78 |
inductive |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
79 |
"Beta_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>\<beta>\<^sup>* _" [80,80] 80) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
80 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
81 |
bs1[intro]: "t \<longrightarrow>\<^sub>\<beta>\<^sup>* t" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
82 |
| bs2[intro]: "t \<longrightarrow>\<^sub>\<beta> s \<Longrightarrow> t \<longrightarrow>\<^sub>\<beta>\<^sup>* s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
83 |
| bs3[intro,trans]: "\<lbrakk>t1\<longrightarrow>\<^sub>\<beta>\<^sup>* t2; t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
84 |
|
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
85 |
section {* One-Reduction *} |
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
86 |
|
23760 | 87 |
inductive |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
88 |
One :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1 _" [80,80] 80) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
89 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
90 |
o1[intro]: "Var x\<longrightarrow>\<^sub>1 Var x" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
91 |
| o2[intro]: "\<lbrakk>t1\<longrightarrow>\<^sub>1t2; s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^sub>1 App t2 s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
92 |
| o3[intro]: "t1\<longrightarrow>\<^sub>1t2 \<Longrightarrow> Lam [x].t1 \<longrightarrow>\<^sub>1 Lam [x].t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
93 |
| o4[intro]: "\<lbrakk>x\<sharp>(s1,s2); t1\<longrightarrow>\<^sub>1t2; s1\<longrightarrow>\<^sub>1s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2[x::=s2]" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
94 |
|
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
95 |
equivariance One |
23159 | 96 |
nominal_inductive One |
26458 | 97 |
by (simp_all add: abs_fresh fresh_fact) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
98 |
|
22855 | 99 |
lemma One_refl: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
100 |
shows "t \<longrightarrow>\<^sub>1 t" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
101 |
by (nominal_induct t rule: lam.strong_induct) (auto) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
102 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
103 |
lemma One_subst: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
104 |
assumes a: "t1 \<longrightarrow>\<^sub>1 t2" "s1 \<longrightarrow>\<^sub>1 s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
105 |
shows "t1[x::=s1] \<longrightarrow>\<^sub>1 t2[x::=s2]" |
29103 | 106 |
using a |
107 |
by (nominal_induct t1 t2 avoiding: s1 s2 x rule: One.strong_induct) |
|
108 |
(auto simp add: substitution_lemma fresh_atm fresh_fact) |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
109 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
110 |
lemma better_o4_intro: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
111 |
assumes a: "t1 \<longrightarrow>\<^sub>1 t2" "s1 \<longrightarrow>\<^sub>1 s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
112 |
shows "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2[x::=s2]" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
113 |
proof - |
23159 | 114 |
obtain y::"name" where fs: "y\<sharp>(x,t1,s1,t2,s2)" by (rule exists_fresh, rule fin_supp, blast) |
115 |
have "App (Lam [x].t1) s1 = App (Lam [y].([(y,x)]\<bullet>t1)) s1" using fs |
|
116 |
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
117 |
also have "\<dots> \<longrightarrow>\<^sub>1 ([(y,x)]\<bullet>t2)[y::=s2]" using fs a by (auto simp add: One.eqvt) |
23159 | 118 |
also have "\<dots> = t2[x::=s2]" using fs by (simp add: subst_rename[symmetric]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
119 |
finally show "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2[x::=s2]" by simp |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
120 |
qed |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
121 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
122 |
lemma One_Var: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
123 |
assumes a: "Var x \<longrightarrow>\<^sub>1 M" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
124 |
shows "M = Var x" |
26055 | 125 |
using a by (cases rule: One.cases) (simp_all) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
126 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
127 |
lemma One_Lam: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
128 |
assumes a: "Lam [x].t \<longrightarrow>\<^sub>1 s" "x\<sharp>s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
129 |
shows "\<exists>t'. s = Lam [x].t' \<and> t \<longrightarrow>\<^sub>1 t'" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
130 |
using a |
26055 | 131 |
by (cases rule: One.strong_cases) |
25831 | 132 |
(auto simp add: lam.inject abs_fresh alpha) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
133 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
134 |
lemma One_App: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
135 |
assumes a: "App t s \<longrightarrow>\<^sub>1 r" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
136 |
shows "(\<exists>t' s'. r = App t' s' \<and> t \<longrightarrow>\<^sub>1 t' \<and> s \<longrightarrow>\<^sub>1 s') \<or> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
137 |
(\<exists>x p p' s'. r = p'[x::=s'] \<and> t = Lam [x].p \<and> p \<longrightarrow>\<^sub>1 p' \<and> s \<longrightarrow>\<^sub>1 s' \<and> x\<sharp>(s,s'))" |
26055 | 138 |
using a by (cases rule: One.cases) (auto simp add: lam.inject) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
139 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
140 |
lemma One_Redex: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
141 |
assumes a: "App (Lam [x].t) s \<longrightarrow>\<^sub>1 r" "x\<sharp>(s,r)" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
142 |
shows "(\<exists>t' s'. r = App (Lam [x].t') s' \<and> t \<longrightarrow>\<^sub>1 t' \<and> s \<longrightarrow>\<^sub>1 s') \<or> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
143 |
(\<exists>t' s'. r = t'[x::=s'] \<and> t \<longrightarrow>\<^sub>1 t' \<and> s \<longrightarrow>\<^sub>1 s')" |
23159 | 144 |
using a |
26055 | 145 |
by (cases rule: One.strong_cases) |
25831 | 146 |
(auto dest: One_Lam simp add: lam.inject abs_fresh alpha fresh_prod) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
147 |
|
22829
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
148 |
section {* Transitive Closure of One *} |
f1db55c7534d
tuned some proofs and changed variable names in some definitions of Nominal.thy
urbanc
parents:
22821
diff
changeset
|
149 |
|
23760 | 150 |
inductive |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
151 |
"One_star" :: "lam\<Rightarrow>lam\<Rightarrow>bool" (" _ \<longrightarrow>\<^sub>1\<^sup>* _" [80,80] 80) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
152 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
153 |
os1[intro]: "t \<longrightarrow>\<^sub>1\<^sup>* t" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
154 |
| os2[intro]: "t \<longrightarrow>\<^sub>1 s \<Longrightarrow> t \<longrightarrow>\<^sub>1\<^sup>* s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
155 |
| os3[intro]: "\<lbrakk>t1\<longrightarrow>\<^sub>1\<^sup>* t2; t2 \<longrightarrow>\<^sub>1\<^sup>* t3\<rbrakk> \<Longrightarrow> t1 \<longrightarrow>\<^sub>1\<^sup>* t3" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
156 |
|
23159 | 157 |
section {* Complete Development Reduction *} |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
158 |
|
23760 | 159 |
inductive |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
160 |
Dev :: "lam \<Rightarrow> lam \<Rightarrow> bool" (" _ \<longrightarrow>\<^sub>d _" [80,80]80) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
161 |
where |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
162 |
d1[intro]: "Var x \<longrightarrow>\<^sub>d Var x" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
163 |
| d2[intro]: "t \<longrightarrow>\<^sub>d s \<Longrightarrow> Lam [x].t \<longrightarrow>\<^sub>d Lam[x].s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
164 |
| d3[intro]: "\<lbrakk>\<not>(\<exists>y t'. t1 = Lam [y].t'); t1 \<longrightarrow>\<^sub>d t2; s1 \<longrightarrow>\<^sub>d s2\<rbrakk> \<Longrightarrow> App t1 s1 \<longrightarrow>\<^sub>d App t2 s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
165 |
| d4[intro]: "\<lbrakk>x\<sharp>(s1,s2); t1 \<longrightarrow>\<^sub>d t2; s1 \<longrightarrow>\<^sub>d s2\<rbrakk> \<Longrightarrow> App (Lam [x].t1) s1 \<longrightarrow>\<^sub>d t2[x::=s2]" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
166 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
167 |
equivariance Dev |
25831 | 168 |
nominal_inductive Dev |
26458 | 169 |
by (simp_all add: abs_fresh fresh_fact) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
170 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
171 |
lemma better_d4_intro: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
172 |
assumes a: "t1 \<longrightarrow>\<^sub>d t2" "s1 \<longrightarrow>\<^sub>d s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
173 |
shows "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>d t2[x::=s2]" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
174 |
proof - |
23159 | 175 |
obtain y::"name" where fs: "y\<sharp>(x,t1,s1,t2,s2)" by (rule exists_fresh, rule fin_supp,blast) |
176 |
have "App (Lam [x].t1) s1 = App (Lam [y].([(y,x)]\<bullet>t1)) s1" using fs |
|
177 |
by (auto simp add: lam.inject alpha' fresh_prod fresh_atm) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
178 |
also have "\<dots> \<longrightarrow>\<^sub>d ([(y,x)]\<bullet>t2)[y::=s2]" using fs a by (auto simp add: Dev.eqvt) |
23159 | 179 |
also have "\<dots> = t2[x::=s2]" using fs by (simp add: subst_rename[symmetric]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
180 |
finally show "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>d t2[x::=s2]" by simp |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
181 |
qed |
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
182 |
|
23159 | 183 |
lemma Dev_preserves_fresh: |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
184 |
fixes x::"name" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
185 |
assumes a: "M\<longrightarrow>\<^sub>d N" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
186 |
shows "x\<sharp>M \<Longrightarrow> x\<sharp>N" |
26055 | 187 |
using a |
26458 | 188 |
by (induct) (auto simp add: abs_fresh fresh_fact) |
23159 | 189 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
190 |
lemma Dev_Lam: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
191 |
assumes a: "Lam [x].M \<longrightarrow>\<^sub>d N" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
192 |
shows "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>\<^sub>d N'" |
25831 | 193 |
proof - |
194 |
from a have "x\<sharp>Lam [x].M" by (simp add: abs_fresh) |
|
195 |
with a have "x\<sharp>N" by (simp add: Dev_preserves_fresh) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
196 |
with a show "\<exists>N'. N = Lam [x].N' \<and> M \<longrightarrow>\<^sub>d N'" |
26055 | 197 |
by (cases rule: Dev.strong_cases) |
25831 | 198 |
(auto simp add: lam.inject abs_fresh alpha) |
199 |
qed |
|
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
200 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
201 |
lemma Development_existence: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
202 |
shows "\<exists>M'. M \<longrightarrow>\<^sub>d M'" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
203 |
by (nominal_induct M rule: lam.strong_induct) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
204 |
(auto dest!: Dev_Lam intro: better_d4_intro) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
205 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
206 |
lemma Triangle: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
207 |
assumes a: "t \<longrightarrow>\<^sub>d t1" "t \<longrightarrow>\<^sub>1 t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
208 |
shows "t2 \<longrightarrow>\<^sub>1 t1" |
25831 | 209 |
using a |
26458 | 210 |
proof(nominal_induct avoiding: t2 rule: Dev.strong_induct) |
211 |
case (d4 x s1 s2 t1 t1' t2) |
|
29103 | 212 |
have fc: "x\<sharp>t2" "x\<sharp>s1" by fact+ |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
213 |
have "App (Lam [x].t1) s1 \<longrightarrow>\<^sub>1 t2" by fact |
29103 | 214 |
then obtain t' s' where reds: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
215 |
"(t2 = App (Lam [x].t') s' \<and> t1 \<longrightarrow>\<^sub>1 t' \<and> s1 \<longrightarrow>\<^sub>1 s') \<or> |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
216 |
(t2 = t'[x::=s'] \<and> t1 \<longrightarrow>\<^sub>1 t' \<and> s1 \<longrightarrow>\<^sub>1 s')" |
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
217 |
using fc by (auto dest!: One_Redex) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
218 |
have ih1: "t1 \<longrightarrow>\<^sub>1 t' \<Longrightarrow> t' \<longrightarrow>\<^sub>1 t1'" by fact |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
219 |
have ih2: "s1 \<longrightarrow>\<^sub>1 s' \<Longrightarrow> s' \<longrightarrow>\<^sub>1 s2" by fact |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
220 |
{ assume "t1 \<longrightarrow>\<^sub>1 t'" "s1 \<longrightarrow>\<^sub>1 s'" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
221 |
then have "App (Lam [x].t') s' \<longrightarrow>\<^sub>1 t1'[x::=s2]" |
29103 | 222 |
using ih1 ih2 by (auto intro: better_o4_intro) |
223 |
} |
|
224 |
moreover |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
225 |
{ assume "t1 \<longrightarrow>\<^sub>1 t'" "s1 \<longrightarrow>\<^sub>1 s'" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
226 |
then have "t'[x::=s'] \<longrightarrow>\<^sub>1 t1'[x::=s2]" |
29103 | 227 |
using ih1 ih2 by (auto intro: One_subst) |
228 |
} |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
229 |
ultimately show "t2 \<longrightarrow>\<^sub>1 t1'[x::=s2]" using reds by auto |
26458 | 230 |
qed (auto dest!: One_Lam One_Var One_App) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
231 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
232 |
lemma Diamond_for_One: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
233 |
assumes a: "t \<longrightarrow>\<^sub>1 t1" "t \<longrightarrow>\<^sub>1 t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
234 |
shows "\<exists>t3. t2 \<longrightarrow>\<^sub>1 t3 \<and> t1 \<longrightarrow>\<^sub>1 t3" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
235 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
236 |
obtain tc where "t \<longrightarrow>\<^sub>d tc" using Development_existence by blast |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
237 |
with a have "t2 \<longrightarrow>\<^sub>1 tc" and "t1 \<longrightarrow>\<^sub>1 tc" by (simp_all add: Triangle) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
238 |
then show "\<exists>t3. t2 \<longrightarrow>\<^sub>1 t3 \<and> t1 \<longrightarrow>\<^sub>1 t3" by blast |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
239 |
qed |
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
240 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
241 |
lemma Rectangle_for_One: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
242 |
assumes a: "t \<longrightarrow>\<^sub>1\<^sup>* t1" "t \<longrightarrow>\<^sub>1 t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
243 |
shows "\<exists>t3. t1 \<longrightarrow>\<^sub>1 t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3" |
23159 | 244 |
using a Diamond_for_One by (induct arbitrary: t2) (blast)+ |
245 |
||
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
246 |
lemma CR_for_One_star: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
247 |
assumes a: "t \<longrightarrow>\<^sub>1\<^sup>* t1" "t \<longrightarrow>\<^sub>1\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
248 |
shows "\<exists>t3. t2 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> t1 \<longrightarrow>\<^sub>1\<^sup>* t3" |
23159 | 249 |
using a Rectangle_for_One by (induct arbitrary: t2) (blast)+ |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
250 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
251 |
section {* Establishing the Equivalence of Beta-star and One-star *} |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
252 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
253 |
lemma Beta_Lam_cong: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
254 |
assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
255 |
shows "Lam [x].t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* Lam [x].t2" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
256 |
using a by (induct) (blast)+ |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
257 |
|
26055 | 258 |
lemma Beta_App_cong_aux: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
259 |
assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
260 |
shows "App t1 s\<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
261 |
and "App s t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App s t2" |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
262 |
using a by (induct) (blast)+ |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
263 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
264 |
lemma Beta_App_cong: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
265 |
assumes a: "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" "s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* s2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
266 |
shows "App t1 s1 \<longrightarrow>\<^sub>\<beta>\<^sup>* App t2 s2" |
26055 | 267 |
using a by (blast intro: Beta_App_cong_aux) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
268 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
269 |
lemmas Beta_congs = Beta_Lam_cong Beta_App_cong |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
270 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
271 |
lemma One_implies_Beta_star: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
272 |
assumes a: "t \<longrightarrow>\<^sub>1 s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
273 |
shows "t \<longrightarrow>\<^sub>\<beta>\<^sup>* s" |
23159 | 274 |
using a by (induct) (auto intro!: Beta_congs) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
275 |
|
26966
071f40487734
made the naming of the induction principles consistent: weak_induct is
urbanc
parents:
26458
diff
changeset
|
276 |
lemma One_congs: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
277 |
assumes a: "t1 \<longrightarrow>\<^sub>1\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
278 |
shows "Lam [x].t1 \<longrightarrow>\<^sub>1\<^sup>* Lam [x].t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
279 |
and "App t1 s \<longrightarrow>\<^sub>1\<^sup>* App t2 s" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
280 |
and "App s t1 \<longrightarrow>\<^sub>1\<^sup>* App s t2" |
23159 | 281 |
using a by (induct) (auto intro: One_refl) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
282 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
283 |
lemma Beta_implies_One_star: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
284 |
assumes a: "t1 \<longrightarrow>\<^sub>\<beta> t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
285 |
shows "t1 \<longrightarrow>\<^sub>1\<^sup>* t2" |
22855 | 286 |
using a by (induct) (auto intro: One_refl One_congs better_o4_intro) |
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
287 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
288 |
lemma Beta_star_equals_One_star: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
289 |
shows "t1 \<longrightarrow>\<^sub>1\<^sup>* t2 = t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
290 |
proof |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
291 |
assume "t1 \<longrightarrow>\<^sub>1\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
292 |
then show "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" by (induct) (auto intro: One_implies_Beta_star) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
293 |
next |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
294 |
assume "t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
295 |
then show "t1 \<longrightarrow>\<^sub>1\<^sup>* t2" by (induct) (auto intro: Beta_implies_One_star) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
296 |
qed |
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
297 |
|
22833
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
298 |
section {* The Church-Rosser Theorem *} |
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
299 |
|
51a16a718820
polished all proofs and made the theory "self-contained"
urbanc
parents:
22829
diff
changeset
|
300 |
theorem CR_for_Beta_star: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
301 |
assumes a: "t \<longrightarrow>\<^sub>\<beta>\<^sup>* t1" "t\<longrightarrow>\<^sub>\<beta>\<^sup>* t2" |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
302 |
shows "\<exists>t3. t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3 \<and> t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3" |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
303 |
proof - |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
304 |
from a have "t \<longrightarrow>\<^sub>1\<^sup>* t1" and "t\<longrightarrow>\<^sub>1\<^sup>* t2" by (simp_all add: Beta_star_equals_One_star) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
305 |
then have "\<exists>t3. t1 \<longrightarrow>\<^sub>1\<^sup>* t3 \<and> t2 \<longrightarrow>\<^sub>1\<^sup>* t3" by (simp add: CR_for_One_star) |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
40945
diff
changeset
|
306 |
then show "\<exists>t3. t1 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3 \<and> t2 \<longrightarrow>\<^sub>\<beta>\<^sup>* t3" by (simp add: Beta_star_equals_One_star) |
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
307 |
qed |
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
308 |
|
29103 | 309 |
|
310 |
||
22821
15b2e7ec1f3b
alternative and much simpler proof for Church-Rosser of Beta-Reduction
urbanc
parents:
diff
changeset
|
311 |
end |