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