author | paulson |
Sat, 31 Aug 2002 14:03:49 +0200 | |
changeset 13551 | b7f64ee8da84 |
child 13652 | 172600c40793 |
permissions | -rw-r--r-- |
13551
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
1 |
(* Title \<in> Zorn.thy |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
2 |
ID \<in> $Id$ |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
3 |
Author \<in> Jacques D. Fleuriot |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
4 |
Copyright \<in> 1998 University of Cambridge |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
5 |
Description \<in> Zorn's Lemma -- See Larry Paulson's Zorn.thy in ZF |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
6 |
*) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
7 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
8 |
header {*Zorn's Lemma*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
9 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
10 |
theory Zorn = Main: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
11 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
12 |
text{*The lemma and section numbers refer to an unpublished article ``Towards |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
13 |
the Mechanization of the Proofs of Some Classical Theorems of Set Theory,'' by |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
14 |
Abrial and Laffitte. *} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
15 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
16 |
constdefs |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
17 |
chain :: "'a::ord set => 'a set set" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
18 |
"chain S == {F. F \<subseteq> S & (\<forall>x \<in> F. \<forall>y \<in> F. x \<subseteq> y | y \<subseteq> x)}" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
19 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
20 |
super :: "['a::ord set,'a set] => 'a set set" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
21 |
"super S c == {d. d \<in> chain(S) & c < d}" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
22 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
23 |
maxchain :: "'a::ord set => 'a set set" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
24 |
"maxchain S == {c. c \<in> chain S & super S c = {}}" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
25 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
26 |
succ :: "['a::ord set,'a set] => 'a set" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
27 |
"succ S c == if (c \<notin> chain S| c \<in> maxchain S) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
28 |
then c else (@c'. c': (super S c))" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
29 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
30 |
consts |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
31 |
"TFin" :: "'a::ord set => 'a set set" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
32 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
33 |
inductive "TFin(S)" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
34 |
intros |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
35 |
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
|
36 |
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
|
37 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
38 |
monos Pow_mono |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
39 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
40 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
41 |
subsection{*Mathematical Preamble*} |
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 |
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
|
44 |
by blast |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
45 |
|
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 |
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
|
48 |
lemma Abrial_axiom1: "x \<subseteq> succ S x" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
49 |
apply (unfold succ_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
50 |
apply (rule split_if [THEN iffD2]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
51 |
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
|
52 |
apply (rule swap, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
53 |
apply (rule someI2, blast+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
54 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
55 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
56 |
lemmas TFin_UnionI = TFin.Pow_UnionI [OF PowI] |
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 |
lemma TFin_induct: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
59 |
"[| n \<in> TFin S; |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
60 |
!!x. [| x \<in> TFin S; P(x) |] ==> P(succ S x); |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
61 |
!!Y. [| Y \<subseteq> TFin S; Ball Y P |] ==> P(Union Y) |] |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
62 |
==> P(n)" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
63 |
apply (erule TFin.induct, blast+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
64 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
65 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
66 |
lemma succ_trans: "x \<subseteq> y ==> x \<subseteq> succ S y" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
67 |
apply (erule subset_trans) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
68 |
apply (rule Abrial_axiom1) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
69 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
70 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
71 |
text{*Lemma 1 of section 3.1*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
72 |
lemma TFin_linear_lemma1: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
73 |
"[| n \<in> TFin S; m \<in> TFin S; |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
74 |
\<forall>x \<in> TFin S. x \<subseteq> m --> x = m | succ S x \<subseteq> m |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
75 |
|] ==> n \<subseteq> m | succ S m \<subseteq> n" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
76 |
apply (erule TFin_induct) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
77 |
apply (erule_tac [2] Union_lemma0) txt{*or just Blast_tac*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
78 |
apply (blast del: subsetI intro: succ_trans) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
79 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
80 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
81 |
text{* Lemma 2 of section 3.2 *} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
82 |
lemma TFin_linear_lemma2: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
83 |
"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
|
84 |
apply (erule TFin_induct) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
85 |
apply (rule impI [THEN ballI]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
86 |
txt{*case split using TFin_linear_lemma1*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
87 |
apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
88 |
assumption+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
89 |
apply (drule_tac x = n in bspec, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
90 |
apply (blast del: subsetI intro: succ_trans, blast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
91 |
txt{*second induction step*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
92 |
apply (rule impI [THEN ballI]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
93 |
apply (rule Union_lemma0 [THEN disjE]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
94 |
apply (rule_tac [3] disjI2) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
95 |
prefer 2 apply blast |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
96 |
apply (rule ballI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
97 |
apply (rule_tac n1 = n and m1 = x in TFin_linear_lemma1 [THEN disjE], |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
98 |
assumption+, auto) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
99 |
apply (blast intro!: Abrial_axiom1 [THEN subsetD]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
100 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
101 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
102 |
text{*Re-ordering the premises of Lemma 2*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
103 |
lemma TFin_subsetD: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
104 |
"[| 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
|
105 |
apply (rule TFin_linear_lemma2 [rule_format]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
106 |
apply (assumption+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
107 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
108 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
109 |
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
|
110 |
lemma TFin_subset_linear: "[| m \<in> TFin S; n \<in> TFin S|] ==> n \<subseteq> m | m \<subseteq> n" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
111 |
apply (rule disjE) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
112 |
apply (rule TFin_linear_lemma1 [OF _ _TFin_linear_lemma2]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
113 |
apply (assumption+, erule disjI2) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
114 |
apply (blast del: subsetI |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
115 |
intro: subsetI Abrial_axiom1 [THEN subset_trans]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
116 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
117 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
118 |
text{*Lemma 3 of section 3.3*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
119 |
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
|
120 |
apply (erule TFin_induct) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
121 |
apply (drule TFin_subsetD) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
122 |
apply (assumption+, force, blast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
123 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
124 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
125 |
text{*Property 3.3 of section 3.3*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
126 |
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
|
127 |
apply (rule iffI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
128 |
apply (rule Union_upper [THEN equalityI]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
129 |
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
|
130 |
apply (assumption+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
131 |
apply (erule ssubst) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
132 |
apply (rule Abrial_axiom1 [THEN equalityI]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
133 |
apply (blast del: subsetI |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
134 |
intro: subsetI TFin_UnionI TFin.succI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
135 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
136 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
137 |
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
|
138 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
139 |
text{*NB: We assume the partial ordering is @{text "\<subseteq>"}, |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
140 |
the subset relation!*} |
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 |
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
|
143 |
by (unfold chain_def, auto) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
144 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
145 |
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
|
146 |
by (unfold super_def, fast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
147 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
148 |
lemma maxchain_subset_chain: "maxchain S \<subseteq> chain S" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
149 |
by (unfold maxchain_def, fast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
150 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
151 |
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
|
152 |
by (unfold super_def maxchain_def, auto) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
153 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
154 |
lemma select_super: "c \<in> chain S - maxchain S ==> |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
155 |
(@c'. c': super S c): super S c" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
156 |
apply (erule mem_super_Ex [THEN exE]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
157 |
apply (rule someI2, auto) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
158 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
159 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
160 |
lemma select_not_equals: "c \<in> chain S - maxchain S ==> |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
161 |
(@c'. c': super S c) \<noteq> c" |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
162 |
apply (rule notI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
163 |
apply (drule select_super) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
164 |
apply (simp add: super_def psubset_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
165 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
166 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
167 |
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
|
168 |
apply (unfold succ_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
169 |
apply (fast intro!: if_not_P) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
170 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
171 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
172 |
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
|
173 |
apply (frule succI3) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
174 |
apply (simp (no_asm_simp)) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
175 |
apply (rule select_not_equals, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
176 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
177 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
178 |
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
|
179 |
apply (erule TFin_induct) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
180 |
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
|
181 |
apply (unfold chain_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
182 |
apply (rule CollectI, safe) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
183 |
apply (drule bspec, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
184 |
apply (rule_tac [2] m1 = Xa and n1 = X in TFin_subset_linear [THEN disjE], |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
185 |
blast+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
186 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
187 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
188 |
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
|
189 |
apply (rule_tac x = "Union (TFin S) " in exI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
190 |
apply (rule classical) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
191 |
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
|
192 |
prefer 2 |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
193 |
apply (blast intro!: TFin_UnionI equal_succ_Union [THEN iffD2, symmetric]) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
194 |
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
|
195 |
apply (drule DiffI [THEN succ_not_equals], blast+) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
196 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
197 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
198 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
199 |
subsection{*Zorn's Lemma: If All Chains Have Upper Bounds Then |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
200 |
There Is a Maximal Element*} |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
201 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
202 |
lemma chain_extend: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
203 |
"[| c \<in> chain S; z \<in> S; |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
204 |
\<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
|
205 |
by (unfold chain_def, blast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
206 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
207 |
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
|
208 |
by (unfold chain_def, auto) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
209 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
210 |
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
|
211 |
by (unfold chain_def, auto) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
212 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
213 |
lemma maxchain_Zorn: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
214 |
"[| 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
|
215 |
apply (rule ccontr) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
216 |
apply (simp add: maxchain_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
217 |
apply (erule conjE) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
218 |
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
|
219 |
apply simp |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
220 |
apply (unfold super_def psubset_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
221 |
apply (blast intro: chain_extend dest: chain_Union_upper) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
222 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
223 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
224 |
theorem Zorn_Lemma: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
225 |
"\<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
|
226 |
apply (cut_tac Hausdorff maxchain_subset_chain) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
227 |
apply (erule exE) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
228 |
apply (drule subsetD, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
229 |
apply (drule bspec, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
230 |
apply (rule_tac x = "Union (c) " in bexI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
231 |
apply (rule ballI, rule impI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
232 |
apply (blast dest!: maxchain_Zorn, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
233 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
234 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
235 |
subsection{*Alternative version of Zorn's Lemma*} |
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 |
lemma Zorn_Lemma2: |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
238 |
"\<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
|
239 |
==> \<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
|
240 |
apply (cut_tac Hausdorff maxchain_subset_chain) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
241 |
apply (erule exE) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
242 |
apply (drule subsetD, assumption) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
243 |
apply (drule bspec, assumption, erule bexE) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
244 |
apply (rule_tac x = y in bexI) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
245 |
prefer 2 apply assumption |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
246 |
apply clarify |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
247 |
apply (rule ccontr) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
248 |
apply (frule_tac z = x in chain_extend) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
249 |
apply (assumption, blast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
250 |
apply (unfold maxchain_def super_def psubset_def) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
251 |
apply (blast elim!: equalityCE) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
252 |
done |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
253 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
254 |
text{*Various other lemmas*} |
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 |
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
|
257 |
by (unfold chain_def, blast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
258 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
259 |
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
|
260 |
by (unfold chain_def, blast) |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
261 |
|
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
262 |
end |
b7f64ee8da84
converted Hyperreal/Zorn to Isar format and moved to Library
paulson
parents:
diff
changeset
|
263 |