src/HOL/Hahn_Banach/Bounds.thy
author hoelzl
Tue, 05 Nov 2013 09:45:02 +0100
changeset 55715 c4159fe6fa46
parent 45759 7ca82df6e951
child 59180 85ec71012df8
permissions -rw-r--r--
move Lubs from HOL to HOL-Library (replaced by conditionally complete lattices)
     1 (*  Title:      HOL/Hahn_Banach/Bounds.thy
     2     Author:     Gertrud Bauer, TU Munich
     3 *)
     4 
     5 header {* Bounds *}
     6 
     7 theory Bounds
     8 imports Main "~~/src/HOL/Library/ContNotDenum"
     9 begin
    10 
    11 locale lub =
    12   fixes A and x
    13   assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
    14     and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
    15 
    16 lemmas [elim?] = lub.least lub.upper
    17 
    18 definition the_lub :: "'a::order set \<Rightarrow> 'a"
    19   where "the_lub A = The (lub A)"
    20 
    21 notation (xsymbols)
    22   the_lub  ("\<Squnion>_" [90] 90)
    23 
    24 lemma the_lub_equality [elim?]:
    25   assumes "lub A x"
    26   shows "\<Squnion>A = (x::'a::order)"
    27 proof -
    28   interpret lub A x by fact
    29   show ?thesis
    30   proof (unfold the_lub_def)
    31     from `lub A x` show "The (lub A) = x"
    32     proof
    33       fix x' assume lub': "lub A x'"
    34       show "x' = x"
    35       proof (rule order_antisym)
    36         from lub' show "x' \<le> x"
    37         proof
    38           fix a assume "a \<in> A"
    39           then show "a \<le> x" ..
    40         qed
    41         show "x \<le> x'"
    42         proof
    43           fix a assume "a \<in> A"
    44           with lub' show "a \<le> x'" ..
    45         qed
    46       qed
    47     qed
    48   qed
    49 qed
    50 
    51 lemma the_lubI_ex:
    52   assumes ex: "\<exists>x. lub A x"
    53   shows "lub A (\<Squnion>A)"
    54 proof -
    55   from ex obtain x where x: "lub A x" ..
    56   also from x have [symmetric]: "\<Squnion>A = x" ..
    57   finally show ?thesis .
    58 qed
    59 
    60 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"
    61   by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def)
    62 
    63 end