wenzelm@14706: (* Title : HOL/Library/Zorn.thy paulson@13652: ID : $Id$ paulson@13652: Author : Jacques D. Fleuriot wenzelm@14706: Description : Zorn's Lemma -- see Larry Paulson's Zorn.thy in ZF wenzelm@14706: *) paulson@13551: wenzelm@14706: header {* Zorn's Lemma *} paulson@13551: nipkow@15131: theory Zorn nipkow@15140: imports Main nipkow@15131: begin paulson@13551: wenzelm@14706: text{* wenzelm@14706: The lemma and section numbers refer to an unpublished article wenzelm@14706: \cite{Abrial-Laffitte}. wenzelm@14706: *} paulson@13551: wenzelm@19736: definition wenzelm@21404: chain :: "'a set set => 'a set set set" where wenzelm@19736: "chain S = {F. F \ S & (\x \ F. \y \ F. x \ y | y \ x)}" paulson@13551: wenzelm@21404: definition wenzelm@21404: super :: "['a set set,'a set set] => 'a set set set" where wenzelm@19736: "super S c = {d. d \ chain S & c \ d}" paulson@13551: wenzelm@21404: definition wenzelm@21404: maxchain :: "'a set set => 'a set set set" where wenzelm@19736: "maxchain S = {c. c \ chain S & super S c = {}}" paulson@13551: wenzelm@21404: definition wenzelm@21404: succ :: "['a set set,'a set set] => 'a set set" where wenzelm@19736: "succ S c = wenzelm@19736: (if c \ chain S | c \ maxchain S wenzelm@19736: then c else SOME c'. c' \ super S c)" paulson@13551: berghofe@23755: inductive_set wenzelm@14706: TFin :: "'a set set => 'a set set set" berghofe@23755: for S :: "'a set set" berghofe@23755: where paulson@13551: succI: "x \ TFin S ==> succ S x \ TFin S" berghofe@23755: | Pow_UnionI: "Y \ Pow(TFin S) ==> Union(Y) \ TFin S" paulson@13551: monos Pow_mono paulson@13551: paulson@13551: paulson@13551: subsection{*Mathematical Preamble*} paulson@13551: wenzelm@17200: lemma Union_lemma0: paulson@18143: "(\x \ C. x \ A | B \ x) ==> Union(C) \ A | B \ Union(C)" wenzelm@17200: by blast paulson@13551: paulson@13551: paulson@13551: text{*This is theorem @{text increasingD2} of ZF/Zorn.thy*} wenzelm@17200: paulson@13551: lemma Abrial_axiom1: "x \ succ S x" wenzelm@17200: apply (unfold succ_def) wenzelm@17200: apply (rule split_if [THEN iffD2]) wenzelm@17200: apply (auto simp add: super_def maxchain_def psubset_def) wenzelm@18585: apply (rule contrapos_np, assumption) wenzelm@17200: apply (rule someI2, blast+) wenzelm@17200: done paulson@13551: paulson@13551: lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] paulson@13551: wenzelm@14706: lemma TFin_induct: wenzelm@14706: "[| n \ TFin S; wenzelm@14706: !!x. [| x \ TFin S; P(x) |] ==> P(succ S x); wenzelm@14706: !!Y. [| Y \ TFin S; Ball Y P |] ==> P(Union Y) |] paulson@13551: ==> P(n)" wenzelm@19736: apply (induct set: TFin) wenzelm@17200: apply blast+ wenzelm@17200: done paulson@13551: paulson@13551: lemma succ_trans: "x \ y ==> x \ succ S y" wenzelm@17200: apply (erule subset_trans) wenzelm@17200: apply (rule Abrial_axiom1) wenzelm@17200: done paulson@13551: paulson@13551: text{*Lemma 1 of section 3.1*} paulson@13551: lemma TFin_linear_lemma1: wenzelm@14706: "[| n \ TFin S; m \ TFin S; wenzelm@14706: \x \ TFin S. x \ m --> x = m | succ S x \ m paulson@13551: |] ==> n \ m | succ S m \ n" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (erule_tac [2] Union_lemma0) wenzelm@17200: apply (blast del: subsetI intro: succ_trans) wenzelm@17200: done paulson@13551: paulson@13551: text{* Lemma 2 of section 3.2 *} paulson@13551: lemma TFin_linear_lemma2: paulson@13551: "m \ TFin S ==> \n \ TFin S. n \ m --> n=m | succ S n \ m" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (rule impI [THEN ballI]) wenzelm@17200: txt{*case split using @{text TFin_linear_lemma1}*} wenzelm@17200: apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], wenzelm@17200: assumption+) wenzelm@17200: apply (drule_tac x = n in bspec, assumption) wenzelm@17200: apply (blast del: subsetI intro: succ_trans, blast) wenzelm@17200: txt{*second induction step*} wenzelm@17200: apply (rule impI [THEN ballI]) wenzelm@17200: apply (rule Union_lemma0 [THEN disjE]) wenzelm@17200: apply (rule_tac [3] disjI2) wenzelm@17200: prefer 2 apply blast wenzelm@17200: apply (rule ballI) wenzelm@17200: apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], wenzelm@17200: assumption+, auto) wenzelm@17200: apply (blast intro!: Abrial_axiom1 [THEN subsetD]) wenzelm@17200: done paulson@13551: paulson@13551: text{*Re-ordering the premises of Lemma 2*} paulson@13551: lemma TFin_subsetD: paulson@13551: "[| n \ m; m \ TFin S; n \ TFin S |] ==> n=m | succ S n \ m" wenzelm@17200: by (rule TFin_linear_lemma2 [rule_format]) paulson@13551: paulson@13551: text{*Consequences from section 3.3 -- Property 3.2, the ordering is total*} paulson@13551: lemma TFin_subset_linear: "[| m \ TFin S; n \ TFin S|] ==> n \ m | m \ n" wenzelm@17200: apply (rule disjE) wenzelm@17200: apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) wenzelm@17200: apply (assumption+, erule disjI2) wenzelm@17200: apply (blast del: subsetI wenzelm@17200: intro: subsetI Abrial_axiom1 [THEN subset_trans]) wenzelm@17200: done paulson@13551: paulson@13551: text{*Lemma 3 of section 3.3*} paulson@13551: lemma eq_succ_upper: "[| n \ TFin S; m \ TFin S; m = succ S m |] ==> n \ m" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (drule TFin_subsetD) wenzelm@17200: apply (assumption+, force, blast) wenzelm@17200: done paulson@13551: paulson@13551: text{*Property 3.3 of section 3.3*} paulson@13551: lemma equal_succ_Union: "m \ TFin S ==> (m = succ S m) = (m = Union(TFin S))" wenzelm@17200: apply (rule iffI) wenzelm@17200: apply (rule Union_upper [THEN equalityI]) paulson@18143: apply assumption paulson@18143: apply (rule eq_succ_upper [THEN Union_least], assumption+) wenzelm@17200: apply (erule ssubst) wenzelm@17200: apply (rule Abrial_axiom1 [THEN equalityI]) wenzelm@17200: apply (blast del: subsetI intro: subsetI TFin_UnionI TFin.succI) wenzelm@17200: done paulson@13551: paulson@13551: subsection{*Hausdorff's Theorem: Every Set Contains a Maximal Chain.*} paulson@13551: wenzelm@14706: text{*NB: We assume the partial ordering is @{text "\"}, paulson@13551: the subset relation!*} paulson@13551: paulson@13551: lemma empty_set_mem_chain: "({} :: 'a set set) \ chain S" wenzelm@17200: by (unfold chain_def) auto paulson@13551: paulson@13551: lemma super_subset_chain: "super S c \ chain S" wenzelm@17200: by (unfold super_def) blast paulson@13551: paulson@13551: lemma maxchain_subset_chain: "maxchain S \ chain S" wenzelm@17200: by (unfold maxchain_def) blast paulson@13551: paulson@13551: lemma mem_super_Ex: "c \ chain S - maxchain S ==> ? d. d \ super S c" wenzelm@17200: by (unfold super_def maxchain_def) auto paulson@13551: paulson@18143: lemma select_super: paulson@18143: "c \ chain S - maxchain S ==> (\c'. c': super S c): super S c" wenzelm@17200: apply (erule mem_super_Ex [THEN exE]) wenzelm@17200: apply (rule someI2, auto) wenzelm@17200: done paulson@13551: paulson@18143: lemma select_not_equals: paulson@18143: "c \ chain S - maxchain S ==> (\c'. c': super S c) \ c" wenzelm@17200: apply (rule notI) wenzelm@17200: apply (drule select_super) wenzelm@17200: apply (simp add: super_def psubset_def) wenzelm@17200: done paulson@13551: wenzelm@17200: lemma succI3: "c \ chain S - maxchain S ==> succ S c = (\c'. c': super S c)" wenzelm@17200: by (unfold succ_def) (blast intro!: if_not_P) paulson@13551: paulson@13551: lemma succ_not_equals: "c \ chain S - maxchain S ==> succ S c \ c" wenzelm@17200: apply (frule succI3) wenzelm@17200: apply (simp (no_asm_simp)) wenzelm@17200: apply (rule select_not_equals, assumption) wenzelm@17200: done paulson@13551: paulson@13551: lemma TFin_chain_lemma4: "c \ TFin S ==> (c :: 'a set set): chain S" wenzelm@17200: apply (erule TFin_induct) wenzelm@17200: apply (simp add: succ_def select_super [THEN super_subset_chain[THEN subsetD]]) wenzelm@17200: apply (unfold chain_def) wenzelm@17200: apply (rule CollectI, safe) wenzelm@17200: apply (drule bspec, assumption) wenzelm@17200: apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE], wenzelm@17200: blast+) wenzelm@17200: done wenzelm@14706: paulson@13551: theorem Hausdorff: "\c. (c :: 'a set set): maxchain S" paulson@18143: apply (rule_tac x = "Union (TFin S)" in exI) wenzelm@17200: apply (rule classical) wenzelm@17200: apply (subgoal_tac "succ S (Union (TFin S)) = Union (TFin S) ") wenzelm@17200: prefer 2 wenzelm@17200: apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) wenzelm@17200: apply (cut_tac subset_refl [THEN TFin_UnionI, THEN TFin_chain_lemma4]) wenzelm@17200: apply (drule DiffI [THEN succ_not_equals], blast+) wenzelm@17200: done paulson@13551: paulson@13551: wenzelm@14706: subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then paulson@13551: There Is a Maximal Element*} paulson@13551: wenzelm@14706: lemma chain_extend: wenzelm@14706: "[| c \ chain S; z \ S; paulson@18143: \x \ c. x \ (z:: 'a set) |] ==> {z} Un c \ chain S" wenzelm@17200: by (unfold chain_def) blast paulson@13551: paulson@13551: lemma chain_Union_upper: "[| c \ chain S; x \ c |] ==> x \ Union(c)" wenzelm@17200: by (unfold chain_def) auto paulson@13551: paulson@13551: lemma chain_ball_Union_upper: "c \ chain S ==> \x \ c. x \ Union(c)" wenzelm@17200: by (unfold chain_def) auto paulson@13551: paulson@13551: lemma maxchain_Zorn: paulson@13551: "[| c \ maxchain S; u \ S; Union(c) \ u |] ==> Union(c) = u" wenzelm@17200: apply (rule ccontr) wenzelm@17200: apply (simp add: maxchain_def) wenzelm@17200: apply (erule conjE) paulson@18143: apply (subgoal_tac "({u} Un c) \ super S c") wenzelm@17200: apply simp wenzelm@17200: apply (unfold super_def psubset_def) wenzelm@17200: apply (blast intro: chain_extend dest: chain_Union_upper) wenzelm@17200: done paulson@13551: paulson@13551: theorem Zorn_Lemma: wenzelm@17200: "\c \ chain S. Union(c): S ==> \y \ S. \z \ S. y \ z --> y = z" wenzelm@17200: apply (cut_tac Hausdorff maxchain_subset_chain) wenzelm@17200: apply (erule exE) wenzelm@17200: apply (drule subsetD, assumption) wenzelm@17200: apply (drule bspec, assumption) paulson@18143: apply (rule_tac x = "Union(c)" in bexI) wenzelm@17200: apply (rule ballI, rule impI) wenzelm@17200: apply (blast dest!: maxchain_Zorn, assumption) wenzelm@17200: done paulson@13551: paulson@13551: subsection{*Alternative version of Zorn's Lemma*} paulson@13551: paulson@13551: lemma Zorn_Lemma2: wenzelm@17200: "\c \ chain S. \y \ S. \x \ c. x \ y wenzelm@17200: ==> \y \ S. \x \ S. (y :: 'a set) \ x --> y = x" wenzelm@17200: apply (cut_tac Hausdorff maxchain_subset_chain) wenzelm@17200: apply (erule exE) wenzelm@17200: apply (drule subsetD, assumption) wenzelm@17200: apply (drule bspec, assumption, erule bexE) wenzelm@17200: apply (rule_tac x = y in bexI) wenzelm@17200: prefer 2 apply assumption wenzelm@17200: apply clarify wenzelm@17200: apply (rule ccontr) wenzelm@17200: apply (frule_tac z = x in chain_extend) wenzelm@17200: apply (assumption, blast) wenzelm@17200: apply (unfold maxchain_def super_def psubset_def) wenzelm@17200: apply (blast elim!: equalityCE) wenzelm@17200: done paulson@13551: paulson@13551: text{*Various other lemmas*} paulson@13551: paulson@13551: lemma chainD: "[| c \ chain S; x \ c; y \ c |] ==> x \ y | y \ x" wenzelm@17200: by (unfold chain_def) blast paulson@13551: paulson@13551: lemma chainD2: "!!(c :: 'a set set). c \ chain S ==> c \ S" wenzelm@17200: by (unfold chain_def) blast paulson@13551: paulson@13551: end