src/HOL/Hahn_Banach/Bounds.thy
author wenzelm
Mon, 24 Oct 2016 14:05:22 +0200
changeset 64371 213cf4215b40
parent 63970 3b6a3632e754
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
more robust;
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/Bounds.thy
7566
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     2
    Author:     Gertrud Bauer, TU Munich
c5a3f980a7af accomodate refined facts handling;
wenzelm
parents: 7535
diff changeset
     3
*)
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
     4
58889
5b7a9633cfa8 modernized header uniformly as section;
wenzelm
parents: 58745
diff changeset
     5
section \<open>Bounds\<close>
7808
fd019ac3485f update from Gertrud;
wenzelm
parents: 7656
diff changeset
     6
25596
ad9e3594f3f3 tuned header
haftmann
parents: 21404
diff changeset
     7
theory Bounds
63970
3b6a3632e754 HOL-Analysis: move Continuum_Not_Denumerable from Library
hoelzl
parents: 63464
diff changeset
     8
imports Main "~~/src/HOL/Analysis/Continuum_Not_Denumerable"
25596
ad9e3594f3f3 tuned header
haftmann
parents: 21404
diff changeset
     9
begin
7535
599d3414b51d The Hahn-Banach theorem for real vectorspaces (Isabelle/Isar)
wenzelm
parents:
diff changeset
    10
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    11
locale lub =
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    12
  fixes A and x
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    13
  assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    14
    and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    15
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    16
lemmas [elim?] = lub.least lub.upper
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    17
58745
wenzelm
parents: 58744
diff changeset
    18
definition the_lub :: "'a::order set \<Rightarrow> 'a"  ("\<Squnion>_" [90] 90)
44887
7ca82df6e951 misc tuning and clarification;
wenzelm
parents: 41413
diff changeset
    19
  where "the_lub A = The (lub A)"
14653
0848ab6fe5fc constdefs: proper order;
wenzelm
parents: 13515
diff changeset
    20
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    21
lemma the_lub_equality [elim?]:
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    22
  assumes "lub A x"
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    23
  shows "\<Squnion>A = (x::'a::order)"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    24
proof -
29234
60f7fb56f8cd More porting to new locales.
ballarin
parents: 29043
diff changeset
    25
  interpret lub A x by fact
27612
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    26
  show ?thesis
d3eb431db035 modernized specifications and proofs;
wenzelm
parents: 27611
diff changeset
    27
  proof (unfold the_lub_def)
58744
c434e37f290e update_cartouches;
wenzelm
parents: 54263
diff changeset
    28
    from \<open>lub A x\<close> show "The (lub A) = x"
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    29
    proof
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    30
      fix x' assume lub': "lub A x'"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    31
      show "x' = x"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    32
      proof (rule order_antisym)
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
    33
        from lub' show "x' \<le> x"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
    34
        proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    35
          fix a assume "a \<in> A"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    36
          then show "a \<le> x" ..
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
    37
        qed
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
    38
        show "x \<le> x'"
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
    39
        proof
27611
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    40
          fix a assume "a \<in> A"
2c01c0bdb385 Removed uses of context element includes.
ballarin
parents: 25596
diff changeset
    41
          with lub' show "a \<le> x'" ..
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 31795
diff changeset
    42
        qed
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    43
      qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    44
    qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    45
  qed
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    46
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    47
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    48
lemma the_lubI_ex:
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    49
  assumes ex: "\<exists>x. lub A x"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    50
  shows "lub A (\<Squnion>A)"
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    51
proof -
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    52
  from ex obtain x where x: "lub A x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    53
  also from x have [symmetric]: "\<Squnion>A = x" ..
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    54
  finally show ?thesis .
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    55
qed
7917
5e5b9813cce7 HahnBanach update by Gertrud Bauer;
wenzelm
parents: 7808
diff changeset
    56
54263
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 44887
diff changeset
    57
lemma real_complete: "\<exists>a::real. a \<in> A \<Longrightarrow> \<exists>y. \<forall>a \<in> A. a \<le> y \<Longrightarrow> \<exists>x. lub A x"
c4159fe6fa46 move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
hoelzl
parents: 44887
diff changeset
    58
  by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def)
13515
a6a7025fd7e8 updated to use locales (still some rough edges);
wenzelm
parents: 10687
diff changeset
    59
10687
c186279eecea tuned HOL/Real/HahnBanach;
wenzelm
parents: 9503
diff changeset
    60
end