src/HOL/Real/HahnBanach/Bounds.thy
changeset 29354 6ef5ddf22d3a
parent 29353 3d2e35c23c66
parent 29350 c7735554d291
child 29355 642cac18e155
child 29371 bab4e907d881
equal deleted inserted replaced
29353:3d2e35c23c66 29354:6ef5ddf22d3a
     1 (*  Title:      HOL/Real/HahnBanach/Bounds.thy
       
     2     ID:         $Id$
       
     3     Author:     Gertrud Bauer, TU Munich
       
     4 *)
       
     5 
       
     6 header {* Bounds *}
       
     7 
       
     8 theory Bounds
       
     9 imports Main ContNotDenum
       
    10 begin
       
    11 
       
    12 locale lub =
       
    13   fixes A and x
       
    14   assumes least [intro?]: "(\<And>a. a \<in> A \<Longrightarrow> a \<le> b) \<Longrightarrow> x \<le> b"
       
    15     and upper [intro?]: "a \<in> A \<Longrightarrow> a \<le> x"
       
    16 
       
    17 lemmas [elim?] = lub.least lub.upper
       
    18 
       
    19 definition
       
    20   the_lub :: "'a::order set \<Rightarrow> 'a" where
       
    21   "the_lub A = The (lub A)"
       
    22 
       
    23 notation (xsymbols)
       
    24   the_lub  ("\<Squnion>_" [90] 90)
       
    25 
       
    26 lemma the_lub_equality [elim?]:
       
    27   assumes "lub A x"
       
    28   shows "\<Squnion>A = (x::'a::order)"
       
    29 proof -
       
    30   interpret lub [A x] by fact
       
    31   show ?thesis
       
    32   proof (unfold the_lub_def)
       
    33     from `lub A x` show "The (lub A) = x"
       
    34     proof
       
    35       fix x' assume lub': "lub A x'"
       
    36       show "x' = x"
       
    37       proof (rule order_antisym)
       
    38 	from lub' show "x' \<le> x"
       
    39 	proof
       
    40           fix a assume "a \<in> A"
       
    41           then show "a \<le> x" ..
       
    42 	qed
       
    43 	show "x \<le> x'"
       
    44 	proof
       
    45           fix a assume "a \<in> A"
       
    46           with lub' show "a \<le> x'" ..
       
    47 	qed
       
    48       qed
       
    49     qed
       
    50   qed
       
    51 qed
       
    52 
       
    53 lemma the_lubI_ex:
       
    54   assumes ex: "\<exists>x. lub A x"
       
    55   shows "lub A (\<Squnion>A)"
       
    56 proof -
       
    57   from ex obtain x where x: "lub A x" ..
       
    58   also from x have [symmetric]: "\<Squnion>A = x" ..
       
    59   finally show ?thesis .
       
    60 qed
       
    61 
       
    62 lemma lub_compat: "lub A x = isLub UNIV A x"
       
    63 proof -
       
    64   have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
       
    65     by (rule ext) (simp only: isUb_def)
       
    66   then show ?thesis
       
    67     by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
       
    68 qed
       
    69 
       
    70 lemma real_complete:
       
    71   fixes A :: "real set"
       
    72   assumes nonempty: "\<exists>a. a \<in> A"
       
    73     and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
       
    74   shows "\<exists>x. lub A x"
       
    75 proof -
       
    76   from ex_upper have "\<exists>y. isUb UNIV A y"
       
    77     unfolding isUb_def setle_def by blast
       
    78   with nonempty have "\<exists>x. isLub UNIV A x"
       
    79     by (rule reals_complete)
       
    80   then show ?thesis by (simp only: lub_compat)
       
    81 qed
       
    82 
       
    83 end