author | wenzelm |
Wed, 12 Mar 2025 11:39:00 +0100 | |
changeset 82265 | 4b875a4c83b0 |
parent 81292 | 137b0b107c4b |
permissions | -rw-r--r-- |
39157
b98909faaea8
more explicit HOL-Proofs sessions, including former ex/Hilbert_Classical.thy which works in parallel mode without the antiquotation option "margin" (which is still critical);
wenzelm
parents:
39126
diff
changeset
|
1 |
(* Title: HOL/Proofs/Lambda/Lambda.thy |
1120 | 2 |
Author: Tobias Nipkow |
3 |
Copyright 1995 TU Muenchen |
|
4 |
*) |
|
5 |
||
61986 | 6 |
section \<open>Basic definitions of Lambda-calculus\<close> |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
7 |
|
58372
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents:
58310
diff
changeset
|
8 |
theory Lambda |
58382 | 9 |
imports Main |
58372
bfd497f2f4c2
moved 'old_datatype' out of 'Main' (but put it in 'HOL-Proofs' because of the inductive realizer)
blanchet
parents:
58310
diff
changeset
|
10 |
begin |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
11 |
|
46512
4f9f61f9b535
simplified configuration options for syntax ambiguity;
wenzelm
parents:
46506
diff
changeset
|
12 |
declare [[syntax_ambiguity_warning = false]] |
39126
ee117c5b3b75
configuration options Syntax.ambiguity_enabled (inverse of former Syntax.ambiguity_is_error), Syntax.ambiguity_level (with Isar attribute "syntax_ambiguity_level"), Syntax.ambiguity_limit;
wenzelm
parents:
36862
diff
changeset
|
13 |
|
1120 | 14 |
|
61986 | 15 |
subsection \<open>Lambda-terms in de Bruijn notation and substitution\<close> |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
16 |
|
58310 | 17 |
datatype dB = |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
18 |
Var nat |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
19 |
| App dB dB (infixl \<open>\<degree>\<close> 200) |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
20 |
| Abs dB |
1120 | 21 |
|
25974 | 22 |
primrec |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
23 |
lift :: "[dB, nat] => dB" |
25974 | 24 |
where |
25 |
"lift (Var i) k = (if i < k then Var i else Var (i + 1))" |
|
26 |
| "lift (s \<degree> t) k = lift s k \<degree> lift t k" |
|
27 |
| "lift (Abs s) k = Abs (lift s (k + 1))" |
|
1153 | 28 |
|
5184 | 29 |
primrec |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
30 |
subst :: "[dB, dB, nat] => dB" (\<open>_[_'/_]\<close> [300, 0, 0] 300) |
25974 | 31 |
where (* FIXME base names *) |
32 |
subst_Var: "(Var i)[s/k] = |
|
33 |
(if k < i then Var (i - 1) else if i = k then s else Var i)" |
|
34 |
| subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]" |
|
35 |
| 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
|
36 |
|
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
37 |
declare subst_Var [simp del] |
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
38 |
|
69597 | 39 |
text \<open>Optimized versions of \<^term>\<open>subst\<close> and \<^term>\<open>lift\<close>.\<close> |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
40 |
|
25974 | 41 |
primrec |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
42 |
liftn :: "[nat, dB, nat] => dB" |
25974 | 43 |
where |
44 |
"liftn n (Var i) k = (if i < k then Var i else Var (i + n))" |
|
45 |
| "liftn n (s \<degree> t) k = liftn n s k \<degree> liftn n t k" |
|
46 |
| "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
|
47 |
|
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
48 |
primrec |
25974 | 49 |
substn :: "[dB, dB, nat] => dB" |
50 |
where |
|
51 |
"substn (Var i) s k = |
|
52 |
(if k < i then Var (i - 1) else if i = k then liftn k s 0 else Var i)" |
|
53 |
| "substn (t \<degree> u) s k = substn t s k \<degree> substn u s k" |
|
54 |
| "substn (Abs t) s k = Abs (substn t s (k + 1))" |
|
1120 | 55 |
|
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
56 |
|
61986 | 57 |
subsection \<open>Beta-reduction\<close> |
1153 | 58 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
59 |
inductive beta :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<close> 50) |
22271 | 60 |
where |
61 |
beta [simp, intro!]: "Abs s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]" |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
62 |
| appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u" |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
63 |
| appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t" |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
64 |
| abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t \<Longrightarrow> Abs s \<rightarrow>\<^sub>\<beta> Abs t" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
65 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21210
diff
changeset
|
66 |
abbreviation |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
69597
diff
changeset
|
67 |
beta_reds :: "[dB, dB] => bool" (infixl \<open>\<rightarrow>\<^sub>\<beta>\<^sup>*\<close> 50) where |
67613 | 68 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* t == beta\<^sup>*\<^sup>* s t" |
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] |
61986 | 76 |
\<comment> \<open>don't add \<open>r_into_rtrancl[intro!]\<close>\<close> |
9811
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 |
|
61986 | 79 |
subsection \<open>Congruence rules\<close> |
9811
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!]: |
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
82 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> 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: |
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
86 |
"s \<rightarrow>\<^sub>\<beta>\<^sup>* s' \<Longrightarrow> 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: |
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
90 |
"t \<rightarrow>\<^sub>\<beta>\<^sup>* t' \<Longrightarrow> 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]: |
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
94 |
"\<lbrakk>s \<rightarrow>\<^sub>\<beta>\<^sup>* s'; t \<rightarrow>\<^sub>\<beta>\<^sup>* t'\<rbrakk> \<Longrightarrow> 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 |
|
61986 | 98 |
subsection \<open>Substitution-lemmas\<close> |
9811
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 |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
103 |
lemma subst_gt [simp]: "i < j \<Longrightarrow> (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 |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
106 |
lemma subst_lt [simp]: "j < i \<Longrightarrow> (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 |
|
61986 | 133 |
subsection \<open>Equivalence proof for optimized substitution\<close> |
9811
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 |
|
61986 | 148 |
subsection \<open>Preservation theorems\<close> |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
149 |
|
61986 | 150 |
text \<open>Not used in Church-Rosser proof, but in Strong |
151 |
Normalization. \medskip\<close> |
|
9811
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]: |
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
154 |
"r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> 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 |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
157 |
theorem subst_preserves_beta': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s \<Longrightarrow> r[t/i] \<rightarrow>\<^sub>\<beta>\<^sup>* s[t/i]" |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
158 |
proof (induct set: rtranclp) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
159 |
case base |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
160 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
161 |
by (iprover intro: rtrancl_refl) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
162 |
next |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
163 |
case (step y z) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
164 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
165 |
by (iprover intro: rtranclp.simps subst_preserves_beta) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
166 |
qed |
14065 | 167 |
|
13915
28ccb51bd2f3
Eliminated most occurrences of rule_format attribute.
berghofe
parents:
13187
diff
changeset
|
168 |
theorem lift_preserves_beta [simp]: |
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
169 |
"r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> lift r i \<rightarrow>\<^sub>\<beta> lift s i" |
20503 | 170 |
by (induct arbitrary: i set: beta) auto |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
171 |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
172 |
theorem lift_preserves_beta': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s \<Longrightarrow> lift r i \<rightarrow>\<^sub>\<beta>\<^sup>* lift s i" |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
173 |
proof (induct set: rtranclp) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
174 |
case base |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
175 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
176 |
by (iprover intro: rtrancl_refl) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
177 |
next |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
178 |
case (step y z) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
179 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
180 |
by (iprover intro: lift_preserves_beta rtranclp.simps) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
181 |
qed |
14065 | 182 |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
183 |
theorem subst_preserves_beta2 [simp]: "r \<rightarrow>\<^sub>\<beta> s \<Longrightarrow> t[r/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t[s/i]" |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
184 |
proof (induct t arbitrary: r s i) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
185 |
case (Var x) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
186 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
187 |
by (simp add: subst_Var r_into_rtranclp) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
188 |
next |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
189 |
case (App t1 t2) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
190 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
191 |
by (simp add: rtrancl_beta_App) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
192 |
next |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
193 |
case (Abs t) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
194 |
then show ?case by (simp add: rtrancl_beta_Abs) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
195 |
qed |
9811
39ffdb8cab03
HOL/Lambda: converted into new-style theory and document;
wenzelm
parents:
6453
diff
changeset
|
196 |
|
81292
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
197 |
|
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
198 |
theorem subst_preserves_beta2': "r \<rightarrow>\<^sub>\<beta>\<^sup>* s \<Longrightarrow> t[r/i] \<rightarrow>\<^sub>\<beta>\<^sup>* t[s/i]" |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
199 |
proof (induct set: rtranclp) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
200 |
case base |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
201 |
then show ?case by (iprover intro: rtrancl_refl) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
202 |
next |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
203 |
case (step y z) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
204 |
then show ?case |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
205 |
by (iprover intro: rtranclp_trans subst_preserves_beta2) |
137b0b107c4b
Trying to clean up these horribly ugly old proofs, but the requirement of constructivity makes it very difficult
paulson <lp15@cam.ac.uk>
parents:
80914
diff
changeset
|
206 |
qed |
14065 | 207 |
|
11638 | 208 |
end |