author | haftmann |
Thu, 11 Mar 2021 07:05:38 +0000 | |
changeset 73411 | 1f1366966296 |
parent 61879 | e4f9d8f094fe |
child 81464 | 2575f1bd226b |
permissions | -rw-r--r-- |
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
1 |
(* Title: HOL/Hahn_Banach/Zorn_Lemma.thy |
7917 | 2 |
Author: Gertrud Bauer, TU Munich |
3 |
*) |
|
4 |
||
58889 | 5 |
section \<open>Zorn's Lemma\<close> |
7917 | 6 |
|
31795
be3e1cc5005c
standard naming conventions for session and theories;
wenzelm
parents:
29252
diff
changeset
|
7 |
theory Zorn_Lemma |
55018
2a526bd279ed
moved 'Zorn' into 'Main', since it's a BNF dependency
blanchet
parents:
52183
diff
changeset
|
8 |
imports Main |
27612 | 9 |
begin |
7917 | 10 |
|
58744 | 11 |
text \<open> |
61879 | 12 |
Zorn's Lemmas states: if every linear ordered subset of an ordered set \<open>S\<close> |
13 |
has an upper bound in \<open>S\<close>, then there exists a maximal element in \<open>S\<close>. In |
|
14 |
our application, \<open>S\<close> is a set of sets ordered by set inclusion. Since the |
|
15 |
union of a chain of sets is an upper bound for all elements of the chain, |
|
16 |
the conditions of Zorn's lemma can be modified: if \<open>S\<close> is non-empty, it |
|
17 |
suffices to show that for every non-empty chain \<open>c\<close> in \<open>S\<close> the union of \<open>c\<close> |
|
18 |
also lies in \<open>S\<close>. |
|
58744 | 19 |
\<close> |
7917 | 20 |
|
10687 | 21 |
theorem Zorn's_Lemma: |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
44887
diff
changeset
|
22 |
assumes r: "\<And>c. c \<in> chains S \<Longrightarrow> \<exists>x. x \<in> c \<Longrightarrow> \<Union>c \<in> S" |
13515 | 23 |
and aS: "a \<in> S" |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
44887
diff
changeset
|
24 |
shows "\<exists>y \<in> S. \<forall>z \<in> S. y \<subseteq> z \<longrightarrow> z = y" |
9035 | 25 |
proof (rule Zorn_Lemma2) |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
44887
diff
changeset
|
26 |
show "\<forall>c \<in> chains S. \<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y" |
9035 | 27 |
proof |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
44887
diff
changeset
|
28 |
fix c assume "c \<in> chains S" |
10687 | 29 |
show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y" |
9035 | 30 |
proof cases |
61879 | 31 |
txt \<open>If \<open>c\<close> is an empty chain, then every element in \<open>S\<close> is an upper |
32 |
bound of \<open>c\<close>.\<close> |
|
7917 | 33 |
|
13515 | 34 |
assume "c = {}" |
9035 | 35 |
with aS show ?thesis by fast |
7917 | 36 |
|
61879 | 37 |
txt \<open>If \<open>c\<close> is non-empty, then \<open>\<Union>c\<close> is an upper bound of \<open>c\<close>, lying in |
38 |
\<open>S\<close>.\<close> |
|
9035 | 39 |
next |
27612 | 40 |
assume "c \<noteq> {}" |
13515 | 41 |
show ?thesis |
42 |
proof |
|
10687 | 43 |
show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast |
13515 | 44 |
show "\<Union>c \<in> S" |
9035 | 45 |
proof (rule r) |
58744 | 46 |
from \<open>c \<noteq> {}\<close> show "\<exists>x. x \<in> c" by fast |
52183
667961fa6a60
fixed files broken due to Zorn changes (cf. 59e5dd7b8f9a)
popescua
parents:
44887
diff
changeset
|
47 |
show "c \<in> chains S" by fact |
9035 | 48 |
qed |
49 |
qed |
|
50 |
qed |
|
51 |
qed |
|
52 |
qed |
|
7917 | 53 |
|
10687 | 54 |
end |