| author | wenzelm | 
| Thu, 02 Jun 2005 18:29:47 +0200 | |
| changeset 16187 | 6ec757011ad6 | 
| parent 15140 | 322485b816ac | 
| child 17200 | 3a4d03d1a31b | 
| 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 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 45 | lemma Union_lemma0: "(\<forall>x \<in> C. x \<subseteq> A | B \<subseteq> x) ==> Union(C)<=A | B \<subseteq> Union(C)" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 46 | by blast | 
| 
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 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 49 | text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 50 | lemma Abrial_axiom1: "x \<subseteq> succ S x" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 51 | apply (unfold succ_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 52 | apply (rule split_if [THEN iffD2]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 53 | apply (auto simp add: super_def maxchain_def psubset_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 54 | apply (rule swap, assumption) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 55 | apply (rule someI2, blast+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 56 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 57 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 58 | lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 59 | |
| 14706 | 60 | lemma TFin_induct: | 
| 61 | "[| n \<in> TFin S; | |
| 62 | !!x. [| x \<in> TFin S; P(x) |] ==> P(succ S x); | |
| 63 | !!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 | 64 | ==> P(n)" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 65 | apply (erule TFin.induct, blast+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 66 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 67 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 68 | lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y" | 
| 14706 | 69 | apply (erule subset_trans) | 
| 70 | apply (rule Abrial_axiom1) | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 71 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 72 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 73 | text{*Lemma 1 of section 3.1*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 74 | lemma TFin_linear_lemma1: | 
| 14706 | 75 | "[| n \<in> TFin S; m \<in> TFin S; | 
| 76 | \<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 | 77 | |] ==> n \<subseteq> m | succ S m \<subseteq> n" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 78 | apply (erule TFin_induct) | 
| 14706 | 79 | apply (erule_tac [2] Union_lemma0) (*or just blast*) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 80 | apply (blast del: subsetI intro: succ_trans) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 81 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 82 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 83 | text{* Lemma 2 of section 3.2 *}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 84 | lemma TFin_linear_lemma2: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 85 | "m \<in> TFin S ==> \<forall>n \<in> TFin S. n \<subseteq> m --> n=m | succ S n \<subseteq> m" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 86 | apply (erule TFin_induct) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 87 | apply (rule impI [THEN ballI]) | 
| 14706 | 88 | txt{*case split using @{text TFin_linear_lemma1}*}
 | 
| 89 | apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 90 | assumption+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 91 | apply (drule_tac x = n in bspec, assumption) | 
| 14706 | 92 | apply (blast del: subsetI intro: succ_trans, blast) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 93 | txt{*second induction step*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 94 | apply (rule impI [THEN ballI]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 95 | apply (rule Union_lemma0 [THEN disjE]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 96 | apply (rule_tac [3] disjI2) | 
| 14706 | 97 | prefer 2 apply blast | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 98 | apply (rule ballI) | 
| 14706 | 99 | apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], | 
| 100 | assumption+, auto) | |
| 101 | apply (blast intro!: Abrial_axiom1 [THEN subsetD]) | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 102 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 103 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 104 | text{*Re-ordering the premises of Lemma 2*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 105 | lemma TFin_subsetD: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 106 | "[| n \<subseteq> m; m \<in> TFin S; n \<in> TFin S |] ==> n=m | succ S n \<subseteq> m" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 107 | apply (rule TFin_linear_lemma2 [rule_format]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 108 | apply (assumption+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 109 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 110 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 111 | 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 | 112 | lemma TFin_subset_linear: "[| m \<in> TFin S; n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n" | 
| 14706 | 113 | apply (rule disjE) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 114 | apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 115 | apply (assumption+, erule disjI2) | 
| 14706 | 116 | apply (blast del: subsetI | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 117 | intro: subsetI Abrial_axiom1 [THEN subset_trans]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 118 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 119 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 120 | text{*Lemma 3 of section 3.3*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 121 | lemma eq_succ_upper: "[| n \<in> TFin S; m \<in> TFin S; m = succ S m |] ==> n \<subseteq> m" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 122 | apply (erule TFin_induct) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 123 | apply (drule TFin_subsetD) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 124 | apply (assumption+, force, blast) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 125 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 126 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 127 | text{*Property 3.3 of section 3.3*}
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 128 | lemma equal_succ_Union: "m \<in> TFin S ==> (m = succ S m) = (m = Union(TFin S))" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 129 | apply (rule iffI) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 130 | apply (rule Union_upper [THEN equalityI]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 131 | apply (rule_tac [2] eq_succ_upper [THEN Union_least]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 132 | apply (assumption+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 133 | apply (erule ssubst) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 134 | apply (rule Abrial_axiom1 [THEN equalityI]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 135 | apply (blast del: subsetI | 
| 14706 | 136 | intro: subsetI TFin_UnionI TFin.succI) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 137 | done | 
| 
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"
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 145 | by (unfold chain_def, auto) | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 148 | by (unfold super_def, fast) | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 151 | by (unfold maxchain_def, fast) | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 154 | by (unfold super_def maxchain_def, auto) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 155 | |
| 14706 | 156 | lemma select_super: "c \<in> chain S - maxchain S ==> | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 157 | (@c'. c': super S c): super S c" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 158 | apply (erule mem_super_Ex [THEN exE]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 159 | apply (rule someI2, auto) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 160 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 161 | |
| 14706 | 162 | lemma select_not_equals: "c \<in> chain S - maxchain S ==> | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 163 | (@c'. c': super S c) \<noteq> c" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 164 | apply (rule notI) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 165 | apply (drule select_super) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 166 | apply (simp add: super_def psubset_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 167 | done | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 168 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 169 | lemma succI3: "c \<in> chain S - maxchain S ==> succ S c = (@c'. c': super S c)" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 170 | apply (unfold succ_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 171 | apply (fast intro!: if_not_P) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 172 | done | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 175 | apply (frule succI3) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 176 | apply (simp (no_asm_simp)) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 177 | apply (rule select_not_equals, assumption) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 178 | done | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 181 | apply (erule TFin_induct) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 182 | apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 183 | apply (unfold chain_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 184 | apply (rule CollectI, safe) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 185 | apply (drule bspec, assumption) | 
| 14706 | 186 | apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE], | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 187 | blast+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 191 | apply (rule_tac x = "Union (TFin S) " in exI) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 192 | apply (rule classical) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 193 | apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ") | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 194 | prefer 2 | 
| 14706 | 195 | apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 196 | apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4]) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 197 | apply (drule DiffI [THEN succ_not_equals], blast+) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 198 | done | 
| 
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: | 
| 205 | "[| c \<in> chain S; z \<in> S; | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 206 |         \<forall>x \<in> c. x<=(z:: 'a set) |] ==> {z} Un c \<in> chain S"
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 207 | by (unfold chain_def, blast) | 
| 
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)" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 210 | by (unfold chain_def, auto) | 
| 
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)" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 213 | by (unfold chain_def, auto) | 
| 
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: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 216 | "[| c \<in> maxchain S; u \<in> S; Union(c) \<subseteq> u |] ==> Union(c) = u" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 217 | apply (rule ccontr) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 218 | apply (simp add: maxchain_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 219 | apply (erule conjE) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 220 | apply (subgoal_tac " ({u} Un c) \<in> super S c")
 | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 221 | apply simp | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 222 | apply (unfold super_def psubset_def) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 223 | apply (blast intro: chain_extend dest: chain_Union_upper) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 224 | done | 
| 
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: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 227 | "\<forall>c \<in> chain S. Union(c): S ==> \<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z --> y = z" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 228 | apply (cut_tac Hausdorff maxchain_subset_chain) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 229 | apply (erule exE) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 230 | apply (drule subsetD, assumption) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 231 | apply (drule bspec, assumption) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 232 | apply (rule_tac x = "Union (c) " in bexI) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 233 | apply (rule ballI, rule impI) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 234 | apply (blast dest!: maxchain_Zorn, assumption) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 235 | done | 
| 
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: | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 240 | "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>x \<in> c. x \<subseteq> y | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 241 | ==> \<exists>y \<in> S. \<forall>x \<in> S. (y :: 'a set) \<subseteq> x --> y = x" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 242 | apply (cut_tac Hausdorff maxchain_subset_chain) | 
| 14706 | 243 | apply (erule exE) | 
| 244 | apply (drule subsetD, assumption) | |
| 245 | apply (drule bspec, assumption, erule bexE) | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 246 | apply (rule_tac x = y in bexI) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 247 | prefer 2 apply assumption | 
| 14706 | 248 | apply clarify | 
| 249 | apply (rule ccontr) | |
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 250 | apply (frule_tac z = x in chain_extend) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 251 | apply (assumption, blast) | 
| 14706 | 252 | apply (unfold maxchain_def super_def psubset_def) | 
| 13551 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 253 | apply (blast elim!: equalityCE) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 254 | done | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 259 | by (unfold chain_def, blast) | 
| 
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" | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 262 | by (unfold chain_def, blast) | 
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 263 | |
| 
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
 paulson parents: diff
changeset | 264 | end |