| author | urbanc | 
| Wed, 27 Aug 2008 04:47:30 +0200 | |
| changeset 28011 | 90074908db16 | 
| parent 25974 | 0cb35fa9c6fa | 
| child 36862 | 952b2b102a0a | 
| permissions | -rw-r--r-- | 
| 1269 | 1  | 
(* Title: HOL/Lambda/Lambda.thy  | 
| 1120 | 2  | 
ID: $Id$  | 
3  | 
Author: Tobias Nipkow  | 
|
4  | 
Copyright 1995 TU Muenchen  | 
|
5  | 
*)  | 
|
6  | 
||
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
7  | 
header {* Basic definitions of Lambda-calculus *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
8  | 
|
| 16417 | 9  | 
theory Lambda imports Main begin  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
10  | 
|
| 1120 | 11  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
12  | 
subsection {* Lambda-terms in de Bruijn notation and substitution *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
13  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
14  | 
datatype dB =  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
15  | 
Var nat  | 
| 12011 | 16  | 
| App dB dB (infixl "\<degree>" 200)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
17  | 
| Abs dB  | 
| 1120 | 18  | 
|
| 25974 | 19  | 
primrec  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
20  | 
lift :: "[dB, nat] => dB"  | 
| 25974 | 21  | 
where  | 
22  | 
"lift (Var i) k = (if i < k then Var i else Var (i + 1))"  | 
|
23  | 
| "lift (s \<degree> t) k = lift s k \<degree> lift t k"  | 
|
24  | 
| "lift (Abs s) k = Abs (lift s (k + 1))"  | 
|
| 1153 | 25  | 
|
| 5184 | 26  | 
primrec  | 
| 25974 | 27  | 
  subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
 | 
28  | 
where (* FIXME base names *)  | 
|
29  | 
subst_Var: "(Var i)[s/k] =  | 
|
30  | 
(if k < i then Var (i - 1) else if i = k then s else Var i)"  | 
|
31  | 
| subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"  | 
|
32  | 
| subst_Abs: "(Abs t)[s/k] = Abs (t[lift s 0 / k+1])"  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
33  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
34  | 
declare subst_Var [simp del]  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
35  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
36  | 
text {* Optimized versions of @{term subst} and @{term lift}. *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
37  | 
|
| 25974 | 38  | 
primrec  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
39  | 
liftn :: "[nat, dB, nat] => dB"  | 
| 25974 | 40  | 
where  | 
41  | 
"liftn n (Var i) k = (if i < k then Var i else Var (i + n))"  | 
|
42  | 
| "liftn n (s \<degree> t) k = liftn n s k \<degree> liftn n t k"  | 
|
43  | 
| "liftn n (Abs s) k = Abs (liftn n s (k + 1))"  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
44  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
45  | 
primrec  | 
| 25974 | 46  | 
substn :: "[dB, dB, nat] => dB"  | 
47  | 
where  | 
|
48  | 
"substn (Var i) s k =  | 
|
49  | 
(if k < i then Var (i - 1) else if i = k then liftn k s 0 else Var i)"  | 
|
50  | 
| "substn (t \<degree> u) s k = substn t s k \<degree> substn u s k"  | 
|
51  | 
| "substn (Abs t) s k = Abs (substn t s (k + 1))"  | 
|
| 1120 | 52  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
53  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
54  | 
subsection {* Beta-reduction *}
 | 
| 1153 | 55  | 
|
| 23750 | 56  | 
inductive beta :: "[dB, dB] => bool" (infixl "\<rightarrow>\<^sub>\<beta>" 50)  | 
| 22271 | 57  | 
where  | 
58  | 
beta [simp, intro!]: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"  | 
|
59  | 
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"  | 
|
60  | 
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"  | 
|
61  | 
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs s \<rightarrow>\<^sub>\<beta> Abs t"  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
62  | 
|
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
63  | 
abbreviation  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
64  | 
beta_reds :: "[dB, dB] => bool" (infixl "->>" 50) where  | 
| 22271 | 65  | 
"s ->> t == beta^** s t"  | 
| 19086 | 66  | 
|
| 21210 | 67  | 
notation (latex)  | 
| 
19656
 
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
 
wenzelm 
parents: 
19380 
diff
changeset
 | 
68  | 
beta_reds (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50)  | 
| 1120 | 69  | 
|
| 23750 | 70  | 
inductive_cases beta_cases [elim!]:  | 
| 14065 | 71  | 
"Var i \<rightarrow>\<^sub>\<beta> t"  | 
72  | 
"Abs r \<rightarrow>\<^sub>\<beta> s"  | 
|
73  | 
"s \<degree> t \<rightarrow>\<^sub>\<beta> u"  | 
|
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
74  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
75  | 
declare if_not_P [simp] not_less_eq [simp]  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
76  | 
  -- {* don't add @{text "r_into_rtrancl[intro!]"} *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
77  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
78  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
79  | 
subsection {* Congruence rules *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
80  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
81  | 
lemma rtrancl_beta_Abs [intro!]:  | 
| 14065 | 82  | 
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' ==> Abs s \<rightarrow>\<^sub>\<beta>\<^sup>* Abs s'"  | 
| 23750 | 83  | 
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
84  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
85  | 
lemma rtrancl_beta_AppL:  | 
| 14065 | 86  | 
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' ==> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t"  | 
| 23750 | 87  | 
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
88  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
89  | 
lemma rtrancl_beta_AppR:  | 
| 14065 | 90  | 
"t \<rightarrow>\<^sub>\<beta>\<^sup>* t' ==> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s \<degree> t'"  | 
| 23750 | 91  | 
by (induct set: rtranclp) (blast intro: rtranclp.rtrancl_into_rtrancl)+  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
92  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
93  | 
lemma rtrancl_beta_App [intro]:  | 
| 14065 | 94  | 
"[| s \<rightarrow>\<^sub>\<beta>\<^sup>* s'; t \<rightarrow>\<^sub>\<beta>\<^sup>* t' |] ==> s \<degree> t \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<degree> t'"  | 
| 23750 | 95  | 
by (blast intro!: rtrancl_beta_AppL rtrancl_beta_AppR intro: rtranclp_trans)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
96  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
97  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
98  | 
subsection {* Substitution-lemmas *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
99  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
100  | 
lemma subst_eq [simp]: "(Var k)[u/k] = u"  | 
| 18241 | 101  | 
by (simp add: subst_Var)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
102  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
103  | 
lemma subst_gt [simp]: "i < j ==> (Var j)[u/i] = Var (j - 1)"  | 
| 18241 | 104  | 
by (simp add: subst_Var)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
105  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
106  | 
lemma subst_lt [simp]: "j < i ==> (Var j)[u/i] = Var j"  | 
| 18241 | 107  | 
by (simp add: subst_Var)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
108  | 
|
| 18241 | 109  | 
lemma lift_lift:  | 
110  | 
"i < k + 1 \<Longrightarrow> lift (lift t i) (Suc k) = lift (lift t k) i"  | 
|
| 20503 | 111  | 
by (induct t arbitrary: i k) auto  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
112  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
113  | 
lemma lift_subst [simp]:  | 
| 18241 | 114  | 
"j < i + 1 \<Longrightarrow> lift (t[s/j]) i = (lift t (i + 1)) [lift s i / j]"  | 
| 20503 | 115  | 
by (induct t arbitrary: i j s)  | 
| 18241 | 116  | 
(simp_all add: diff_Suc subst_Var lift_lift split: nat.split)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
117  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
118  | 
lemma lift_subst_lt:  | 
| 18241 | 119  | 
"i < j + 1 \<Longrightarrow> lift (t[s/j]) i = (lift t i) [lift s i / j + 1]"  | 
| 20503 | 120  | 
by (induct t arbitrary: i j s) (simp_all add: subst_Var lift_lift)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
121  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
122  | 
lemma subst_lift [simp]:  | 
| 18241 | 123  | 
"(lift t k)[s/k] = t"  | 
| 20503 | 124  | 
by (induct t arbitrary: k s) simp_all  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
125  | 
|
| 18241 | 126  | 
lemma subst_subst:  | 
127  | 
"i < j + 1 \<Longrightarrow> t[lift v i / Suc j][u[v/j]/i] = t[u/i][v/j]"  | 
|
| 20503 | 128  | 
by (induct t arbitrary: i j u v)  | 
| 18241 | 129  | 
(simp_all add: diff_Suc subst_Var lift_lift [symmetric] lift_subst_lt  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
130  | 
split: nat.split)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
131  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
132  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
133  | 
subsection {* Equivalence proof for optimized substitution *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
134  | 
|
| 18241 | 135  | 
lemma liftn_0 [simp]: "liftn 0 t k = t"  | 
| 20503 | 136  | 
by (induct t arbitrary: k) (simp_all add: subst_Var)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
137  | 
|
| 18241 | 138  | 
lemma liftn_lift [simp]: "liftn (Suc n) t k = lift (liftn n t k) k"  | 
| 20503 | 139  | 
by (induct t arbitrary: k) (simp_all add: subst_Var)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
140  | 
|
| 18241 | 141  | 
lemma substn_subst_n [simp]: "substn t s n = t[liftn n s 0 / n]"  | 
| 20503 | 142  | 
by (induct t arbitrary: n) (simp_all add: subst_Var)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
143  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
144  | 
theorem substn_subst_0: "substn t s 0 = t[s/0]"  | 
| 18241 | 145  | 
by simp  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
146  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
147  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
148  | 
subsection {* Preservation theorems *}
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
149  | 
|
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
150  | 
text {* Not used in Church-Rosser proof, but in Strong
 | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
151  | 
Normalization. \medskip *}  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
152  | 
|
| 
13915
 
28ccb51bd2f3
Eliminated most occurrences of rule_format attribute.
 
berghofe 
parents: 
13187 
diff
changeset
 | 
153  | 
theorem subst_preserves_beta [simp]:  | 
| 18257 | 154  | 
"r \<rightarrow>\<^sub>\<beta> s ==> r[t/i] \<rightarrow>\<^sub>\<beta> s[t/i]"  | 
| 20503 | 155  | 
by (induct arbitrary: t i set: beta) (simp_all add: subst_subst [symmetric])  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
156  | 
|
| 14065 | 157  | 
theorem subst_preserves_beta': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s ==> r[t/i] \<rightarrow>\<^sub>\<beta>\<^sup>* s[t/i]"  | 
| 23750 | 158  | 
apply (induct set: rtranclp)  | 
159  | 
apply (rule rtranclp.rtrancl_refl)  | 
|
160  | 
apply (erule rtranclp.rtrancl_into_rtrancl)  | 
|
| 14065 | 161  | 
apply (erule subst_preserves_beta)  | 
162  | 
done  | 
|
163  | 
||
| 
13915
 
28ccb51bd2f3
Eliminated most occurrences of rule_format attribute.
 
berghofe 
parents: 
13187 
diff
changeset
 | 
164  | 
theorem lift_preserves_beta [simp]:  | 
| 18257 | 165  | 
"r \<rightarrow>\<^sub>\<beta> s ==> lift r i \<rightarrow>\<^sub>\<beta> lift s i"  | 
| 20503 | 166  | 
by (induct arbitrary: i set: beta) auto  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
167  | 
|
| 14065 | 168  | 
theorem lift_preserves_beta': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s ==> lift r i \<rightarrow>\<^sub>\<beta>\<^sup>* lift s i"  | 
| 23750 | 169  | 
apply (induct set: rtranclp)  | 
170  | 
apply (rule rtranclp.rtrancl_refl)  | 
|
171  | 
apply (erule rtranclp.rtrancl_into_rtrancl)  | 
|
| 14065 | 172  | 
apply (erule lift_preserves_beta)  | 
173  | 
done  | 
|
174  | 
||
| 18241 | 175  | 
theorem subst_preserves_beta2 [simp]: "r \<rightarrow>\<^sub>\<beta> s ==> t[r/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t[s/i]"  | 
| 20503 | 176  | 
apply (induct t arbitrary: r s i)  | 
| 23750 | 177  | 
apply (simp add: subst_Var r_into_rtranclp)  | 
| 
9811
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
178  | 
apply (simp add: rtrancl_beta_App)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
179  | 
apply (simp add: rtrancl_beta_Abs)  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
180  | 
done  | 
| 
 
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
 
wenzelm 
parents: 
6453 
diff
changeset
 | 
181  | 
|
| 14065 | 182  | 
theorem subst_preserves_beta2': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s ==> t[r/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t[s/i]"  | 
| 23750 | 183  | 
apply (induct set: rtranclp)  | 
184  | 
apply (rule rtranclp.rtrancl_refl)  | 
|
185  | 
apply (erule rtranclp_trans)  | 
|
| 14065 | 186  | 
apply (erule subst_preserves_beta2)  | 
187  | 
done  | 
|
188  | 
||
| 11638 | 189  | 
end  |