src/HOL/Hahn_Banach/Bounds.thy
changeset 55715 c4159fe6fa46
parent 45759 7ca82df6e951
child 59180 85ec71012df8
     1.1 --- a/src/HOL/Hahn_Banach/Bounds.thy	Tue Nov 05 09:45:00 2013 +0100
     1.2 +++ b/src/HOL/Hahn_Banach/Bounds.thy	Tue Nov 05 09:45:02 2013 +0100
     1.3 @@ -57,25 +57,7 @@
     1.4    finally show ?thesis .
     1.5  qed
     1.6  
     1.7 -lemma lub_compat: "lub A x = isLub UNIV A x"
     1.8 -proof -
     1.9 -  have "isUb UNIV A = (\<lambda>x. A *<= x \<and> x \<in> UNIV)"
    1.10 -    by (rule ext) (simp only: isUb_def)
    1.11 -  then show ?thesis
    1.12 -    by (simp only: lub_def isLub_def leastP_def setge_def setle_def) blast
    1.13 -qed
    1.14 -
    1.15 -lemma real_complete:
    1.16 -  fixes A :: "real set"
    1.17 -  assumes nonempty: "\<exists>a. a \<in> A"
    1.18 -    and ex_upper: "\<exists>y. \<forall>a \<in> A. a \<le> y"
    1.19 -  shows "\<exists>x. lub A x"
    1.20 -proof -
    1.21 -  from ex_upper have "\<exists>y. isUb UNIV A y"
    1.22 -    unfolding isUb_def setle_def by blast
    1.23 -  with nonempty have "\<exists>x. isLub UNIV A x"
    1.24 -    by (rule reals_complete)
    1.25 -  then show ?thesis by (simp only: lub_compat)
    1.26 -qed
    1.27 +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"
    1.28 +  by (intro exI[of _ "Sup A"]) (auto intro!: cSup_upper cSup_least simp: lub_def)
    1.29  
    1.30  end