author | wenzelm |
Sun, 14 Mar 2021 22:55:52 +0100 | |
changeset 73439 | cb127ce2c092 |
parent 69913 | ca515cf61651 |
child 80932 | 261cd8722677 |
permissions | -rw-r--r-- |
7700 | 1 |
(* Title: HOL/Inductive.thy |
10402 | 2 |
Author: Markus Wenzel, TU Muenchen |
11688 | 3 |
*) |
10727 | 4 |
|
60758 | 5 |
section \<open>Knaster-Tarski Fixpoint Theorem and inductive definitions\<close> |
1187 | 6 |
|
54398
100c0eaf63d5
moved 'Ctr_Sugar' further up the theory hierarchy, so that 'Datatype' can use it
blanchet
parents:
52143
diff
changeset
|
7 |
theory Inductive |
63588 | 8 |
imports Complete_Lattices Ctr_Sugar |
9 |
keywords |
|
69913 | 10 |
"inductive" "coinductive" "inductive_cases" "inductive_simps" :: thy_defn and |
63588 | 11 |
"monos" and |
12 |
"print_inductives" :: diag and |
|
13 |
"old_rep_datatype" :: thy_goal and |
|
69913 | 14 |
"primrec" :: thy_defn |
15131 | 15 |
begin |
10727 | 16 |
|
63979 | 17 |
subsection \<open>Least fixed points\<close> |
24915 | 18 |
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
19 |
context complete_lattice |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
20 |
begin |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
21 |
|
63979 | 22 |
definition lfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" |
63400 | 23 |
where "lfp f = Inf {u. f u \<le> u}" |
24915 | 24 |
|
63400 | 25 |
lemma lfp_lowerbound: "f A \<le> A \<Longrightarrow> lfp f \<le> A" |
63981 | 26 |
unfolding lfp_def by (rule Inf_lower) simp |
24915 | 27 |
|
63400 | 28 |
lemma lfp_greatest: "(\<And>u. f u \<le> u \<Longrightarrow> A \<le> u) \<Longrightarrow> A \<le> lfp f" |
63981 | 29 |
unfolding lfp_def by (rule Inf_greatest) simp |
24915 | 30 |
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
31 |
end |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
32 |
|
63979 | 33 |
lemma lfp_fixpoint: |
34 |
assumes "mono f" |
|
35 |
shows "f (lfp f) = lfp f" |
|
36 |
unfolding lfp_def |
|
37 |
proof (rule order_antisym) |
|
38 |
let ?H = "{u. f u \<le> u}" |
|
39 |
let ?a = "\<Sqinter>?H" |
|
40 |
show "f ?a \<le> ?a" |
|
41 |
proof (rule Inf_greatest) |
|
42 |
fix x |
|
43 |
assume "x \<in> ?H" |
|
44 |
then have "?a \<le> x" by (rule Inf_lower) |
|
45 |
with \<open>mono f\<close> have "f ?a \<le> f x" .. |
|
46 |
also from \<open>x \<in> ?H\<close> have "f x \<le> x" .. |
|
47 |
finally show "f ?a \<le> x" . |
|
48 |
qed |
|
49 |
show "?a \<le> f ?a" |
|
50 |
proof (rule Inf_lower) |
|
51 |
from \<open>mono f\<close> and \<open>f ?a \<le> ?a\<close> have "f (f ?a) \<le> f ?a" .. |
|
52 |
then show "f ?a \<in> ?H" .. |
|
53 |
qed |
|
54 |
qed |
|
24915 | 55 |
|
63400 | 56 |
lemma lfp_unfold: "mono f \<Longrightarrow> lfp f = f (lfp f)" |
63979 | 57 |
by (rule lfp_fixpoint [symmetric]) |
24915 | 58 |
|
59 |
lemma lfp_const: "lfp (\<lambda>x. t) = t" |
|
63400 | 60 |
by (rule lfp_unfold) (simp add: mono_def) |
24915 | 61 |
|
63588 | 62 |
lemma lfp_eqI: "mono F \<Longrightarrow> F x = x \<Longrightarrow> (\<And>z. F z = z \<Longrightarrow> x \<le> z) \<Longrightarrow> lfp F = x" |
63 |
by (rule antisym) (simp_all add: lfp_lowerbound lfp_unfold[symmetric]) |
|
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
64 |
|
24915 | 65 |
|
60758 | 66 |
subsection \<open>General induction rules for least fixed points\<close> |
24915 | 67 |
|
63400 | 68 |
lemma lfp_ordinal_induct [case_names mono step union]: |
61076 | 69 |
fixes f :: "'a::complete_lattice \<Rightarrow> 'a" |
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
70 |
assumes mono: "mono f" |
63400 | 71 |
and P_f: "\<And>S. P S \<Longrightarrow> S \<le> lfp f \<Longrightarrow> P (f S)" |
72 |
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Sup M)" |
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
73 |
shows "P (lfp f)" |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
74 |
proof - |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
75 |
let ?M = "{S. S \<le> lfp f \<and> P S}" |
63588 | 76 |
from P_Union have "P (Sup ?M)" by simp |
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
77 |
also have "Sup ?M = lfp f" |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
78 |
proof (rule antisym) |
63588 | 79 |
show "Sup ?M \<le> lfp f" |
80 |
by (blast intro: Sup_least) |
|
63400 | 81 |
then have "f (Sup ?M) \<le> f (lfp f)" |
82 |
by (rule mono [THEN monoD]) |
|
83 |
then have "f (Sup ?M) \<le> lfp f" |
|
84 |
using mono [THEN lfp_unfold] by simp |
|
85 |
then have "f (Sup ?M) \<in> ?M" |
|
86 |
using P_Union by simp (intro P_f Sup_least, auto) |
|
87 |
then have "f (Sup ?M) \<le> Sup ?M" |
|
88 |
by (rule Sup_upper) |
|
89 |
then show "lfp f \<le> Sup ?M" |
|
90 |
by (rule lfp_lowerbound) |
|
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
91 |
qed |
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
92 |
finally show ?thesis . |
63400 | 93 |
qed |
26013
8764a1f1253b
Theorem Inductive.lfp_ordinal_induct generalized to complete lattices
haftmann
parents:
25557
diff
changeset
|
94 |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
95 |
theorem lfp_induct: |
63400 | 96 |
assumes mono: "mono f" |
97 |
and ind: "f (inf (lfp f) P) \<le> P" |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
98 |
shows "lfp f \<le> P" |
63588 | 99 |
proof (induct rule: lfp_ordinal_induct) |
100 |
case mono |
|
101 |
show ?case by fact |
|
102 |
next |
|
63400 | 103 |
case (step S) |
104 |
then show ?case |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
105 |
by (intro order_trans[OF _ ind] monoD[OF mono]) auto |
63588 | 106 |
next |
107 |
case (union M) |
|
108 |
then show ?case |
|
109 |
by (auto intro: Sup_least) |
|
110 |
qed |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
111 |
|
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
112 |
lemma lfp_induct_set: |
63400 | 113 |
assumes lfp: "a \<in> lfp f" |
114 |
and mono: "mono f" |
|
115 |
and hyp: "\<And>x. x \<in> f (lfp f \<inter> {x. P x}) \<Longrightarrow> P x" |
|
116 |
shows "P a" |
|
117 |
by (rule lfp_induct [THEN subsetD, THEN CollectD, OF mono _ lfp]) (auto intro: hyp) |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
118 |
|
63400 | 119 |
lemma lfp_ordinal_induct_set: |
24915 | 120 |
assumes mono: "mono f" |
63400 | 121 |
and P_f: "\<And>S. P S \<Longrightarrow> P (f S)" |
122 |
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (\<Union>M)" |
|
123 |
shows "P (lfp f)" |
|
46008
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
45907
diff
changeset
|
124 |
using assms by (rule lfp_ordinal_induct) |
24915 | 125 |
|
126 |
||
63400 | 127 |
text \<open>Definition forms of \<open>lfp_unfold\<close> and \<open>lfp_induct\<close>, to control unfolding.\<close> |
24915 | 128 |
|
63400 | 129 |
lemma def_lfp_unfold: "h \<equiv> lfp f \<Longrightarrow> mono f \<Longrightarrow> h = f h" |
45899 | 130 |
by (auto intro!: lfp_unfold) |
24915 | 131 |
|
63400 | 132 |
lemma def_lfp_induct: "A \<equiv> lfp f \<Longrightarrow> mono f \<Longrightarrow> f (inf A P) \<le> P \<Longrightarrow> A \<le> P" |
24915 | 133 |
by (blast intro: lfp_induct) |
134 |
||
63400 | 135 |
lemma def_lfp_induct_set: |
136 |
"A \<equiv> lfp f \<Longrightarrow> mono f \<Longrightarrow> a \<in> A \<Longrightarrow> (\<And>x. x \<in> f (A \<inter> {x. P x}) \<Longrightarrow> P x) \<Longrightarrow> P a" |
|
24915 | 137 |
by (blast intro: lfp_induct_set) |
138 |
||
63400 | 139 |
text \<open>Monotonicity of \<open>lfp\<close>!\<close> |
140 |
lemma lfp_mono: "(\<And>Z. f Z \<le> g Z) \<Longrightarrow> lfp f \<le> lfp g" |
|
141 |
by (rule lfp_lowerbound [THEN lfp_greatest]) (blast intro: order_trans) |
|
24915 | 142 |
|
143 |
||
63979 | 144 |
subsection \<open>Greatest fixed points\<close> |
24915 | 145 |
|
63979 | 146 |
context complete_lattice |
147 |
begin |
|
148 |
||
149 |
definition gfp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a" |
|
150 |
where "gfp f = Sup {u. u \<le> f u}" |
|
24915 | 151 |
|
63400 | 152 |
lemma gfp_upperbound: "X \<le> f X \<Longrightarrow> X \<le> gfp f" |
24915 | 153 |
by (auto simp add: gfp_def intro: Sup_upper) |
154 |
||
63400 | 155 |
lemma gfp_least: "(\<And>u. u \<le> f u \<Longrightarrow> u \<le> X) \<Longrightarrow> gfp f \<le> X" |
24915 | 156 |
by (auto simp add: gfp_def intro: Sup_least) |
157 |
||
63979 | 158 |
end |
159 |
||
160 |
lemma lfp_le_gfp: "mono f \<Longrightarrow> lfp f \<le> gfp f" |
|
161 |
by (rule gfp_upperbound) (simp add: lfp_fixpoint) |
|
24915 | 162 |
|
63979 | 163 |
lemma gfp_fixpoint: |
164 |
assumes "mono f" |
|
165 |
shows "f (gfp f) = gfp f" |
|
166 |
unfolding gfp_def |
|
167 |
proof (rule order_antisym) |
|
168 |
let ?H = "{u. u \<le> f u}" |
|
169 |
let ?a = "\<Squnion>?H" |
|
170 |
show "?a \<le> f ?a" |
|
171 |
proof (rule Sup_least) |
|
172 |
fix x |
|
173 |
assume "x \<in> ?H" |
|
174 |
then have "x \<le> f x" .. |
|
175 |
also from \<open>x \<in> ?H\<close> have "x \<le> ?a" by (rule Sup_upper) |
|
176 |
with \<open>mono f\<close> have "f x \<le> f ?a" .. |
|
177 |
finally show "x \<le> f ?a" . |
|
178 |
qed |
|
179 |
show "f ?a \<le> ?a" |
|
180 |
proof (rule Sup_upper) |
|
181 |
from \<open>mono f\<close> and \<open>?a \<le> f ?a\<close> have "f ?a \<le> f (f ?a)" .. |
|
182 |
then show "f ?a \<in> ?H" .. |
|
183 |
qed |
|
184 |
qed |
|
24915 | 185 |
|
63400 | 186 |
lemma gfp_unfold: "mono f \<Longrightarrow> gfp f = f (gfp f)" |
63979 | 187 |
by (rule gfp_fixpoint [symmetric]) |
24915 | 188 |
|
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
189 |
lemma gfp_const: "gfp (\<lambda>x. t) = t" |
63588 | 190 |
by (rule gfp_unfold) (simp add: mono_def) |
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
191 |
|
63588 | 192 |
lemma gfp_eqI: "mono F \<Longrightarrow> F x = x \<Longrightarrow> (\<And>z. F z = z \<Longrightarrow> z \<le> x) \<Longrightarrow> gfp F = x" |
193 |
by (rule antisym) (simp_all add: gfp_upperbound gfp_unfold[symmetric]) |
|
63561
fba08009ff3e
add lemmas contributed by Peter Gammie
Andreas Lochbihler
parents:
63540
diff
changeset
|
194 |
|
24915 | 195 |
|
60758 | 196 |
subsection \<open>Coinduction rules for greatest fixed points\<close> |
24915 | 197 |
|
63400 | 198 |
text \<open>Weak version.\<close> |
199 |
lemma weak_coinduct: "a \<in> X \<Longrightarrow> X \<subseteq> f X \<Longrightarrow> a \<in> gfp f" |
|
45899 | 200 |
by (rule gfp_upperbound [THEN subsetD]) auto |
24915 | 201 |
|
63400 | 202 |
lemma weak_coinduct_image: "a \<in> X \<Longrightarrow> g`X \<subseteq> f (g`X) \<Longrightarrow> g a \<in> gfp f" |
45899 | 203 |
apply (erule gfp_upperbound [THEN subsetD]) |
204 |
apply (erule imageI) |
|
205 |
done |
|
24915 | 206 |
|
63400 | 207 |
lemma coinduct_lemma: "X \<le> f (sup X (gfp f)) \<Longrightarrow> mono f \<Longrightarrow> sup X (gfp f) \<le> f (sup X (gfp f))" |
63979 | 208 |
apply (frule gfp_unfold [THEN eq_refl]) |
24915 | 209 |
apply (drule mono_sup) |
210 |
apply (rule le_supI) |
|
63588 | 211 |
apply assumption |
24915 | 212 |
apply (rule order_trans) |
63588 | 213 |
apply (rule order_trans) |
214 |
apply assumption |
|
215 |
apply (rule sup_ge2) |
|
24915 | 216 |
apply assumption |
217 |
done |
|
218 |
||
63400 | 219 |
text \<open>Strong version, thanks to Coen and Frost.\<close> |
220 |
lemma coinduct_set: "mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (X \<union> gfp f) \<Longrightarrow> a \<in> gfp f" |
|
55604 | 221 |
by (rule weak_coinduct[rotated], rule coinduct_lemma) blast+ |
24915 | 222 |
|
63400 | 223 |
lemma gfp_fun_UnI2: "mono f \<Longrightarrow> a \<in> gfp f \<Longrightarrow> a \<in> f (X \<union> gfp f)" |
63979 | 224 |
by (blast dest: gfp_fixpoint mono_Un) |
24915 | 225 |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
226 |
lemma gfp_ordinal_induct[case_names mono step union]: |
61076 | 227 |
fixes f :: "'a::complete_lattice \<Rightarrow> 'a" |
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
228 |
assumes mono: "mono f" |
63400 | 229 |
and P_f: "\<And>S. P S \<Longrightarrow> gfp f \<le> S \<Longrightarrow> P (f S)" |
230 |
and P_Union: "\<And>M. \<forall>S\<in>M. P S \<Longrightarrow> P (Inf M)" |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
231 |
shows "P (gfp f)" |
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
232 |
proof - |
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
233 |
let ?M = "{S. gfp f \<le> S \<and> P S}" |
63588 | 234 |
from P_Union have "P (Inf ?M)" by simp |
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
235 |
also have "Inf ?M = gfp f" |
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
236 |
proof (rule antisym) |
63400 | 237 |
show "gfp f \<le> Inf ?M" |
238 |
by (blast intro: Inf_greatest) |
|
239 |
then have "f (gfp f) \<le> f (Inf ?M)" |
|
240 |
by (rule mono [THEN monoD]) |
|
241 |
then have "gfp f \<le> f (Inf ?M)" |
|
242 |
using mono [THEN gfp_unfold] by simp |
|
243 |
then have "f (Inf ?M) \<in> ?M" |
|
244 |
using P_Union by simp (intro P_f Inf_greatest, auto) |
|
245 |
then have "Inf ?M \<le> f (Inf ?M)" |
|
246 |
by (rule Inf_lower) |
|
247 |
then show "Inf ?M \<le> gfp f" |
|
248 |
by (rule gfp_upperbound) |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
249 |
qed |
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
250 |
finally show ?thesis . |
63400 | 251 |
qed |
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
252 |
|
63400 | 253 |
lemma coinduct: |
254 |
assumes mono: "mono f" |
|
255 |
and ind: "X \<le> f (sup X (gfp f))" |
|
256 |
shows "X \<le> gfp f" |
|
63588 | 257 |
proof (induct rule: gfp_ordinal_induct) |
258 |
case mono |
|
259 |
then show ?case by fact |
|
260 |
next |
|
261 |
case (step S) |
|
262 |
then show ?case |
|
60174
70d8f7abde8f
strengthened lfp_ordinal_induct; added dual gfp variant
hoelzl
parents:
60173
diff
changeset
|
263 |
by (intro order_trans[OF ind _] monoD[OF mono]) auto |
63588 | 264 |
next |
265 |
case (union M) |
|
266 |
then show ?case |
|
267 |
by (auto intro: mono Inf_greatest) |
|
268 |
qed |
|
24915 | 269 |
|
63400 | 270 |
|
60758 | 271 |
subsection \<open>Even Stronger Coinduction Rule, by Martin Coen\<close> |
24915 | 272 |
|
69593 | 273 |
text \<open>Weakens the condition \<^term>\<open>X \<subseteq> f X\<close> to one expressed using both |
274 |
\<^term>\<open>lfp\<close> and \<^term>\<open>gfp\<close>\<close> |
|
63400 | 275 |
lemma coinduct3_mono_lemma: "mono f \<Longrightarrow> mono (\<lambda>x. f x \<union> X \<union> B)" |
276 |
by (iprover intro: subset_refl monoI Un_mono monoD) |
|
24915 | 277 |
|
278 |
lemma coinduct3_lemma: |
|
63400 | 279 |
"X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f)) \<Longrightarrow> mono f \<Longrightarrow> |
280 |
lfp (\<lambda>x. f x \<union> X \<union> gfp f) \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f))" |
|
281 |
apply (rule subset_trans) |
|
63979 | 282 |
apply (erule coinduct3_mono_lemma [THEN lfp_unfold [THEN eq_refl]]) |
63400 | 283 |
apply (rule Un_least [THEN Un_least]) |
63588 | 284 |
apply (rule subset_refl, assumption) |
63400 | 285 |
apply (rule gfp_unfold [THEN equalityD1, THEN subset_trans], assumption) |
286 |
apply (rule monoD, assumption) |
|
287 |
apply (subst coinduct3_mono_lemma [THEN lfp_unfold], auto) |
|
288 |
done |
|
24915 | 289 |
|
63400 | 290 |
lemma coinduct3: "mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> gfp f)) \<Longrightarrow> a \<in> gfp f" |
291 |
apply (rule coinduct3_lemma [THEN [2] weak_coinduct]) |
|
63588 | 292 |
apply (rule coinduct3_mono_lemma [THEN lfp_unfold, THEN ssubst]) |
293 |
apply simp_all |
|
63400 | 294 |
done |
24915 | 295 |
|
63400 | 296 |
text \<open>Definition forms of \<open>gfp_unfold\<close> and \<open>coinduct\<close>, to control unfolding.\<close> |
24915 | 297 |
|
63400 | 298 |
lemma def_gfp_unfold: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> A = f A" |
45899 | 299 |
by (auto intro!: gfp_unfold) |
24915 | 300 |
|
63400 | 301 |
lemma def_coinduct: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> X \<le> f (sup X A) \<Longrightarrow> X \<le> A" |
45899 | 302 |
by (iprover intro!: coinduct) |
24915 | 303 |
|
63400 | 304 |
lemma def_coinduct_set: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (X \<union> A) \<Longrightarrow> a \<in> A" |
45899 | 305 |
by (auto intro!: coinduct_set) |
24915 | 306 |
|
307 |
lemma def_Collect_coinduct: |
|
63400 | 308 |
"A \<equiv> gfp (\<lambda>w. Collect (P w)) \<Longrightarrow> mono (\<lambda>w. Collect (P w)) \<Longrightarrow> a \<in> X \<Longrightarrow> |
309 |
(\<And>z. z \<in> X \<Longrightarrow> P (X \<union> A) z) \<Longrightarrow> a \<in> A" |
|
45899 | 310 |
by (erule def_coinduct_set) auto |
24915 | 311 |
|
63400 | 312 |
lemma def_coinduct3: "A \<equiv> gfp f \<Longrightarrow> mono f \<Longrightarrow> a \<in> X \<Longrightarrow> X \<subseteq> f (lfp (\<lambda>x. f x \<union> X \<union> A)) \<Longrightarrow> a \<in> A" |
45899 | 313 |
by (auto intro!: coinduct3) |
24915 | 314 |
|
69593 | 315 |
text \<open>Monotonicity of \<^term>\<open>gfp\<close>!\<close> |
63400 | 316 |
lemma gfp_mono: "(\<And>Z. f Z \<le> g Z) \<Longrightarrow> gfp f \<le> gfp g" |
317 |
by (rule gfp_upperbound [THEN gfp_least]) (blast intro: order_trans) |
|
318 |
||
24915 | 319 |
|
60758 | 320 |
subsection \<open>Rules for fixed point calculus\<close> |
60173 | 321 |
|
322 |
lemma lfp_rolling: |
|
323 |
assumes "mono g" "mono f" |
|
324 |
shows "g (lfp (\<lambda>x. f (g x))) = lfp (\<lambda>x. g (f x))" |
|
325 |
proof (rule antisym) |
|
326 |
have *: "mono (\<lambda>x. f (g x))" |
|
327 |
using assms by (auto simp: mono_def) |
|
328 |
show "lfp (\<lambda>x. g (f x)) \<le> g (lfp (\<lambda>x. f (g x)))" |
|
329 |
by (rule lfp_lowerbound) (simp add: lfp_unfold[OF *, symmetric]) |
|
330 |
show "g (lfp (\<lambda>x. f (g x))) \<le> lfp (\<lambda>x. g (f x))" |
|
331 |
proof (rule lfp_greatest) |
|
63400 | 332 |
fix u |
63540 | 333 |
assume u: "g (f u) \<le> u" |
334 |
then have "g (lfp (\<lambda>x. f (g x))) \<le> g (f u)" |
|
60173 | 335 |
by (intro assms[THEN monoD] lfp_lowerbound) |
63540 | 336 |
with u show "g (lfp (\<lambda>x. f (g x))) \<le> u" |
60173 | 337 |
by auto |
338 |
qed |
|
339 |
qed |
|
340 |
||
341 |
lemma lfp_lfp: |
|
342 |
assumes f: "\<And>x y w z. x \<le> y \<Longrightarrow> w \<le> z \<Longrightarrow> f x w \<le> f y z" |
|
343 |
shows "lfp (\<lambda>x. lfp (f x)) = lfp (\<lambda>x. f x x)" |
|
344 |
proof (rule antisym) |
|
345 |
have *: "mono (\<lambda>x. f x x)" |
|
346 |
by (blast intro: monoI f) |
|
347 |
show "lfp (\<lambda>x. lfp (f x)) \<le> lfp (\<lambda>x. f x x)" |
|
348 |
by (intro lfp_lowerbound) (simp add: lfp_unfold[OF *, symmetric]) |
|
349 |
show "lfp (\<lambda>x. lfp (f x)) \<ge> lfp (\<lambda>x. f x x)" (is "?F \<ge> _") |
|
350 |
proof (intro lfp_lowerbound) |
|
351 |
have *: "?F = lfp (f ?F)" |
|
352 |
by (rule lfp_unfold) (blast intro: monoI lfp_mono f) |
|
353 |
also have "\<dots> = f ?F (lfp (f ?F))" |
|
354 |
by (rule lfp_unfold) (blast intro: monoI lfp_mono f) |
|
355 |
finally show "f ?F ?F \<le> ?F" |
|
356 |
by (simp add: *[symmetric]) |
|
357 |
qed |
|
358 |
qed |
|
359 |
||
360 |
lemma gfp_rolling: |
|
361 |
assumes "mono g" "mono f" |
|
362 |
shows "g (gfp (\<lambda>x. f (g x))) = gfp (\<lambda>x. g (f x))" |
|
363 |
proof (rule antisym) |
|
364 |
have *: "mono (\<lambda>x. f (g x))" |
|
365 |
using assms by (auto simp: mono_def) |
|
366 |
show "g (gfp (\<lambda>x. f (g x))) \<le> gfp (\<lambda>x. g (f x))" |
|
367 |
by (rule gfp_upperbound) (simp add: gfp_unfold[OF *, symmetric]) |
|
368 |
show "gfp (\<lambda>x. g (f x)) \<le> g (gfp (\<lambda>x. f (g x)))" |
|
369 |
proof (rule gfp_least) |
|
63540 | 370 |
fix u |
371 |
assume u: "u \<le> g (f u)" |
|
372 |
then have "g (f u) \<le> g (gfp (\<lambda>x. f (g x)))" |
|
60173 | 373 |
by (intro assms[THEN monoD] gfp_upperbound) |
63540 | 374 |
with u show "u \<le> g (gfp (\<lambda>x. f (g x)))" |
60173 | 375 |
by auto |
376 |
qed |
|
377 |
qed |
|
378 |
||
379 |
lemma gfp_gfp: |
|
380 |
assumes f: "\<And>x y w z. x \<le> y \<Longrightarrow> w \<le> z \<Longrightarrow> f x w \<le> f y z" |
|
381 |
shows "gfp (\<lambda>x. gfp (f x)) = gfp (\<lambda>x. f x x)" |
|
382 |
proof (rule antisym) |
|
383 |
have *: "mono (\<lambda>x. f x x)" |
|
384 |
by (blast intro: monoI f) |
|
385 |
show "gfp (\<lambda>x. f x x) \<le> gfp (\<lambda>x. gfp (f x))" |
|
386 |
by (intro gfp_upperbound) (simp add: gfp_unfold[OF *, symmetric]) |
|
387 |
show "gfp (\<lambda>x. gfp (f x)) \<le> gfp (\<lambda>x. f x x)" (is "?F \<le> _") |
|
388 |
proof (intro gfp_upperbound) |
|
389 |
have *: "?F = gfp (f ?F)" |
|
390 |
by (rule gfp_unfold) (blast intro: monoI gfp_mono f) |
|
391 |
also have "\<dots> = f ?F (gfp (f ?F))" |
|
392 |
by (rule gfp_unfold) (blast intro: monoI gfp_mono f) |
|
393 |
finally show "?F \<le> f ?F ?F" |
|
394 |
by (simp add: *[symmetric]) |
|
395 |
qed |
|
396 |
qed |
|
24915 | 397 |
|
63400 | 398 |
|
60758 | 399 |
subsection \<open>Inductive predicates and sets\<close> |
11688 | 400 |
|
60758 | 401 |
text \<open>Package setup.\<close> |
10402 | 402 |
|
61337 | 403 |
lemmas basic_monos = |
22218 | 404 |
subset_refl imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj |
11688 | 405 |
Collect_mono in_mono vimage_mono |
406 |
||
63863
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
407 |
lemma le_rel_bool_arg_iff: "X \<le> Y \<longleftrightarrow> X False \<le> Y False \<and> X True \<le> Y True" |
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
408 |
unfolding le_fun_def le_bool_def using bool_induct by auto |
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
409 |
|
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
410 |
lemma imp_conj_iff: "((P \<longrightarrow> Q) \<and> P) = (P \<and> Q)" |
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
411 |
by blast |
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
412 |
|
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
413 |
lemma meta_fun_cong: "P \<equiv> Q \<Longrightarrow> P a \<equiv> Q a" |
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
414 |
by auto |
d14e580c3b8f
don't expose internal construction in the coinduction rule for mutual coinductive predicates
traytel
parents:
63588
diff
changeset
|
415 |
|
69605 | 416 |
ML_file \<open>Tools/inductive.ML\<close> |
21018
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
417 |
|
61337 | 418 |
lemmas [mono] = |
22218 | 419 |
imp_refl disj_mono conj_mono ex_mono all_mono if_bool_eq_conj |
33934
25d6a8982e37
Streamlined setup for monotonicity rules (no longer requires classical rules).
berghofe
parents:
32701
diff
changeset
|
420 |
imp_mono not_mono |
21018
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
421 |
Ball_def Bex_def |
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
422 |
induct_rulify_fallback |
e6b8d6784792
Added new package for inductive definitions, moved old package
berghofe
parents:
20604
diff
changeset
|
423 |
|
11688 | 424 |
|
63980
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
425 |
subsection \<open>The Schroeder-Bernstein Theorem\<close> |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
426 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
427 |
text \<open> |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
428 |
See also: |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
429 |
\<^item> \<^file>\<open>$ISABELLE_HOME/src/HOL/ex/Set_Theory.thy\<close> |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
430 |
\<^item> \<^url>\<open>http://planetmath.org/proofofschroederbernsteintheoremusingtarskiknastertheorem\<close> |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
431 |
\<^item> Springer LNCS 828 (cover page) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
432 |
\<close> |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
433 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
434 |
theorem Schroeder_Bernstein: |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
435 |
fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'a" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
436 |
and A :: "'a set" and B :: "'b set" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
437 |
assumes inj1: "inj_on f A" and sub1: "f ` A \<subseteq> B" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
438 |
and inj2: "inj_on g B" and sub2: "g ` B \<subseteq> A" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
439 |
shows "\<exists>h. bij_betw h A B" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
440 |
proof (rule exI, rule bij_betw_imageI) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
441 |
define X where "X = lfp (\<lambda>X. A - (g ` (B - (f ` X))))" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
442 |
define g' where "g' = the_inv_into (B - (f ` X)) g" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
443 |
let ?h = "\<lambda>z. if z \<in> X then f z else g' z" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
444 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
445 |
have X: "X = A - (g ` (B - (f ` X)))" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
446 |
unfolding X_def by (rule lfp_unfold) (blast intro: monoI) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
447 |
then have X_compl: "A - X = g ` (B - (f ` X))" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
448 |
using sub2 by blast |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
449 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
450 |
from inj2 have inj2': "inj_on g (B - (f ` X))" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
451 |
by (rule inj_on_subset) auto |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
452 |
with X_compl have *: "g' ` (A - X) = B - (f ` X)" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
453 |
by (simp add: g'_def) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
454 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
455 |
from X have X_sub: "X \<subseteq> A" by auto |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
456 |
from X sub1 have fX_sub: "f ` X \<subseteq> B" by auto |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
457 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
458 |
show "?h ` A = B" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
459 |
proof - |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
460 |
from X_sub have "?h ` A = ?h ` (X \<union> (A - X))" by auto |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
461 |
also have "\<dots> = ?h ` X \<union> ?h ` (A - X)" by (simp only: image_Un) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
462 |
also have "?h ` X = f ` X" by auto |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
463 |
also from * have "?h ` (A - X) = B - (f ` X)" by auto |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
464 |
also from fX_sub have "f ` X \<union> (B - f ` X) = B" by blast |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
465 |
finally show ?thesis . |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
466 |
qed |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
467 |
show "inj_on ?h A" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
468 |
proof - |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
469 |
from inj1 X_sub have on_X: "inj_on f X" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
470 |
by (rule subset_inj_on) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
471 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
472 |
have on_X_compl: "inj_on g' (A - X)" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
473 |
unfolding g'_def X_compl |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
474 |
by (rule inj_on_the_inv_into) (rule inj2') |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
475 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
476 |
have impossible: False if eq: "f a = g' b" and a: "a \<in> X" and b: "b \<in> A - X" for a b |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
477 |
proof - |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
478 |
from a have fa: "f a \<in> f ` X" by (rule imageI) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
479 |
from b have "g' b \<in> g' ` (A - X)" by (rule imageI) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
480 |
with * have "g' b \<in> - (f ` X)" by simp |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
481 |
with eq fa show False by simp |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
482 |
qed |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
483 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
484 |
show ?thesis |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
485 |
proof (rule inj_onI) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
486 |
fix a b |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
487 |
assume h: "?h a = ?h b" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
488 |
assume "a \<in> A" and "b \<in> A" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
489 |
then consider "a \<in> X" "b \<in> X" | "a \<in> A - X" "b \<in> A - X" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
490 |
| "a \<in> X" "b \<in> A - X" | "a \<in> A - X" "b \<in> X" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
491 |
by blast |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
492 |
then show "a = b" |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
493 |
proof cases |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
494 |
case 1 |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
495 |
with h on_X show ?thesis by (simp add: inj_on_eq_iff) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
496 |
next |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
497 |
case 2 |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
498 |
with h on_X_compl show ?thesis by (simp add: inj_on_eq_iff) |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
499 |
next |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
500 |
case 3 |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
501 |
with h impossible [of a b] have False by simp |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
502 |
then show ?thesis .. |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
503 |
next |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
504 |
case 4 |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
505 |
with h impossible [of b a] have False by simp |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
506 |
then show ?thesis .. |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
507 |
qed |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
508 |
qed |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
509 |
qed |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
510 |
qed |
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
511 |
|
f8e556c8ad6f
Isar proof of Schroeder_Bernstein without using Hilbert_Choice (and metis);
wenzelm
parents:
63979
diff
changeset
|
512 |
|
60758 | 513 |
subsection \<open>Inductive datatypes and primitive recursion\<close> |
11688 | 514 |
|
60758 | 515 |
text \<open>Package setup.\<close> |
11825 | 516 |
|
69605 | 517 |
ML_file \<open>Tools/Old_Datatype/old_datatype_aux.ML\<close> |
518 |
ML_file \<open>Tools/Old_Datatype/old_datatype_prop.ML\<close> |
|
519 |
ML_file \<open>Tools/Old_Datatype/old_datatype_data.ML\<close> |
|
520 |
ML_file \<open>Tools/Old_Datatype/old_rep_datatype.ML\<close> |
|
521 |
ML_file \<open>Tools/Old_Datatype/old_datatype_codegen.ML\<close> |
|
522 |
ML_file \<open>Tools/BNF/bnf_fp_rec_sugar_util.ML\<close> |
|
523 |
ML_file \<open>Tools/Old_Datatype/old_primrec.ML\<close> |
|
524 |
ML_file \<open>Tools/BNF/bnf_lfp_rec_sugar.ML\<close> |
|
55575
a5e33e18fb5c
moved 'primrec' up (for real this time) and removed temporary 'old_primrec'
blanchet
parents:
55534
diff
changeset
|
525 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
526 |
text \<open>Lambda-abstractions with pattern matching:\<close> |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
527 |
syntax (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
528 |
"_lam_pats_syntax" :: "cases_syn \<Rightarrow> 'a \<Rightarrow> 'b" ("(%_)" 10) |
23526 | 529 |
syntax |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61952
diff
changeset
|
530 |
"_lam_pats_syntax" :: "cases_syn \<Rightarrow> 'a \<Rightarrow> 'b" ("(\<lambda>_)" 10) |
60758 | 531 |
parse_translation \<open> |
52143 | 532 |
let |
533 |
fun fun_tr ctxt [cs] = |
|
534 |
let |
|
535 |
val x = Syntax.free (fst (Name.variant "x" (Term.declare_term_frees cs Name.context))); |
|
536 |
val ft = Case_Translation.case_tr true ctxt [x, cs]; |
|
537 |
in lambda x ft end |
|
69593 | 538 |
in [(\<^syntax_const>\<open>_lam_pats_syntax\<close>, fun_tr)] end |
60758 | 539 |
\<close> |
23526 | 540 |
|
541 |
end |