src/HOL/Hahn_Banach/Zorn_Lemma.thy
author nipkow
Fri, 11 Dec 2020 17:58:01 +0100
changeset 72884 50f18a822ee9
parent 61879 e4f9d8f094fe
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
31795
be3e1cc5005c standard naming conventions for session and theories;
wenzelm
parents: 29252
diff changeset
     1
(*  Title:      HOL/Hahn_Banach/Zorn_Lemma.thy
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     2
    Author:     Gertrud Bauer, TU Munich
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     3
*)
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     4
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58744
diff changeset
     5
section \<open>Zorn's Lemma\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
     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
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
     9
begin
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    10
58744
c434e37f290e update_cartouches;
wenzelm
parents: 55018
diff changeset
    11
text \<open>
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    12
  Zorn's Lemmas states: if every linear ordered subset of an ordered set \<open>S\<close>
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    13
  has an upper bound in \<open>S\<close>, then there exists a maximal element in \<open>S\<close>. In
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    14
  our application, \<open>S\<close> is a set of sets ordered by set inclusion. Since the
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    15
  union of a chain of sets is an upper bound for all elements of the chain,
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    16
  the conditions of Zorn's lemma can be modified: if \<open>S\<close> is non-empty, it
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    17
  suffices to show that for every non-empty chain \<open>c\<close> in \<open>S\<close> the union of \<open>c\<close>
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    18
  also lies in \<open>S\<close>.
58744
c434e37f290e update_cartouches;
wenzelm
parents: 55018
diff changeset
    19
\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    20
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    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
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    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
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    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
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    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
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    29
    show "\<exists>y \<in> S. \<forall>z \<in> c. z \<subseteq> y"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    30
    proof cases
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    31
      txt \<open>If \<open>c\<close> is an empty chain, then every element in \<open>S\<close> is an upper
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    32
        bound of \<open>c\<close>.\<close>
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    33
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    34
      assume "c = {}"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    35
      with aS show ?thesis by fast
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    36
61879
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    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
e4f9d8f094fe tuned whitespace;
wenzelm
parents: 61539
diff changeset
    38
        \<open>S\<close>.\<close>
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    39
    next
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 23378
diff changeset
    40
      assume "c \<noteq> {}"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    41
      show ?thesis
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    42
      proof
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    43
        show "\<forall>z \<in> c. z \<subseteq> \<Union>c" by fast
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    44
        show "\<Union>c \<in> S"
9035
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    45
        proof (rule r)
58744
c434e37f290e update_cartouches;
wenzelm
parents: 55018
diff changeset
    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
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    48
        qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    49
      qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    50
    qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    51
  qed
371f023d3dbd removed explicit terminator (";");
wenzelm
parents: 8280
diff changeset
    52
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents:
diff changeset
    53
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9035
diff changeset
    54
end