7917
|
1 |
(* Title: HOL/Real/HahnBanach/ZornLemma.thy
|
|
2 |
ID: $Id$
|
|
3 |
Author: Gertrud Bauer, TU Munich
|
|
4 |
*)
|
|
5 |
|
9035
|
6 |
header {* Zorn's Lemma *}
|
7917
|
7 |
|
16417
|
8 |
theory ZornLemma imports Zorn begin
|
7917
|
9 |
|
10687
|
10 |
text {*
|
|
11 |
Zorn's Lemmas states: if every linear ordered subset of an ordered
|
|
12 |
set @{text S} has an upper bound in @{text S}, then there exists a
|
|
13 |
maximal element in @{text S}. In our application, @{text S} is a
|
|
14 |
set of sets ordered by set inclusion. Since the union of a chain of
|
|
15 |
sets is an upper bound for all elements of the chain, the conditions
|
|
16 |
of Zorn's lemma can be modified: if @{text S} is non-empty, it
|
|
17 |
suffices to show that for every non-empty chain @{text c} in @{text
|
|
18 |
S} the union of @{text c} also lies in @{text S}.
|
|
19 |
*}
|
7917
|
20 |
|
10687
|
21 |
theorem Zorn's_Lemma:
|
13515
|
22 |
assumes r: "\<And>c. c \<in> chain S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S"
|
|
23 |
and aS: "a \<in> S"
|
|
24 |
shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> y = z"
|
9035
|
25 |
proof (rule Zorn_Lemma2)
|
8104
|
26 |
txt_raw {* \footnote{See
|
14024
|
27 |
\url{http://isabelle.in.tum.de/library/HOL/HOL-Complex/Zorn.html}} \isanewline *}
|
10687
|
28 |
show "\<forall>c \<in> chain S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
|
9035
|
29 |
proof
|
10687
|
30 |
fix c assume "c \<in> chain S"
|
|
31 |
show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
|
9035
|
32 |
proof cases
|
13515
|
33 |
|
10687
|
34 |
txt {* If @{text c} is an empty chain, then every element in
|
|
35 |
@{text S} is an upper bound of @{text c}. *}
|
7917
|
36 |
|
13515
|
37 |
assume "c = {}"
|
9035
|
38 |
with aS show ?thesis by fast
|
7917
|
39 |
|
10687
|
40 |
txt {* If @{text c} is non-empty, then @{text "\<Union>c"} is an upper
|
|
41 |
bound of @{text c}, lying in @{text S}. *}
|
7917
|
42 |
|
9035
|
43 |
next
|
10687
|
44 |
assume c: "c \<noteq> {}"
|
13515
|
45 |
show ?thesis
|
|
46 |
proof
|
10687
|
47 |
show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
|
13515
|
48 |
show "\<Union>c \<in> S"
|
9035
|
49 |
proof (rule r)
|
13515
|
50 |
from c show "\<exists>x. x \<in> c" by fast
|
23378
|
51 |
show "c \<in> chain S" by fact
|
9035
|
52 |
qed
|
|
53 |
qed
|
|
54 |
qed
|
|
55 |
qed
|
|
56 |
qed
|
7917
|
57 |
|
10687
|
58 |
end
|