| author | wenzelm | 
| Fri, 13 Jan 2006 01:13:03 +0100 | |
| changeset 18666 | f37a43cec547 | 
| parent 18585 | 5d379fe2eb74 | 
| child 19736 | d8d0f8f51d69 | 
| permissions | -rw-r--r-- | 
| 14706 | 1 | (* Title : HOL/Library/Zorn.thy | 
| 13652 | 2 | ID : $Id$ | 
| 3 | Author : Jacques D. Fleuriot | |
| 14706 | 4 | Description : Zorn's Lemma -- see Larry Paulson's Zorn.thy in ZF | 
| 5 | *) | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 6 | |
| 14706 | 7 | header {* Zorn's Lemma *}
 | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 8 | |
| 15131 | 9 | theory Zorn | 
| 15140 | 10 | imports Main | 
| 15131 | 11 | begin | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 12 | |
| 14706 | 13 | text{*
 | 
| 14 | The lemma and section numbers refer to an unpublished article | |
| 15 |   \cite{Abrial-Laffitte}.
 | |
| 16 | *} | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 17 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 18 | constdefs | 
| 13652 | 19 | chain :: "'a set set => 'a set set set" | 
| 14706 | 20 |   "chain S  == {F. F \<subseteq> S & (\<forall>x \<in> F. \<forall>y \<in> F. x \<subseteq> y | y \<subseteq> x)}"
 | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 21 | |
| 13652 | 22 | super :: "['a set set,'a set set] => 'a set set set" | 
| 14706 | 23 |   "super S c == {d. d \<in> chain S & c \<subset> d}"
 | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 24 | |
| 13652 | 25 | maxchain :: "'a set set => 'a set set set" | 
| 14706 | 26 |   "maxchain S == {c. c \<in> chain S & super S c = {}}"
 | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 27 | |
| 13652 | 28 | succ :: "['a set set,'a set set] => 'a set set" | 
| 14706 | 29 | "succ S c == | 
| 30 | if c \<notin> chain S | c \<in> maxchain S | |
| 31 | then c else SOME c'. c' \<in> super S c" | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 32 | |
| 14706 | 33 | consts | 
| 34 | TFin :: "'a set set => 'a set set set" | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 35 | |
| 14706 | 36 | inductive "TFin S" | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 37 | intros | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 38 | succI: "x \<in> TFin S ==> succ S x \<in> TFin S" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 39 | Pow_UnionI: "Y \<in> Pow(TFin S) ==> Union(Y) \<in> TFin S" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 40 | monos Pow_mono | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 41 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 42 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 43 | subsection{*Mathematical Preamble*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 44 | |
| 17200 | 45 | lemma Union_lemma0: | 
| 18143 | 46 | "(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C) \<subseteq> A | B \<subseteq> Union(C)" | 
| 17200 | 47 | by blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 48 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 49 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 50 | text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*}
 | 
| 17200 | 51 | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 52 | lemma Abrial_axiom1: "x \<subseteq> succ S x" | 
| 17200 | 53 | apply (unfold succ_def) | 
| 54 | apply (rule split_if [THEN iffD2]) | |
| 55 | apply (auto simp add: super_def maxchain_def psubset_def) | |
| 18585 | 56 | apply (rule contrapos_np, assumption) | 
| 17200 | 57 | apply (rule someI2, blast+) | 
| 58 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 59 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 60 | lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 61 | |
| 14706 | 62 | lemma TFin_induct: | 
| 63 | "[| n \<in> TFin S; | |
| 64 | !!x. [| x \<in> TFin S; P(x) |] ==> P(succ S x); | |
| 65 | !!Y. [| Y \<subseteq> TFin S; Ball Y P |] ==> P(Union Y) |] | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 66 | ==> P(n)" | 
| 17200 | 67 | apply (erule TFin.induct) | 
| 68 | apply blast+ | |
| 69 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 70 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 71 | lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y" | 
| 17200 | 72 | apply (erule subset_trans) | 
| 73 | apply (rule Abrial_axiom1) | |
| 74 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 75 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 76 | text{*Lemma 1 of section 3.1*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 77 | lemma TFin_linear_lemma1: | 
| 14706 | 78 | "[| n \<in> TFin S; m \<in> TFin S; | 
| 79 | \<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 | 80 | |] ==> n \<subseteq> m | succ S m \<subseteq> n" | 
| 17200 | 81 | apply (erule TFin_induct) | 
| 82 | apply (erule_tac [2] Union_lemma0) | |
| 83 | apply (blast del: subsetI intro: succ_trans) | |
| 84 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 85 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 86 | text{* Lemma 2 of section 3.2 *}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 87 | lemma TFin_linear_lemma2: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 88 | "m \<in> TFin S ==> \<forall>n \<in> TFin S. n \<subseteq> m --> n=m | succ S n \<subseteq> m" | 
| 17200 | 89 | apply (erule TFin_induct) | 
| 90 | apply (rule impI [THEN ballI]) | |
| 91 |    txt{*case split using @{text TFin_linear_lemma1}*}
 | |
| 92 | apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], | |
| 93 | assumption+) | |
| 94 | apply (drule_tac x = n in bspec, assumption) | |
| 95 | apply (blast del: subsetI intro: succ_trans, blast) | |
| 96 |   txt{*second induction step*}
 | |
| 97 | apply (rule impI [THEN ballI]) | |
| 98 | apply (rule Union_lemma0 [THEN disjE]) | |
| 99 | apply (rule_tac [3] disjI2) | |
| 100 | prefer 2 apply blast | |
| 101 | apply (rule ballI) | |
| 102 | apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], | |
| 103 | assumption+, auto) | |
| 104 | apply (blast intro!: Abrial_axiom1 [THEN subsetD]) | |
| 105 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 106 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 107 | text{*Re-ordering the premises of Lemma 2*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 108 | lemma TFin_subsetD: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 109 | "[| n \<subseteq> m; m \<in> TFin S; n \<in> TFin S |] ==> n=m | succ S n \<subseteq> m" | 
| 17200 | 110 | by (rule TFin_linear_lemma2 [rule_format]) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 111 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 112 | 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 | 113 | lemma TFin_subset_linear: "[| m \<in> TFin S; n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n" | 
| 17200 | 114 | apply (rule disjE) | 
| 115 | apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) | |
| 116 | apply (assumption+, erule disjI2) | |
| 117 | apply (blast del: subsetI | |
| 118 | intro: subsetI Abrial_axiom1 [THEN subset_trans]) | |
| 119 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 120 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 121 | text{*Lemma 3 of section 3.3*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 122 | lemma eq_succ_upper: "[| n \<in> TFin S; m \<in> TFin S; m = succ S m |] ==> n \<subseteq> m" | 
| 17200 | 123 | apply (erule TFin_induct) | 
| 124 | apply (drule TFin_subsetD) | |
| 125 | apply (assumption+, force, blast) | |
| 126 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 127 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 128 | text{*Property 3.3 of section 3.3*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 129 | lemma equal_succ_Union: "m \<in> TFin S ==> (m = succ S m) = (m = Union(TFin S))" | 
| 17200 | 130 | apply (rule iffI) | 
| 131 | apply (rule Union_upper [THEN equalityI]) | |
| 18143 | 132 | apply assumption | 
| 133 | apply (rule eq_succ_upper [THEN Union_least], assumption+) | |
| 17200 | 134 | apply (erule ssubst) | 
| 135 | apply (rule Abrial_axiom1 [THEN equalityI]) | |
| 136 | apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI) | |
| 137 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 138 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 139 | 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 | 140 | |
| 14706 | 141 | 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 | 142 | the subset relation!*} | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 143 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 144 | lemma empty_set_mem_chain: "({} :: 'a set set) \<in> chain S"
 | 
| 17200 | 145 | by (unfold chain_def) auto | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 146 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 147 | lemma super_subset_chain: "super S c \<subseteq> chain S" | 
| 17200 | 148 | by (unfold super_def) blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 149 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 150 | lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S" | 
| 17200 | 151 | by (unfold maxchain_def) blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 152 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 153 | lemma mem_super_Ex: "c \<in> chain S - maxchain S ==> ? d. d \<in> super S c" | 
| 17200 | 154 | by (unfold super_def maxchain_def) auto | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 155 | |
| 18143 | 156 | lemma select_super: | 
| 157 | "c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c): super S c" | |
| 17200 | 158 | apply (erule mem_super_Ex [THEN exE]) | 
| 159 | apply (rule someI2, auto) | |
| 160 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 161 | |
| 18143 | 162 | lemma select_not_equals: | 
| 163 | "c \<in> chain S - maxchain S ==> (\<some>c'. c': super S c) \<noteq> c" | |
| 17200 | 164 | apply (rule notI) | 
| 165 | apply (drule select_super) | |
| 166 | apply (simp add: super_def psubset_def) | |
| 167 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 168 | |
| 17200 | 169 | lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (\<some>c'. c': super S c)" | 
| 170 | 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 | 171 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 172 | lemma succ_not_equals: "c \<in> chain S - maxchain S ==> succ S c \<noteq> c" | 
| 17200 | 173 | apply (frule succI3) | 
| 174 | apply (simp (no_asm_simp)) | |
| 175 | apply (rule select_not_equals, assumption) | |
| 176 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 177 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 178 | lemma TFin_chain_lemma4: "c \<in> TFin S ==> (c :: 'a set set): chain S" | 
| 17200 | 179 | apply (erule TFin_induct) | 
| 180 | apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]]) | |
| 181 | apply (unfold chain_def) | |
| 182 | apply (rule CollectI, safe) | |
| 183 | apply (drule bspec, assumption) | |
| 184 | apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE], | |
| 185 | blast+) | |
| 186 | done | |
| 14706 | 187 | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 188 | theorem Hausdorff: "\<exists>c. (c :: 'a set set): maxchain S" | 
| 18143 | 189 | apply (rule_tac x = "Union (TFin S)" in exI) | 
| 17200 | 190 | apply (rule classical) | 
| 191 | apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ") | |
| 192 | prefer 2 | |
| 193 | apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) | |
| 194 | apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4]) | |
| 195 | apply (drule DiffI [THEN succ_not_equals], blast+) | |
| 196 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 197 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 198 | |
| 14706 | 199 | 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 | 200 | There Is a Maximal Element*} | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 201 | |
| 14706 | 202 | lemma chain_extend: | 
| 203 | "[| c \<in> chain S; z \<in> S; | |
| 18143 | 204 |         \<forall>x \<in> c. x \<subseteq> (z:: 'a set) |] ==> {z} Un c \<in> chain S"
 | 
| 17200 | 205 | by (unfold chain_def) blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 206 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 207 | lemma chain_Union_upper: "[| c \<in> chain S; x \<in> c |] ==> x \<subseteq> Union(c)" | 
| 17200 | 208 | by (unfold chain_def) auto | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 209 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 210 | lemma chain_ball_Union_upper: "c \<in> chain S ==> \<forall>x \<in> c. x \<subseteq> Union(c)" | 
| 17200 | 211 | by (unfold chain_def) auto | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 212 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 213 | lemma maxchain_Zorn: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 214 | "[| c \<in> maxchain S; u \<in> S; Union(c) \<subseteq> u |] ==> Union(c) = u" | 
| 17200 | 215 | apply (rule ccontr) | 
| 216 | apply (simp add: maxchain_def) | |
| 217 | apply (erule conjE) | |
| 18143 | 218 |   apply (subgoal_tac "({u} Un c) \<in> super S c")
 | 
| 17200 | 219 | apply simp | 
| 220 | apply (unfold super_def psubset_def) | |
| 221 | apply (blast intro: chain_extend dest: chain_Union_upper) | |
| 222 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 223 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 224 | theorem Zorn_Lemma: | 
| 17200 | 225 | "\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z" | 
| 226 | apply (cut_tac Hausdorff maxchain_subset_chain) | |
| 227 | apply (erule exE) | |
| 228 | apply (drule subsetD, assumption) | |
| 229 | apply (drule bspec, assumption) | |
| 18143 | 230 | apply (rule_tac x = "Union(c)" in bexI) | 
| 17200 | 231 | apply (rule ballI, rule impI) | 
| 232 | apply (blast dest!: maxchain_Zorn, assumption) | |
| 233 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 234 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 235 | subsection{*Alternative version of Zorn's Lemma*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 236 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 237 | lemma Zorn_Lemma2: | 
| 17200 | 238 | "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y | 
| 239 | ==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x" | |
| 240 | apply (cut_tac Hausdorff maxchain_subset_chain) | |
| 241 | apply (erule exE) | |
| 242 | apply (drule subsetD, assumption) | |
| 243 | apply (drule bspec, assumption, erule bexE) | |
| 244 | apply (rule_tac x = y in bexI) | |
| 245 | prefer 2 apply assumption | |
| 246 | apply clarify | |
| 247 | apply (rule ccontr) | |
| 248 | apply (frule_tac z = x in chain_extend) | |
| 249 | apply (assumption, blast) | |
| 250 | apply (unfold maxchain_def super_def psubset_def) | |
| 251 | apply (blast elim!: equalityCE) | |
| 252 | done | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 253 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 254 | text{*Various other lemmas*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 255 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 256 | lemma chainD: "[| c \<in> chain S; x \<in> c; y \<in> c |] ==> x \<subseteq> y | y \<subseteq> x" | 
| 17200 | 257 | by (unfold chain_def) blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 258 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 259 | lemma chainD2: "!!(c :: 'a set set). c \<in> chain S ==> c \<subseteq> S" | 
| 17200 | 260 | by (unfold chain_def) blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 261 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 262 | end |