(* Title : HOL/Library/Zorn.thy ID : $Id: Zorn.thy,v 1.19 2008/06/03 10:38:39 ballarin Exp $ Author : Jacques D. Fleuriot, Tobias Nipkow Description : Zorn's Lemma (ported from Larry Paulson's Zorn.thy in ZF) The well-ordering theorem *) header {* Zorn's Lemma *} theory Zorn imports Order_Relation begin (* Define globally? In Set.thy? *) definition chain_subset :: "'a set set => bool" ("chain⊆") where "chain⊆ C ≡ ∀A∈C.∀B∈C. A ⊆ B ∨ B ⊆ A" text{* The lemma and section numbers refer to an unpublished article \cite{Abrial-Laffitte}. *} definition chain :: "'a set set => 'a set set set" where "chain S = {F. F ⊆ S & chain⊆ F}" definition super :: "['a set set,'a set set] => 'a set set set" where "super S c = {d. d ∈ chain S & c ⊂ d}" definition maxchain :: "'a set set => 'a set set set" where "maxchain S = {c. c ∈ chain S & super S c = {}}" definition succ :: "['a set set,'a set set] => 'a set set" where "succ S c = (if c ∉ chain S | c ∈ maxchain S then c else SOME c'. c' ∈ super S c)" inductive_set TFin :: "'a set set => 'a set set set" for S :: "'a set set" where succI: "x ∈ TFin S ==> succ S x ∈ TFin S" | Pow_UnionI: "Y ∈ Pow(TFin S) ==> Union(Y) ∈ TFin S" subsection{*Mathematical Preamble*} lemma Union_lemma0: "(∀x ∈ C. x ⊆ A | B ⊆ x) ==> Union(C) ⊆ A | B ⊆ Union(C)" by blast text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*} lemma Abrial_axiom1: "x ⊆ succ S x" apply (auto simp add: succ_def super_def maxchain_def) apply (rule contrapos_np, assumption) apply (rule_tac Q="λS. xa ∈ S" in someI2, blast+) done lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] lemma TFin_induct: assumes H: "n ∈ TFin S" and I: "!!x. x ∈ TFin S ==> P x ==> P (succ S x)" "!!Y. Y ⊆ TFin S ==> Ball Y P ==> P(Union Y)" shows "P n" using H apply (induct rule: TFin.induct [where P=P]) apply (blast intro: I)+ done lemma succ_trans: "x ⊆ y ==> x ⊆ succ S y" apply (erule subset_trans) apply (rule Abrial_axiom1) done text{*Lemma 1 of section 3.1*} lemma TFin_linear_lemma1: "[| n ∈ TFin S; m ∈ TFin S; ∀x ∈ TFin S. x ⊆ m --> x = m | succ S x ⊆ m |] ==> n ⊆ m | succ S m ⊆ n" apply (erule TFin_induct) apply (erule_tac [2] Union_lemma0) apply (blast del: subsetI intro: succ_trans) done text{* Lemma 2 of section 3.2 *} lemma TFin_linear_lemma2: "m ∈ TFin S ==> ∀n ∈ TFin S. n ⊆ m --> n=m | succ S n ⊆ m" apply (erule TFin_induct) apply (rule impI [THEN ballI]) txt{*case split using @{text TFin_linear_lemma1}*} apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], assumption+) apply (drule_tac x = n in bspec, assumption) apply (blast del: subsetI intro: succ_trans, blast) txt{*second induction step*} apply (rule impI [THEN ballI]) apply (rule Union_lemma0 [THEN disjE]) apply (rule_tac [3] disjI2) prefer 2 apply blast apply (rule ballI) apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], assumption+, auto) apply (blast intro!: Abrial_axiom1 [THEN subsetD]) done text{*Re-ordering the premises of Lemma 2*} lemma TFin_subsetD: "[| n ⊆ m; m ∈ TFin S; n ∈ TFin S |] ==> n=m | succ S n ⊆ m" by (rule TFin_linear_lemma2 [rule_format]) text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*} lemma TFin_subset_linear: "[| m ∈ TFin S; n ∈ TFin S|] ==> n ⊆ m | m ⊆ n" apply (rule disjE) apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) apply (assumption+, erule disjI2) apply (blast del: subsetI intro: subsetI Abrial_axiom1 [THEN subset_trans]) done text{*Lemma 3 of section 3.3*} lemma eq_succ_upper: "[| n ∈ TFin S; m ∈ TFin S; m = succ S m |] ==> n ⊆ m" apply (erule TFin_induct) apply (drule TFin_subsetD) apply (assumption+, force, blast) done text{*Property 3.3 of section 3.3*} lemma equal_succ_Union: "m ∈ TFin S ==> (m = succ S m) = (m = Union(TFin S))" apply (rule iffI) apply (rule Union_upper [THEN equalityI]) apply assumption apply (rule eq_succ_upper [THEN Union_least], assumption+) apply (erule ssubst) apply (rule Abrial_axiom1 [THEN equalityI]) apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI) done subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*} text{*NB: We assume the partial ordering is @{text "⊆"}, the subset relation!*} lemma empty_set_mem_chain: "({} :: 'a set set) ∈ chain S" by (unfold chain_def chain_subset_def) auto lemma super_subset_chain: "super S c ⊆ chain S" by (unfold super_def) blast lemma maxchain_subset_chain: "maxchain S ⊆ chain S" by (unfold maxchain_def) blast lemma mem_super_Ex: "c ∈ chain S - maxchain S ==> EX d. d ∈ super S c" by (unfold super_def maxchain_def) auto lemma select_super: "c ∈ chain S - maxchain S ==> (\<some>c'. c': super S c): super S c" apply (erule mem_super_Ex [THEN exE]) apply (rule someI2 [where Q="%X. X : super S c"], auto) done lemma select_not_equals: "c ∈ chain S - maxchain S ==> (\<some>c'. c': super S c) ≠ c" apply (rule notI) apply (drule select_super) apply (simp add: super_def less_le) done lemma succI3: "c ∈ chain S - maxchain S ==> succ S c = (\<some>c'. c': super S c)" by (unfold succ_def) (blast intro!: if_not_P) lemma succ_not_equals: "c ∈ chain S - maxchain S ==> succ S c ≠ c" apply (frule succI3) apply (simp (no_asm_simp)) apply (rule select_not_equals, assumption) done lemma TFin_chain_lemma4: "c ∈ TFin S ==> (c :: 'a set set): chain S" apply (erule TFin_induct) apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]]) apply (unfold chain_def chain_subset_def) apply (rule CollectI, safe) apply (drule bspec, assumption) apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE], best+) done theorem Hausdorff: "∃c. (c :: 'a set set): maxchain S" apply (rule_tac x = "Union (TFin S)" in exI) apply (rule classical) apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ") prefer 2 apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4]) apply (drule DiffI [THEN succ_not_equals], blast+) done subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then There Is a Maximal Element*} lemma chain_extend: "[| c ∈ chain S; z ∈ S; ∀x ∈ c. x ⊆ (z:: 'a set) |] ==> {z} Un c ∈ chain S" by (unfold chain_def chain_subset_def) blast lemma chain_Union_upper: "[| c ∈ chain S; x ∈ c |] ==> x ⊆ Union(c)" by auto lemma chain_ball_Union_upper: "c ∈ chain S ==> ∀x ∈ c. x ⊆ Union(c)" by auto lemma maxchain_Zorn: "[| c ∈ maxchain S; u ∈ S; Union(c) ⊆ u |] ==> Union(c) = u" apply (rule ccontr) apply (simp add: maxchain_def) apply (erule conjE) apply (subgoal_tac "({u} Un c) ∈ super S c") apply simp apply (unfold super_def less_le) apply (blast intro: chain_extend dest: chain_Union_upper) done theorem Zorn_Lemma: "∀c ∈ chain S. Union(c): S ==> ∃y ∈ S. ∀z ∈ S. y ⊆ z --> y = z" apply (cut_tac Hausdorff maxchain_subset_chain) apply (erule exE) apply (drule subsetD, assumption) apply (drule bspec, assumption) apply (rule_tac x = "Union(c)" in bexI) apply (rule ballI, rule impI) apply (blast dest!: maxchain_Zorn, assumption) done subsection{*Alternative version of Zorn's Lemma*} lemma Zorn_Lemma2: "∀c ∈ chain S. ∃y ∈ S. ∀x ∈ c. x ⊆ y ==> ∃y ∈ S. ∀x ∈ S. (y :: 'a set) ⊆ x --> y = x" apply (cut_tac Hausdorff maxchain_subset_chain) apply (erule exE) apply (drule subsetD, assumption) apply (drule bspec, assumption, erule bexE) apply (rule_tac x = y in bexI) prefer 2 apply assumption apply clarify apply (rule ccontr) apply (frule_tac z = x in chain_extend) apply (assumption, blast) apply (unfold maxchain_def super_def less_le) apply (blast elim!: equalityCE) done text{*Various other lemmas*} lemma chainD: "[| c ∈ chain S; x ∈ c; y ∈ c |] ==> x ⊆ y | y ⊆ x" by (unfold chain_def chain_subset_def) blast lemma chainD2: "!!(c :: 'a set set). c ∈ chain S ==> c ⊆ S" by (unfold chain_def) blast (* Define globally? In Relation.thy? *) definition Chain :: "('a*'a)set => 'a set set" where "Chain r ≡ {A. ∀a∈A.∀b∈A. (a,b) : r ∨ (b,a) ∈ r}" lemma mono_Chain: "r ⊆ s ==> Chain r ⊆ Chain s" unfolding Chain_def by blast text{* Zorn's lemma for partial orders: *} lemma Zorns_po_lemma: assumes po: "Partial_order r" and u: "∀C∈Chain r. ∃u∈Field r. ∀a∈C. (a,u):r" shows "∃m∈Field r. ∀a∈Field r. (m,a):r --> a=m" proof- have "Preorder r" using po by(simp add:partial_order_on_def) --{* Mirror r in the set of subsets below (wrt r) elements of A*} let ?B = "%x. r^-1 `` {x}" let ?S = "?B ` Field r" have "∀C ∈ chain ?S. EX U:?S. ALL A:C. A⊆U" proof (auto simp:chain_def chain_subset_def) fix C assume 1: "C ⊆ ?S" and 2: "∀A∈C.∀B∈C. A⊆B | B⊆A" let ?A = "{x∈Field r. ∃M∈C. M = ?B x}" have "C = ?B ` ?A" using 1 by(auto simp: image_def) have "?A∈Chain r" proof (simp add:Chain_def, intro allI impI, elim conjE) fix a b assume "a ∈ Field r" "?B a ∈ C" "b ∈ Field r" "?B b ∈ C" hence "?B a ⊆ ?B b ∨ ?B b ⊆ ?B a" using 2 by auto thus "(a, b) ∈ r ∨ (b, a) ∈ r" using `Preorder r` `a:Field r` `b:Field r` by(simp add:subset_Image1_Image1_iff) qed then obtain u where uA: "u:Field r" "∀a∈?A. (a,u) : r" using u by auto have "∀A∈C. A ⊆ r^-1 `` {u}" (is "?P u") proof auto fix a B assume aB: "B:C" "a:B" with 1 obtain x where "x:Field r" "B = r^-1 `` {x}" by auto thus "(a,u) : r" using uA aB `Preorder r` by (auto simp add: preorder_on_def refl_def) (metis transD) qed thus "EX u:Field r. ?P u" using `u:Field r` by blast qed from Zorn_Lemma2[OF this] obtain m B where "m:Field r" "B = r^-1 `` {m}" "∀x∈Field r. B ⊆ r^-1 `` {x} --> B = r^-1 `` {x}" by auto hence "∀a∈Field r. (m, a) ∈ r --> a = m" using po `Preorder r` `m:Field r` by(auto simp:subset_Image1_Image1_iff Partial_order_eq_Image1_Image1_iff) thus ?thesis using `m:Field r` by blast qed (* The initial segment of a relation appears generally useful. Move to Relation.thy? Definition correct/most general? Naming? *) definition init_seg_of :: "(('a*'a)set * ('a*'a)set)set" where "init_seg_of == {(r,s). r ⊆ s ∧ (∀a b c. (a,b):s ∧ (b,c):r --> (a,b):r)}" abbreviation initialSegmentOf :: "('a*'a)set => ('a*'a)set => bool" (infix "initial'_segment'_of" 55) where "r initial_segment_of s == (r,s):init_seg_of" lemma refl_init_seg_of[simp]: "r initial_segment_of r" by(simp add:init_seg_of_def) lemma trans_init_seg_of: "r initial_segment_of s ==> s initial_segment_of t ==> r initial_segment_of t" by(simp (no_asm_use) add: init_seg_of_def) (metis Domain_iff UnCI Un_absorb2 subset_trans) lemma antisym_init_seg_of: "r initial_segment_of s ==> s initial_segment_of r ==> r=s" by(auto simp:init_seg_of_def) lemma Chain_init_seg_of_Union: "R ∈ Chain init_seg_of ==> r∈R ==> r initial_segment_of \<Union>R" by(auto simp add:init_seg_of_def Chain_def Ball_def) blast lemma chain_subset_trans_Union: "chain⊆ R ==> ∀r∈R. trans r ==> trans(\<Union>R)" apply(auto simp add:chain_subset_def) apply(simp (no_asm_use) add:trans_def) apply (metis subsetD) done lemma chain_subset_antisym_Union: "chain⊆ R ==> ∀r∈R. antisym r ==> antisym(\<Union>R)" apply(auto simp add:chain_subset_def antisym_def) apply (metis subsetD) done lemma chain_subset_Total_Union: assumes "chain⊆ R" "∀r∈R. Total r" shows "Total (\<Union>R)" proof (simp add: total_on_def Ball_def, auto del:disjCI) fix r s a b assume A: "r:R" "s:R" "a:Field r" "b:Field s" "a≠b" from `chain⊆ R` `r:R` `s:R` have "r⊆s ∨ s⊆r" by(simp add:chain_subset_def) thus "(∃r∈R. (a,b) ∈ r) ∨ (∃r∈R. (b,a) ∈ r)" proof assume "r⊆s" hence "(a,b):s ∨ (b,a):s" using assms(2) A by(simp add:total_on_def)(metis mono_Field subsetD) thus ?thesis using `s:R` by blast next assume "s⊆r" hence "(a,b):r ∨ (b,a):r" using assms(2) A by(simp add:total_on_def)(metis mono_Field subsetD) thus ?thesis using `r:R` by blast qed qed lemma wf_Union_wf_init_segs: assumes "R ∈ Chain init_seg_of" and "∀r∈R. wf r" shows "wf(\<Union>R)" proof(simp add:wf_iff_no_infinite_down_chain, rule ccontr, auto) fix f assume 1: "∀i. ∃r∈R. (f(Suc i), f i) ∈ r" then obtain r where "r:R" and "(f(Suc 0), f 0) : r" by auto { fix i have "(f(Suc i), f i) ∈ r" proof(induct i) case 0 show ?case by fact next case (Suc i) moreover obtain s where "s∈R" and "(f(Suc(Suc i)), f(Suc i)) ∈ s" using 1 by auto moreover hence "s initial_segment_of r ∨ r initial_segment_of s" using assms(1) `r:R` by(simp add: Chain_def) ultimately show ?case by(simp add:init_seg_of_def) blast qed } thus False using assms(2) `r:R` by(simp add:wf_iff_no_infinite_down_chain) blast qed lemma Chain_inits_DiffI: "R ∈ Chain init_seg_of ==> {r - s |r. r ∈ R} ∈ Chain init_seg_of" apply(auto simp:Chain_def init_seg_of_def) apply (metis subsetD) apply (metis subsetD) done theorem well_ordering: "∃r::('a*'a)set. Well_order r ∧ Field r = UNIV" proof- -- {*The initial segment relation on well-orders: *} let ?WO = "{r::('a*'a)set. Well_order r}" def I ≡ "init_seg_of ∩ ?WO × ?WO" have I_init: "I ⊆ init_seg_of" by(auto simp:I_def) hence subch: "!!R. R : Chain I ==> chain⊆ R" by(auto simp:init_seg_of_def chain_subset_def Chain_def) have Chain_wo: "!!R r. R ∈ Chain I ==> r ∈ R ==> Well_order r" by(simp add:Chain_def I_def) blast have FI: "Field I = ?WO" by(auto simp add:I_def init_seg_of_def Field_def) hence 0: "Partial_order I" by(auto simp: partial_order_on_def preorder_on_def antisym_def antisym_init_seg_of refl_def trans_def I_def elim!: trans_init_seg_of) -- {*I-chains have upper bounds in ?WO wrt I: their Union*} { fix R assume "R ∈ Chain I" hence Ris: "R ∈ Chain init_seg_of" using mono_Chain[OF I_init] by blast have subch: "chain⊆ R" using `R : Chain I` I_init by(auto simp:init_seg_of_def chain_subset_def Chain_def) have "∀r∈R. Refl r" "∀r∈R. trans r" "∀r∈R. antisym r" "∀r∈R. Total r" "∀r∈R. wf(r-Id)" using Chain_wo[OF `R ∈ Chain I`] by(simp_all add:order_on_defs) have "Refl (\<Union>R)" using `∀r∈R. Refl r` by(auto simp:refl_def) moreover have "trans (\<Union>R)" by(rule chain_subset_trans_Union[OF subch `∀r∈R. trans r`]) moreover have "antisym(\<Union>R)" by(rule chain_subset_antisym_Union[OF subch `∀r∈R. antisym r`]) moreover have "Total (\<Union>R)" by(rule chain_subset_Total_Union[OF subch `∀r∈R. Total r`]) moreover have "wf((\<Union>R)-Id)" proof- have "(\<Union>R)-Id = \<Union>{r-Id|r. r ∈ R}" by blast with `∀r∈R. wf(r-Id)` wf_Union_wf_init_segs[OF Chain_inits_DiffI[OF Ris]] show ?thesis by (simp (no_asm_simp)) blast qed ultimately have "Well_order (\<Union>R)" by(simp add:order_on_defs) moreover have "∀r ∈ R. r initial_segment_of \<Union>R" using Ris by(simp add: Chain_init_seg_of_Union) ultimately have "\<Union>R : ?WO ∧ (∀r∈R. (r,\<Union>R) : I)" using mono_Chain[OF I_init] `R ∈ Chain I` by(simp (no_asm) add:I_def del:Field_Union)(metis Chain_wo subsetD) } hence 1: "∀R ∈ Chain I. ∃u∈Field I. ∀r∈R. (r,u) : I" by (subst FI) blast --{*Zorn's Lemma yields a maximal well-order m:*} then obtain m::"('a*'a)set" where "Well_order m" and max: "∀r. Well_order r ∧ (m,r):I --> r=m" using Zorns_po_lemma[OF 0 1] by (auto simp:FI) --{*Now show by contradiction that m covers the whole type:*} { fix x::'a assume "x ∉ Field m" --{*We assume that x is not covered and extend m at the top with x*} have "m ≠ {}" proof assume "m={}" moreover have "Well_order {(x,x)}" by(simp add:order_on_defs refl_def trans_def antisym_def total_on_def Field_def Domain_def Range_def) ultimately show False using max by (auto simp:I_def init_seg_of_def simp del:Field_insert) qed hence "Field m ≠ {}" by(auto simp:Field_def) moreover have "wf(m-Id)" using `Well_order m` by(simp add:well_order_on_def) --{*The extension of m by x:*} let ?s = "{(a,x)|a. a : Field m}" let ?m = "insert (x,x) m Un ?s" have Fm: "Field ?m = insert x (Field m)" apply(simp add:Field_insert Field_Un) unfolding Field_def by auto have "Refl m" "trans m" "antisym m" "Total m" "wf(m-Id)" using `Well_order m` by(simp_all add:order_on_defs) --{*We show that the extension is a well-order*} have "Refl ?m" using `Refl m` Fm by(auto simp:refl_def) moreover have "trans ?m" using `trans m` `x ∉ Field m` unfolding trans_def Field_def Domain_def Range_def by blast moreover have "antisym ?m" using `antisym m` `x ∉ Field m` unfolding antisym_def Field_def Domain_def Range_def by blast moreover have "Total ?m" using `Total m` Fm by(auto simp: total_on_def) moreover have "wf(?m-Id)" proof- have "wf ?s" using `x ∉ Field m` by(auto simp add:wf_eq_minimal Field_def Domain_def Range_def) metis thus ?thesis using `wf(m-Id)` `x ∉ Field m` wf_subset[OF `wf ?s` Diff_subset] by (fastsimp intro!: wf_Un simp add: Un_Diff Field_def) qed ultimately have "Well_order ?m" by(simp add:order_on_defs) --{*We show that the extension is above m*} moreover hence "(m,?m) : I" using `Well_order m` `x ∉ Field m` by(fastsimp simp:I_def init_seg_of_def Field_def Domain_def Range_def) ultimately --{*This contradicts maximality of m:*} have False using max `x ∉ Field m` unfolding Field_def by blast } hence "Field m = UNIV" by auto moreover with `Well_order m` have "Well_order m" by simp ultimately show ?thesis by blast qed corollary well_order_on: "∃r::('a*'a)set. well_order_on A r" proof - obtain r::"('a*'a)set" where wo: "Well_order r" and univ: "Field r = UNIV" using well_ordering[where 'a = "'a"] by blast let ?r = "{(x,y). x:A & y:A & (x,y):r}" have 1: "Field ?r = A" using wo univ by(fastsimp simp: Field_def Domain_def Range_def order_on_defs refl_def) have "Refl r" "trans r" "antisym r" "Total r" "wf(r-Id)" using `Well_order r` by(simp_all add:order_on_defs) have "Refl ?r" using `Refl r` by(auto simp:refl_def 1 univ) moreover have "trans ?r" using `trans r` unfolding trans_def by blast moreover have "antisym ?r" using `antisym r` unfolding antisym_def by blast moreover have "Total ?r" using `Total r` by(simp add:total_on_def 1 univ) moreover have "wf(?r - Id)" by(rule wf_subset[OF `wf(r-Id)`]) blast ultimately have "Well_order ?r" by(simp add:order_on_defs) with 1 show ?thesis by metis qed end
lemma Union_lemma0:
∀x∈C. x ⊆ A ∨ B ⊆ x ==> Union C ⊆ A ∨ B ⊆ Union C
lemma Abrial_axiom1:
x ⊆ succ S x
lemma TFin_UnionI:
Y ⊆ TFin S ==> Union Y ∈ TFin S
lemma TFin_induct:
[| n ∈ TFin S; !!x. [| x ∈ TFin S; P x |] ==> P (succ S x);
!!Y. [| Y ⊆ TFin S; Ball Y P |] ==> P (Union Y) |]
==> P n
lemma succ_trans:
x ⊆ y ==> x ⊆ succ S y
lemma TFin_linear_lemma1:
[| n ∈ TFin S; m ∈ TFin S; ∀x∈TFin S. x ⊆ m --> x = m ∨ succ S x ⊆ m |]
==> n ⊆ m ∨ succ S m ⊆ n
lemma TFin_linear_lemma2:
m ∈ TFin S ==> ∀n∈TFin S. n ⊆ m --> n = m ∨ succ S n ⊆ m
lemma TFin_subsetD:
[| n ⊆ m; m ∈ TFin S; n ∈ TFin S |] ==> n = m ∨ succ S n ⊆ m
lemma TFin_subset_linear:
[| m ∈ TFin S; n ∈ TFin S |] ==> n ⊆ m ∨ m ⊆ n
lemma eq_succ_upper:
[| n ∈ TFin S; m ∈ TFin S; m = succ S m |] ==> n ⊆ m
lemma equal_succ_Union:
m ∈ TFin S ==> (m = succ S m) = (m = Union (TFin S))
lemma empty_set_mem_chain:
{} ∈ chain S
lemma super_subset_chain:
super S c ⊆ chain S
lemma maxchain_subset_chain:
maxchain S ⊆ chain S
lemma mem_super_Ex:
c ∈ chain S - maxchain S ==> ∃d. d ∈ super S c
lemma select_super:
c ∈ chain S - maxchain S ==> (SOME c'. c' ∈ super S c) ∈ super S c
lemma select_not_equals:
c ∈ chain S - maxchain S ==> (SOME c'. c' ∈ super S c) ≠ c
lemma succI3:
c ∈ chain S - maxchain S ==> succ S c = (SOME c'. c' ∈ super S c)
lemma succ_not_equals:
c ∈ chain S - maxchain S ==> succ S c ≠ c
lemma TFin_chain_lemma4:
c ∈ TFin S ==> c ∈ chain S
theorem Hausdorff:
∃c. c ∈ maxchain S
lemma chain_extend:
[| c ∈ chain S; z ∈ S; ∀x∈c. x ⊆ z |] ==> {z} ∪ c ∈ chain S
lemma chain_Union_upper:
[| c ∈ chain S; x ∈ c |] ==> x ⊆ Union c
lemma chain_ball_Union_upper:
c ∈ chain S ==> ∀x∈c. x ⊆ Union c
lemma maxchain_Zorn:
[| c ∈ maxchain S; u ∈ S; Union c ⊆ u |] ==> Union c = u
theorem Zorn_Lemma:
∀c∈chain S. Union c ∈ S ==> ∃y∈S. ∀z∈S. y ⊆ z --> y = z
lemma Zorn_Lemma2:
∀c∈chain S. ∃y∈S. ∀x∈c. x ⊆ y ==> ∃y∈S. ∀x∈S. y ⊆ x --> y = x
lemma chainD:
[| c ∈ chain S; x ∈ c; y ∈ c |] ==> x ⊆ y ∨ y ⊆ x
lemma chainD2:
c ∈ chain S ==> c ⊆ S
lemma mono_Chain:
r ⊆ s ==> Chain r ⊆ Chain s
lemma Zorns_po_lemma:
[| Partial_order r; ∀C∈Chain r. ∃u∈Field r. ∀a∈C. (a, u) ∈ r |]
==> ∃m∈Field r. ∀a∈Field r. (m, a) ∈ r --> a = m
lemma refl_init_seg_of:
r initial_segment_of r
lemma trans_init_seg_of:
[| r initial_segment_of s; s initial_segment_of t |] ==> r initial_segment_of t
lemma antisym_init_seg_of:
[| r initial_segment_of s; s initial_segment_of r |] ==> r = s
lemma Chain_init_seg_of_Union:
[| R ∈ Chain init_seg_of; r ∈ R |] ==> r initial_segment_of Union R
lemma chain_subset_trans_Union:
[| chain⊆ R; ∀r∈R. trans r |] ==> trans (Union R)
lemma chain_subset_antisym_Union:
[| chain⊆ R; ∀r∈R. antisym r |] ==> antisym (Union R)
lemma chain_subset_Total_Union:
[| chain⊆ R; ∀r∈R. Total r |] ==> Total (Union R)
lemma wf_Union_wf_init_segs:
[| R ∈ Chain init_seg_of; ∀r∈R. wf r |] ==> wf (Union R)
lemma Chain_inits_DiffI:
R ∈ Chain init_seg_of ==> {r - s |r. r ∈ R} ∈ Chain init_seg_of
theorem well_ordering:
∃r. Well_order r ∧ Field r = UNIV
corollary well_order_on:
∃r. well_order_on A r