src/HOL/Hahn_Banach/Bounds.thy
changeset 55715 c4159fe6fa46
parent 45759 7ca82df6e951
child 59180 85ec71012df8
equal deleted inserted replaced
55714:326fd7103cb4 55715:c4159fe6fa46
    55   from ex obtain x where x: "lub A x" ..
    55   from ex obtain x where x: "lub A x" ..
    56   also from x have [symmetric]: "\<Squnion>A = x" ..
    56   also from x have [symmetric]: "\<Squnion>A = x" ..
    57   finally show ?thesis .
    57   finally show ?thesis .
    58 qed
    58 qed
    59 
    59 
    60 lemma lub_compat: "lub A x = isLub UNIV A x"
    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 proof -
    61   by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def)
    62   have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
       
    63     by (rule ext) (simp only: isUb_def)
       
    64   then show ?thesis
       
    65     by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
       
    66 qed
       
    67 
       
    68 lemma real_complete:
       
    69   fixes A :: "real set"
       
    70   assumes nonempty: "\<exists>a. a \<in> A"
       
    71     and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
       
    72   shows "\<exists>x. lub A x"
       
    73 proof -
       
    74   from ex_upper have "\<exists>y. isUb UNIV A y"
       
    75     unfolding isUb_def setle_def by blast
       
    76   with nonempty have "\<exists>x. isLub UNIV A x"
       
    77     by (rule reals_complete)
       
    78   then show ?thesis by (simp only: lub_compat)
       
    79 qed
       
    80 
    62 
    81 end
    63 end