author | Christian Sternagel |
Wed, 29 Aug 2012 12:24:26 +0900 | |
changeset 49084 | e3973567ed4f |
parent 48750 | a151db85a62b |
child 51500 | 01fe31f05aa8 |
permissions | -rw-r--r-- |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
1 |
(* Title: HOL/Library/Zorn.thy |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
2 |
Author: Jacques D. Fleuriot, Tobias Nipkow |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
3 |
|
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
4 |
Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF). |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
5 |
The well-ordering theorem. |
14706 | 6 |
*) |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
7 |
|
14706 | 8 |
header {* Zorn's Lemma *} |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
9 |
|
15131 | 10 |
theory Zorn |
48750 | 11 |
imports Order_Relation |
15131 | 12 |
begin |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
13 |
|
26272 | 14 |
(* Define globally? In Set.thy? *) |
46980 | 15 |
definition chain_subset :: "'a set set \<Rightarrow> bool" ("chain\<^bsub>\<subseteq>\<^esub>") |
16 |
where |
|
17 |
"chain\<^bsub>\<subseteq>\<^esub> C \<equiv> \<forall>A\<in>C.\<forall>B\<in>C. A \<subseteq> B \<or> B \<subseteq> A" |
|
26272 | 18 |
|
14706 | 19 |
text{* |
20 |
The lemma and section numbers refer to an unpublished article |
|
21 |
\cite{Abrial-Laffitte}. |
|
22 |
*} |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
23 |
|
46980 | 24 |
definition chain :: "'a set set \<Rightarrow> 'a set set set" |
25 |
where |
|
26 |
"chain S = {F. F \<subseteq> S \<and> chain\<^bsub>\<subseteq>\<^esub> F}" |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
27 |
|
46980 | 28 |
definition super :: "'a set set \<Rightarrow> 'a set set \<Rightarrow> 'a set set set" |
29 |
where |
|
30 |
"super S c = {d. d \<in> chain S \<and> c \<subset> d}" |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
31 |
|
46980 | 32 |
definition maxchain :: "'a set set \<Rightarrow> 'a set set set" |
33 |
where |
|
34 |
"maxchain S = {c. c \<in> chain S \<and> super S c = {}}" |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
35 |
|
46980 | 36 |
definition succ :: "'a set set \<Rightarrow> 'a set set \<Rightarrow> 'a set set" |
37 |
where |
|
38 |
"succ S c = (if c \<notin> chain S \<or> c \<in> maxchain S then c else SOME c'. c' \<in> super S c)" |
|
39 |
||
40 |
inductive_set TFin :: "'a set set \<Rightarrow> 'a set set set" |
|
41 |
for S :: "'a set set" |
|
42 |
where |
|
43 |
succI: "x \<in> TFin S \<Longrightarrow> succ S x \<in> TFin S" |
|
44 |
| Pow_UnionI: "Y \<in> Pow (TFin S) \<Longrightarrow> \<Union>Y \<in> TFin S" |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
45 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
46 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
47 |
subsection{*Mathematical Preamble*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
48 |
|
17200 | 49 |
lemma Union_lemma0: |
18143 | 50 |
"(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C) \<subseteq> A | B \<subseteq> Union(C)" |
17200 | 51 |
by blast |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
52 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
53 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
54 |
text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*} |
17200 | 55 |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
56 |
lemma Abrial_axiom1: "x \<subseteq> succ S x" |
26806 | 57 |
apply (auto simp add: succ_def super_def maxchain_def) |
18585 | 58 |
apply (rule contrapos_np, assumption) |
46008
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
59 |
apply (rule someI2) |
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
60 |
apply blast+ |
17200 | 61 |
done |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
62 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
63 |
lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
64 |
|
14706 | 65 |
lemma TFin_induct: |
46008
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
66 |
assumes H: "n \<in> TFin S" and |
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
67 |
I: "!!x. x \<in> TFin S ==> P x ==> P (succ S x)" |
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
68 |
"!!Y. Y \<subseteq> TFin S ==> Ball Y P ==> P (Union Y)" |
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
69 |
shows "P n" |
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
70 |
using H by induct (blast intro: I)+ |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
71 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
72 |
lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y" |
17200 | 73 |
apply (erule subset_trans) |
74 |
apply (rule Abrial_axiom1) |
|
75 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
76 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
77 |
text{*Lemma 1 of section 3.1*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
78 |
lemma TFin_linear_lemma1: |
14706 | 79 |
"[| n \<in> TFin S; m \<in> TFin S; |
80 |
\<forall>x \<in> TFin S. x \<subseteq> m --> x = m | succ S x \<subseteq> m |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
81 |
|] ==> n \<subseteq> m | succ S m \<subseteq> n" |
17200 | 82 |
apply (erule TFin_induct) |
83 |
apply (erule_tac [2] Union_lemma0) |
|
84 |
apply (blast del: subsetI intro: succ_trans) |
|
85 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
86 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
87 |
text{* Lemma 2 of section 3.2 *} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
88 |
lemma TFin_linear_lemma2: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
89 |
"m \<in> TFin S ==> \<forall>n \<in> TFin S. n \<subseteq> m --> n=m | succ S n \<subseteq> m" |
17200 | 90 |
apply (erule TFin_induct) |
91 |
apply (rule impI [THEN ballI]) |
|
92 |
txt{*case split using @{text TFin_linear_lemma1}*} |
|
93 |
apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], |
|
94 |
assumption+) |
|
95 |
apply (drule_tac x = n in bspec, assumption) |
|
96 |
apply (blast del: subsetI intro: succ_trans, blast) |
|
97 |
txt{*second induction step*} |
|
98 |
apply (rule impI [THEN ballI]) |
|
99 |
apply (rule Union_lemma0 [THEN disjE]) |
|
100 |
apply (rule_tac [3] disjI2) |
|
101 |
prefer 2 apply blast |
|
102 |
apply (rule ballI) |
|
103 |
apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], |
|
104 |
assumption+, auto) |
|
105 |
apply (blast intro!: Abrial_axiom1 [THEN subsetD]) |
|
106 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
107 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
108 |
text{*Re-ordering the premises of Lemma 2*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
109 |
lemma TFin_subsetD: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
110 |
"[| n \<subseteq> m; m \<in> TFin S; n \<in> TFin S |] ==> n=m | succ S n \<subseteq> m" |
17200 | 111 |
by (rule TFin_linear_lemma2 [rule_format]) |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
112 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
113 |
text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
114 |
lemma TFin_subset_linear: "[| m \<in> TFin S; n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n" |
17200 | 115 |
apply (rule disjE) |
116 |
apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) |
|
117 |
apply (assumption+, erule disjI2) |
|
118 |
apply (blast del: subsetI |
|
119 |
intro: subsetI Abrial_axiom1 [THEN subset_trans]) |
|
120 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
121 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
122 |
text{*Lemma 3 of section 3.3*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
123 |
lemma eq_succ_upper: "[| n \<in> TFin S; m \<in> TFin S; m = succ S m |] ==> n \<subseteq> m" |
17200 | 124 |
apply (erule TFin_induct) |
125 |
apply (drule TFin_subsetD) |
|
126 |
apply (assumption+, force, blast) |
|
127 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
128 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
129 |
text{*Property 3.3 of section 3.3*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
130 |
lemma equal_succ_Union: "m \<in> TFin S ==> (m = succ S m) = (m = Union(TFin S))" |
17200 | 131 |
apply (rule iffI) |
132 |
apply (rule Union_upper [THEN equalityI]) |
|
18143 | 133 |
apply assumption |
134 |
apply (rule eq_succ_upper [THEN Union_least], assumption+) |
|
17200 | 135 |
apply (erule ssubst) |
136 |
apply (rule Abrial_axiom1 [THEN equalityI]) |
|
137 |
apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI) |
|
138 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
139 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
140 |
subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
141 |
|
14706 | 142 |
text{*NB: We assume the partial ordering is @{text "\<subseteq>"}, |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
143 |
the subset relation!*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
144 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
145 |
lemma empty_set_mem_chain: "({} :: 'a set set) \<in> chain S" |
48750 | 146 |
by (unfold chain_def chain_subset_def) simp |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
147 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
148 |
lemma super_subset_chain: "super S c \<subseteq> chain S" |
17200 | 149 |
by (unfold super_def) blast |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
150 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
151 |
lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S" |
17200 | 152 |
by (unfold maxchain_def) blast |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
153 |
|
26191 | 154 |
lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> EX d. d \<in> super S c" |
48750 | 155 |
by (unfold super_def maxchain_def) simp |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
156 |
|
18143 | 157 |
lemma select_super: |
158 |
"c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c): super S c" |
|
17200 | 159 |
apply (erule mem_super_Ex [THEN exE]) |
46008
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
160 |
apply (rule someI2) |
48750 | 161 |
apply simp+ |
17200 | 162 |
done |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
163 |
|
18143 | 164 |
lemma select_not_equals: |
165 |
"c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c) \<noteq> c" |
|
17200 | 166 |
apply (rule notI) |
167 |
apply (drule select_super) |
|
26806 | 168 |
apply (simp add: super_def less_le) |
17200 | 169 |
done |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
170 |
|
17200 | 171 |
lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (\<some>c'. c': super S c)" |
172 |
by (unfold succ_def) (blast intro!: if_not_P) |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
173 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
174 |
lemma succ_not_equals: "c \<in> chain S - maxchain S ==> succ S c \<noteq> c" |
17200 | 175 |
apply (frule succI3) |
176 |
apply (simp (no_asm_simp)) |
|
177 |
apply (rule select_not_equals, assumption) |
|
178 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
179 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
180 |
lemma TFin_chain_lemma4: "c \<in> TFin S ==> (c :: 'a set set): chain S" |
17200 | 181 |
apply (erule TFin_induct) |
182 |
apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]]) |
|
26272 | 183 |
apply (unfold chain_def chain_subset_def) |
17200 | 184 |
apply (rule CollectI, safe) |
185 |
apply (drule bspec, assumption) |
|
46008
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
186 |
apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE]) |
c296c75f4cf4
reverted some changes for set->predicate transition, according to "hg log -u berghofe -r Isabelle2007:Isabelle2008";
wenzelm
parents:
44890
diff
changeset
|
187 |
apply blast+ |
17200 | 188 |
done |
14706 | 189 |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
190 |
theorem Hausdorff: "\<exists>c. (c :: 'a set set): maxchain S" |
18143 | 191 |
apply (rule_tac x = "Union (TFin S)" in exI) |
17200 | 192 |
apply (rule classical) |
193 |
apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ") |
|
194 |
prefer 2 |
|
195 |
apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) |
|
196 |
apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4]) |
|
197 |
apply (drule DiffI [THEN succ_not_equals], blast+) |
|
198 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
199 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
200 |
|
14706 | 201 |
subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
202 |
There Is a Maximal Element*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
203 |
|
14706 | 204 |
lemma chain_extend: |
26272 | 205 |
"[| c \<in> chain S; z \<in> S; \<forall>x \<in> c. x \<subseteq> (z:: 'a set) |] ==> {z} Un c \<in> chain S" |
206 |
by (unfold chain_def chain_subset_def) blast |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
207 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
208 |
lemma chain_Union_upper: "[| c \<in> chain S; x \<in> c |] ==> x \<subseteq> Union(c)" |
26272 | 209 |
by auto |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
210 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
211 |
lemma chain_ball_Union_upper: "c \<in> chain S ==> \<forall>x \<in> c. x \<subseteq> Union(c)" |
26272 | 212 |
by auto |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
213 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
214 |
lemma maxchain_Zorn: |
26272 | 215 |
"[| c \<in> maxchain S; u \<in> S; Union(c) \<subseteq> u |] ==> Union(c) = u" |
216 |
apply (rule ccontr) |
|
217 |
apply (simp add: maxchain_def) |
|
218 |
apply (erule conjE) |
|
219 |
apply (subgoal_tac "({u} Un c) \<in> super S c") |
|
220 |
apply simp |
|
26806 | 221 |
apply (unfold super_def less_le) |
26272 | 222 |
apply (blast intro: chain_extend dest: chain_Union_upper) |
223 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
224 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
225 |
theorem Zorn_Lemma: |
26272 | 226 |
"\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z" |
227 |
apply (cut_tac Hausdorff maxchain_subset_chain) |
|
228 |
apply (erule exE) |
|
229 |
apply (drule subsetD, assumption) |
|
230 |
apply (drule bspec, assumption) |
|
231 |
apply (rule_tac x = "Union(c)" in bexI) |
|
232 |
apply (rule ballI, rule impI) |
|
233 |
apply (blast dest!: maxchain_Zorn, assumption) |
|
234 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
235 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
236 |
subsection{*Alternative version of Zorn's Lemma*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
237 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
238 |
lemma Zorn_Lemma2: |
17200 | 239 |
"\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y |
240 |
==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x" |
|
26272 | 241 |
apply (cut_tac Hausdorff maxchain_subset_chain) |
242 |
apply (erule exE) |
|
243 |
apply (drule subsetD, assumption) |
|
244 |
apply (drule bspec, assumption, erule bexE) |
|
245 |
apply (rule_tac x = y in bexI) |
|
246 |
prefer 2 apply assumption |
|
247 |
apply clarify |
|
248 |
apply (rule ccontr) |
|
249 |
apply (frule_tac z = x in chain_extend) |
|
250 |
apply (assumption, blast) |
|
26806 | 251 |
apply (unfold maxchain_def super_def less_le) |
26272 | 252 |
apply (blast elim!: equalityCE) |
253 |
done |
|
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
254 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
255 |
text{*Various other lemmas*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
256 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
257 |
lemma chainD: "[| c \<in> chain S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x" |
26272 | 258 |
by (unfold chain_def chain_subset_def) blast |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
259 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
260 |
lemma chainD2: "!!(c :: 'a set set). c \<in> chain S ==> c \<subseteq> S" |
26272 | 261 |
by (unfold chain_def) blast |
26191 | 262 |
|
263 |
||
264 |
(* Define globally? In Relation.thy? *) |
|
265 |
definition Chain :: "('a*'a)set \<Rightarrow> 'a set set" where |
|
266 |
"Chain r \<equiv> {A. \<forall>a\<in>A.\<forall>b\<in>A. (a,b) : r \<or> (b,a) \<in> r}" |
|
267 |
||
268 |
lemma mono_Chain: "r \<subseteq> s \<Longrightarrow> Chain r \<subseteq> Chain s" |
|
269 |
unfolding Chain_def by blast |
|
270 |
||
271 |
text{* Zorn's lemma for partial orders: *} |
|
272 |
||
273 |
lemma Zorns_po_lemma: |
|
274 |
assumes po: "Partial_order r" and u: "\<forall>C\<in>Chain r. \<exists>u\<in>Field r. \<forall>a\<in>C. (a,u):r" |
|
275 |
shows "\<exists>m\<in>Field r. \<forall>a\<in>Field r. (m,a):r \<longrightarrow> a=m" |
|
276 |
proof- |
|
26295 | 277 |
have "Preorder r" using po by(simp add:partial_order_on_def) |
26191 | 278 |
--{* Mirror r in the set of subsets below (wrt r) elements of A*} |
279 |
let ?B = "%x. r^-1 `` {x}" let ?S = "?B ` Field r" |
|
280 |
have "\<forall>C \<in> chain ?S. EX U:?S. ALL A:C. A\<subseteq>U" |
|
26272 | 281 |
proof (auto simp:chain_def chain_subset_def) |
26191 | 282 |
fix C assume 1: "C \<subseteq> ?S" and 2: "\<forall>A\<in>C.\<forall>B\<in>C. A\<subseteq>B | B\<subseteq>A" |
283 |
let ?A = "{x\<in>Field r. \<exists>M\<in>C. M = ?B x}" |
|
284 |
have "C = ?B ` ?A" using 1 by(auto simp: image_def) |
|
285 |
have "?A\<in>Chain r" |
|
286 |
proof (simp add:Chain_def, intro allI impI, elim conjE) |
|
287 |
fix a b |
|
288 |
assume "a \<in> Field r" "?B a \<in> C" "b \<in> Field r" "?B b \<in> C" |
|
48750 | 289 |
hence "?B a \<subseteq> ?B b \<or> ?B b \<subseteq> ?B a" using 2 by simp |
26191 | 290 |
thus "(a, b) \<in> r \<or> (b, a) \<in> r" using `Preorder r` `a:Field r` `b:Field r` |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
291 |
by (simp add:subset_Image1_Image1_iff) |
26191 | 292 |
qed |
293 |
then obtain u where uA: "u:Field r" "\<forall>a\<in>?A. (a,u) : r" using u by auto |
|
294 |
have "\<forall>A\<in>C. A \<subseteq> r^-1 `` {u}" (is "?P u") |
|
295 |
proof auto |
|
296 |
fix a B assume aB: "B:C" "a:B" |
|
297 |
with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto |
|
298 |
thus "(a,u) : r" using uA aB `Preorder r` |
|
48750 | 299 |
by (simp add: preorder_on_def refl_on_def) (rule transD, blast+) |
26191 | 300 |
qed |
301 |
thus "EX u:Field r. ?P u" using `u:Field r` by blast |
|
302 |
qed |
|
303 |
from Zorn_Lemma2[OF this] |
|
304 |
obtain m B where "m:Field r" "B = r^-1 `` {m}" |
|
305 |
"\<forall>x\<in>Field r. B \<subseteq> r^-1 `` {x} \<longrightarrow> B = r^-1 `` {x}" |
|
27064 | 306 |
by auto |
26191 | 307 |
hence "\<forall>a\<in>Field r. (m, a) \<in> r \<longrightarrow> a = m" using po `Preorder r` `m:Field r` |
308 |
by(auto simp:subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff) |
|
309 |
thus ?thesis using `m:Field r` by blast |
|
310 |
qed |
|
311 |
||
312 |
(* The initial segment of a relation appears generally useful. |
|
313 |
Move to Relation.thy? |
|
314 |
Definition correct/most general? |
|
315 |
Naming? |
|
316 |
*) |
|
317 |
definition init_seg_of :: "(('a*'a)set * ('a*'a)set)set" where |
|
318 |
"init_seg_of == {(r,s). r \<subseteq> s \<and> (\<forall>a b c. (a,b):s \<and> (b,c):r \<longrightarrow> (a,b):r)}" |
|
319 |
||
320 |
abbreviation initialSegmentOf :: "('a*'a)set \<Rightarrow> ('a*'a)set \<Rightarrow> bool" |
|
321 |
(infix "initial'_segment'_of" 55) where |
|
322 |
"r initial_segment_of s == (r,s):init_seg_of" |
|
323 |
||
30198 | 324 |
lemma refl_on_init_seg_of[simp]: "r initial_segment_of r" |
26191 | 325 |
by(simp add:init_seg_of_def) |
326 |
||
327 |
lemma trans_init_seg_of: |
|
328 |
"r initial_segment_of s \<Longrightarrow> s initial_segment_of t \<Longrightarrow> r initial_segment_of t" |
|
48750 | 329 |
by (simp (no_asm_use) add: init_seg_of_def) (metis (no_types) in_mono order_trans) |
26191 | 330 |
|
331 |
lemma antisym_init_seg_of: |
|
332 |
"r initial_segment_of s \<Longrightarrow> s initial_segment_of r \<Longrightarrow> r=s" |
|
35175 | 333 |
unfolding init_seg_of_def by safe |
26191 | 334 |
|
335 |
lemma Chain_init_seg_of_Union: |
|
336 |
"R \<in> Chain init_seg_of \<Longrightarrow> r\<in>R \<Longrightarrow> r initial_segment_of \<Union>R" |
|
48750 | 337 |
by(simp add: init_seg_of_def Chain_def Ball_def) blast |
26191 | 338 |
|
26272 | 339 |
lemma chain_subset_trans_Union: |
340 |
"chain\<^bsub>\<subseteq>\<^esub> R \<Longrightarrow> \<forall>r\<in>R. trans r \<Longrightarrow> trans(\<Union>R)" |
|
48750 | 341 |
apply(simp add:chain_subset_def) |
26191 | 342 |
apply(simp (no_asm_use) add:trans_def) |
48750 | 343 |
by (metis subsetD) |
26191 | 344 |
|
26272 | 345 |
lemma chain_subset_antisym_Union: |
346 |
"chain\<^bsub>\<subseteq>\<^esub> R \<Longrightarrow> \<forall>r\<in>R. antisym r \<Longrightarrow> antisym(\<Union>R)" |
|
48750 | 347 |
by (simp add:chain_subset_def antisym_def) (metis subsetD) |
26191 | 348 |
|
26272 | 349 |
lemma chain_subset_Total_Union: |
350 |
assumes "chain\<^bsub>\<subseteq>\<^esub> R" "\<forall>r\<in>R. Total r" |
|
26191 | 351 |
shows "Total (\<Union>R)" |
26295 | 352 |
proof (simp add: total_on_def Ball_def, auto del:disjCI) |
26191 | 353 |
fix r s a b assume A: "r:R" "s:R" "a:Field r" "b:Field s" "a\<noteq>b" |
26272 | 354 |
from `chain\<^bsub>\<subseteq>\<^esub> R` `r:R` `s:R` have "r\<subseteq>s \<or> s\<subseteq>r" |
355 |
by(simp add:chain_subset_def) |
|
26191 | 356 |
thus "(\<exists>r\<in>R. (a,b) \<in> r) \<or> (\<exists>r\<in>R. (b,a) \<in> r)" |
357 |
proof |
|
358 |
assume "r\<subseteq>s" hence "(a,b):s \<or> (b,a):s" using assms(2) A |
|
26295 | 359 |
by(simp add:total_on_def)(metis mono_Field subsetD) |
26191 | 360 |
thus ?thesis using `s:R` by blast |
361 |
next |
|
362 |
assume "s\<subseteq>r" hence "(a,b):r \<or> (b,a):r" using assms(2) A |
|
26295 | 363 |
by(simp add:total_on_def)(metis mono_Field subsetD) |
26191 | 364 |
thus ?thesis using `r:R` by blast |
365 |
qed |
|
366 |
qed |
|
367 |
||
368 |
lemma wf_Union_wf_init_segs: |
|
369 |
assumes "R \<in> Chain init_seg_of" and "\<forall>r\<in>R. wf r" shows "wf(\<Union>R)" |
|
370 |
proof(simp add:wf_iff_no_infinite_down_chain, rule ccontr, auto) |
|
371 |
fix f assume 1: "\<forall>i. \<exists>r\<in>R. (f(Suc i), f i) \<in> r" |
|
372 |
then obtain r where "r:R" and "(f(Suc 0), f 0) : r" by auto |
|
373 |
{ fix i have "(f(Suc i), f i) \<in> r" |
|
374 |
proof(induct i) |
|
375 |
case 0 show ?case by fact |
|
376 |
next |
|
377 |
case (Suc i) |
|
378 |
moreover obtain s where "s\<in>R" and "(f(Suc(Suc i)), f(Suc i)) \<in> s" |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
379 |
using 1 by auto |
26191 | 380 |
moreover hence "s initial_segment_of r \<or> r initial_segment_of s" |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
381 |
using assms(1) `r:R` by(simp add: Chain_def) |
26191 | 382 |
ultimately show ?case by(simp add:init_seg_of_def) blast |
383 |
qed |
|
384 |
} |
|
385 |
thus False using assms(2) `r:R` |
|
386 |
by(simp add:wf_iff_no_infinite_down_chain) blast |
|
387 |
qed |
|
388 |
||
27476 | 389 |
lemma initial_segment_of_Diff: |
390 |
"p initial_segment_of q \<Longrightarrow> p - s initial_segment_of q - s" |
|
391 |
unfolding init_seg_of_def by blast |
|
392 |
||
26191 | 393 |
lemma Chain_inits_DiffI: |
394 |
"R \<in> Chain init_seg_of \<Longrightarrow> {r - s |r. r \<in> R} \<in> Chain init_seg_of" |
|
27476 | 395 |
unfolding Chain_def by (blast intro: initial_segment_of_Diff) |
26191 | 396 |
|
26272 | 397 |
theorem well_ordering: "\<exists>r::('a*'a)set. Well_order r \<and> Field r = UNIV" |
26191 | 398 |
proof- |
399 |
-- {*The initial segment relation on well-orders: *} |
|
400 |
let ?WO = "{r::('a*'a)set. Well_order r}" |
|
401 |
def I \<equiv> "init_seg_of \<inter> ?WO \<times> ?WO" |
|
48750 | 402 |
have I_init: "I \<subseteq> init_seg_of" by(simp add: I_def) |
26272 | 403 |
hence subch: "!!R. R : Chain I \<Longrightarrow> chain\<^bsub>\<subseteq>\<^esub> R" |
404 |
by(auto simp:init_seg_of_def chain_subset_def Chain_def) |
|
26191 | 405 |
have Chain_wo: "!!R r. R \<in> Chain I \<Longrightarrow> r \<in> R \<Longrightarrow> Well_order r" |
406 |
by(simp add:Chain_def I_def) blast |
|
407 |
have FI: "Field I = ?WO" by(auto simp add:I_def init_seg_of_def Field_def) |
|
408 |
hence 0: "Partial_order I" |
|
30198 | 409 |
by(auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_on_def trans_def I_def elim!: trans_init_seg_of) |
26191 | 410 |
-- {*I-chains have upper bounds in ?WO wrt I: their Union*} |
411 |
{ fix R assume "R \<in> Chain I" |
|
412 |
hence Ris: "R \<in> Chain init_seg_of" using mono_Chain[OF I_init] by blast |
|
26272 | 413 |
have subch: "chain\<^bsub>\<subseteq>\<^esub> R" using `R : Chain I` I_init |
414 |
by(auto simp:init_seg_of_def chain_subset_def Chain_def) |
|
26191 | 415 |
have "\<forall>r\<in>R. Refl r" "\<forall>r\<in>R. trans r" "\<forall>r\<in>R. antisym r" "\<forall>r\<in>R. Total r" |
416 |
"\<forall>r\<in>R. wf(r-Id)" |
|
26295 | 417 |
using Chain_wo[OF `R \<in> Chain I`] by(simp_all add:order_on_defs) |
30198 | 418 |
have "Refl (\<Union>R)" using `\<forall>r\<in>R. Refl r` by(auto simp:refl_on_def) |
26191 | 419 |
moreover have "trans (\<Union>R)" |
26272 | 420 |
by(rule chain_subset_trans_Union[OF subch `\<forall>r\<in>R. trans r`]) |
26191 | 421 |
moreover have "antisym(\<Union>R)" |
26272 | 422 |
by(rule chain_subset_antisym_Union[OF subch `\<forall>r\<in>R. antisym r`]) |
26191 | 423 |
moreover have "Total (\<Union>R)" |
26272 | 424 |
by(rule chain_subset_Total_Union[OF subch `\<forall>r\<in>R. Total r`]) |
26191 | 425 |
moreover have "wf((\<Union>R)-Id)" |
426 |
proof- |
|
427 |
have "(\<Union>R)-Id = \<Union>{r-Id|r. r \<in> R}" by blast |
|
428 |
with `\<forall>r\<in>R. wf(r-Id)` wf_Union_wf_init_segs[OF Chain_inits_DiffI[OF Ris]] |
|
429 |
show ?thesis by (simp (no_asm_simp)) blast |
|
430 |
qed |
|
26295 | 431 |
ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs) |
26191 | 432 |
moreover have "\<forall>r \<in> R. r initial_segment_of \<Union>R" using Ris |
433 |
by(simp add: Chain_init_seg_of_Union) |
|
434 |
ultimately have "\<Union>R : ?WO \<and> (\<forall>r\<in>R. (r,\<Union>R) : I)" |
|
435 |
using mono_Chain[OF I_init] `R \<in> Chain I` |
|
48750 | 436 |
by(simp (no_asm) add:I_def del:Field_Union)(metis Chain_wo) |
26191 | 437 |
} |
438 |
hence 1: "\<forall>R \<in> Chain I. \<exists>u\<in>Field I. \<forall>r\<in>R. (r,u) : I" by (subst FI) blast |
|
439 |
--{*Zorn's Lemma yields a maximal well-order m:*} |
|
440 |
then obtain m::"('a*'a)set" where "Well_order m" and |
|
441 |
max: "\<forall>r. Well_order r \<and> (m,r):I \<longrightarrow> r=m" |
|
442 |
using Zorns_po_lemma[OF 0 1] by (auto simp:FI) |
|
443 |
--{*Now show by contradiction that m covers the whole type:*} |
|
444 |
{ fix x::'a assume "x \<notin> Field m" |
|
445 |
--{*We assume that x is not covered and extend m at the top with x*} |
|
446 |
have "m \<noteq> {}" |
|
447 |
proof |
|
448 |
assume "m={}" |
|
449 |
moreover have "Well_order {(x,x)}" |
|
46752
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
haftmann
parents:
46008
diff
changeset
|
450 |
by(simp add:order_on_defs refl_on_def trans_def antisym_def total_on_def Field_def Domain_unfold Domain_converse [symmetric]) |
26191 | 451 |
ultimately show False using max |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
452 |
by (auto simp:I_def init_seg_of_def simp del:Field_insert) |
26191 | 453 |
qed |
454 |
hence "Field m \<noteq> {}" by(auto simp:Field_def) |
|
26295 | 455 |
moreover have "wf(m-Id)" using `Well_order m` |
456 |
by(simp add:well_order_on_def) |
|
26191 | 457 |
--{*The extension of m by x:*} |
458 |
let ?s = "{(a,x)|a. a : Field m}" let ?m = "insert (x,x) m Un ?s" |
|
459 |
have Fm: "Field ?m = insert x (Field m)" |
|
460 |
apply(simp add:Field_insert Field_Un) |
|
461 |
unfolding Field_def by auto |
|
462 |
have "Refl m" "trans m" "antisym m" "Total m" "wf(m-Id)" |
|
26295 | 463 |
using `Well_order m` by(simp_all add:order_on_defs) |
26191 | 464 |
--{*We show that the extension is a well-order*} |
30198 | 465 |
have "Refl ?m" using `Refl m` Fm by(auto simp:refl_on_def) |
26191 | 466 |
moreover have "trans ?m" using `trans m` `x \<notin> Field m` |
46752
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
haftmann
parents:
46008
diff
changeset
|
467 |
unfolding trans_def Field_def Domain_unfold Domain_converse [symmetric] by blast |
26191 | 468 |
moreover have "antisym ?m" using `antisym m` `x \<notin> Field m` |
46752
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
haftmann
parents:
46008
diff
changeset
|
469 |
unfolding antisym_def Field_def Domain_unfold Domain_converse [symmetric] by blast |
26295 | 470 |
moreover have "Total ?m" using `Total m` Fm by(auto simp: total_on_def) |
26191 | 471 |
moreover have "wf(?m-Id)" |
472 |
proof- |
|
473 |
have "wf ?s" using `x \<notin> Field m` |
|
48750 | 474 |
by(simp add:wf_eq_minimal Field_def Domain_unfold Domain_converse [symmetric]) metis |
26191 | 475 |
thus ?thesis using `wf(m-Id)` `x \<notin> Field m` |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30663
diff
changeset
|
476 |
wf_subset[OF `wf ?s` Diff_subset] |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
35175
diff
changeset
|
477 |
by (fastforce intro!: wf_Un simp add: Un_Diff Field_def) |
26191 | 478 |
qed |
26295 | 479 |
ultimately have "Well_order ?m" by(simp add:order_on_defs) |
26191 | 480 |
--{*We show that the extension is above m*} |
481 |
moreover hence "(m,?m) : I" using `Well_order m` `x \<notin> Field m` |
|
46752
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
haftmann
parents:
46008
diff
changeset
|
482 |
by(fastforce simp:I_def init_seg_of_def Field_def Domain_unfold Domain_converse [symmetric]) |
26191 | 483 |
ultimately |
484 |
--{*This contradicts maximality of m:*} |
|
485 |
have False using max `x \<notin> Field m` unfolding Field_def by blast |
|
486 |
} |
|
487 |
hence "Field m = UNIV" by auto |
|
26272 | 488 |
moreover with `Well_order m` have "Well_order m" by simp |
489 |
ultimately show ?thesis by blast |
|
490 |
qed |
|
491 |
||
26295 | 492 |
corollary well_order_on: "\<exists>r::('a*'a)set. well_order_on A r" |
26272 | 493 |
proof - |
494 |
obtain r::"('a*'a)set" where wo: "Well_order r" and univ: "Field r = UNIV" |
|
495 |
using well_ordering[where 'a = "'a"] by blast |
|
496 |
let ?r = "{(x,y). x:A & y:A & (x,y):r}" |
|
497 |
have 1: "Field ?r = A" using wo univ |
|
46752
e9e7209eb375
more fundamental pred-to-set conversions, particularly by means of inductive_set; associated consolidation of some theorem names (c.f. NEWS)
haftmann
parents:
46008
diff
changeset
|
498 |
by(fastforce simp: Field_def Domain_unfold Domain_converse [symmetric] order_on_defs refl_on_def) |
26272 | 499 |
have "Refl r" "trans r" "antisym r" "Total r" "wf(r-Id)" |
26295 | 500 |
using `Well_order r` by(simp_all add:order_on_defs) |
30198 | 501 |
have "Refl ?r" using `Refl r` by(auto simp:refl_on_def 1 univ) |
26272 | 502 |
moreover have "trans ?r" using `trans r` |
503 |
unfolding trans_def by blast |
|
504 |
moreover have "antisym ?r" using `antisym r` |
|
505 |
unfolding antisym_def by blast |
|
26295 | 506 |
moreover have "Total ?r" using `Total r` by(simp add:total_on_def 1 univ) |
26272 | 507 |
moreover have "wf(?r - Id)" by(rule wf_subset[OF `wf(r-Id)`]) blast |
26295 | 508 |
ultimately have "Well_order ?r" by(simp add:order_on_defs) |
509 |
with 1 show ?thesis by metis |
|
26191 | 510 |
qed |
511 |
||
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
512 |
end |