(* Title : HOL/Library/Zorn.thy
ID : $Id: Zorn.thy,v 1.6 2005/08/31 13:46:37 wenzelm Exp $
Author : Jacques D. Fleuriot
Description : Zorn's Lemma -- see Larry Paulson's Zorn.thy in ZF
*)
header {* Zorn's Lemma *}
theory Zorn
imports Main
begin
text{*
The lemma and section numbers refer to an unpublished article
\cite{Abrial-Laffitte}.
*}
constdefs
chain :: "'a set set => 'a set set set"
"chain S == {F. F ⊆ S & (∀x ∈ F. ∀y ∈ F. x ⊆ y | y ⊆ x)}"
super :: "['a set set,'a set set] => 'a set set set"
"super S c == {d. d ∈ chain S & c ⊂ d}"
maxchain :: "'a set set => 'a set set set"
"maxchain S == {c. c ∈ chain S & super S c = {}}"
succ :: "['a set set,'a set set] => 'a set set"
"succ S c ==
if c ∉ chain S | c ∈ maxchain S
then c else SOME c'. c' ∈ super S c"
consts
TFin :: "'a set set => 'a set set set"
inductive "TFin S"
intros
succI: "x ∈ TFin S ==> succ S x ∈ TFin S"
Pow_UnionI: "Y ∈ Pow(TFin S) ==> Union(Y) ∈ TFin S"
monos Pow_mono
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 (unfold succ_def)
apply (rule split_if [THEN iffD2])
apply (auto simp add: super_def maxchain_def psubset_def)
apply (rule swap, assumption)
apply (rule someI2, blast+)
done
lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI]
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)"
apply (erule TFin.induct)
apply blast+
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 (rule_tac [2] eq_succ_upper [THEN Union_least])
apply (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) 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 ==> ? 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, 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 psubset_def)
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)
apply (rule CollectI, safe)
apply (drule bspec, assumption)
apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE],
blast+)
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) blast
lemma chain_Union_upper: "[| c ∈ chain S; x ∈ c |] ==> x ⊆ Union(c)"
by (unfold chain_def) auto
lemma chain_ball_Union_upper: "c ∈ chain S ==> ∀x ∈ c. x ⊆ Union(c)"
by (unfold chain_def) 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 psubset_def)
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 psubset_def)
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) blast
lemma chainD2: "!!(c :: 'a set set). c ∈ chain S ==> c ⊆ S"
by (unfold chain_def) blast
end
lemma Union_lemma0:
∀x∈C. x ⊆ A ∨ B ⊆ x ==> Union C ⊆ A ∨ B ⊆ Union C
lemma Abrial_axiom1:
x ⊆ succ S x
lemmas TFin_UnionI:
Y ⊆ TFin S ==> Union Y ∈ TFin S
lemmas 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