use bdd_above and bdd_below for conditionally complete lattices
authorhoelzl
Tue, 05 Nov 2013 09:44:58 +0100
changeset 54258 adfc759263ab
parent 54257 5c7a3b6b05a9
child 54259 71c701dc5bf9
use bdd_above and bdd_below for conditionally complete lattices
src/HOL/Conditionally_Complete_Lattices.thy
src/HOL/Library/FSet.thy
src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy
src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy
src/HOL/Multivariate_Analysis/Integration.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/Real.thy
src/HOL/Topological_Spaces.thy
--- a/src/HOL/Conditionally_Complete_Lattices.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Conditionally_Complete_Lattices.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -1,6 +1,7 @@
 (*  Title:      HOL/Conditionally_Complete_Lattices.thy
     Author:     Amine Chaieb and L C Paulson, University of Cambridge
     Author:     Johannes Hölzl, TU München
+    Author:     Luke S. Serafin, Carnegie Mellon University
 *)
 
 header {* Conditionally-complete Lattices *}
@@ -15,6 +16,118 @@
 lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
   by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
 
+context preorder
+begin
+
+definition "bdd_above A \<longleftrightarrow> (\<exists>M. \<forall>x \<in> A. x \<le> M)"
+definition "bdd_below A \<longleftrightarrow> (\<exists>m. \<forall>x \<in> A. m \<le> x)"
+
+lemma bdd_aboveI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> x \<le> M) \<Longrightarrow> bdd_above A"
+  by (auto simp: bdd_above_def)
+
+lemma bdd_belowI[intro]: "(\<And>x. x \<in> A \<Longrightarrow> m \<le> x) \<Longrightarrow> bdd_below A"
+  by (auto simp: bdd_below_def)
+
+lemma bdd_above_empty [simp, intro]: "bdd_above {}"
+  unfolding bdd_above_def by auto
+
+lemma bdd_below_empty [simp, intro]: "bdd_below {}"
+  unfolding bdd_below_def by auto
+
+lemma bdd_above_mono: "bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_above A"
+  by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD)
+
+lemma bdd_below_mono: "bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> bdd_below A"
+  by (metis bdd_below_def order_class.le_neq_trans psubsetD)
+
+lemma bdd_above_Int1 [simp]: "bdd_above A \<Longrightarrow> bdd_above (A \<inter> B)"
+  using bdd_above_mono by auto
+
+lemma bdd_above_Int2 [simp]: "bdd_above B \<Longrightarrow> bdd_above (A \<inter> B)"
+  using bdd_above_mono by auto
+
+lemma bdd_below_Int1 [simp]: "bdd_below A \<Longrightarrow> bdd_below (A \<inter> B)"
+  using bdd_below_mono by auto
+
+lemma bdd_below_Int2 [simp]: "bdd_below B \<Longrightarrow> bdd_below (A \<inter> B)"
+  using bdd_below_mono by auto
+
+lemma bdd_above_Ioo [simp, intro]: "bdd_above {a <..< b}"
+  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Ico [simp, intro]: "bdd_above {a ..< b}"
+  by (auto simp add: bdd_above_def intro!: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Iio [simp, intro]: "bdd_above {..< b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Ioc [simp, intro]: "bdd_above {a <.. b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Icc [simp, intro]: "bdd_above {a .. b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_above_Iic [simp, intro]: "bdd_above {.. b}"
+  by (auto simp add: bdd_above_def intro: exI[of _ b] less_imp_le)
+
+lemma bdd_below_Ioo [simp, intro]: "bdd_below {a <..< b}"
+  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ioc [simp, intro]: "bdd_below {a <.. b}"
+  by (auto simp add: bdd_below_def intro!: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ioi [simp, intro]: "bdd_below {a <..}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ico [simp, intro]: "bdd_below {a ..< b}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Icc [simp, intro]: "bdd_below {a .. b}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+lemma bdd_below_Ici [simp, intro]: "bdd_below {a ..}"
+  by (auto simp add: bdd_below_def intro: exI[of _ a] less_imp_le)
+
+end
+
+context lattice
+begin
+
+lemma bdd_above_insert [simp]: "bdd_above (insert a A) = bdd_above A"
+  by (auto simp: bdd_above_def intro: le_supI2 sup_ge1)
+
+lemma bdd_below_insert [simp]: "bdd_below (insert a A) = bdd_below A"
+  by (auto simp: bdd_below_def intro: le_infI2 inf_le1)
+
+lemma bdd_finite [simp]:
+  assumes "finite A" shows bdd_above_finite: "bdd_above A" and bdd_below_finite: "bdd_below A"
+  using assms by (induct rule: finite_induct, auto)
+
+lemma bdd_above_Un [simp]: "bdd_above (A \<union> B) = (bdd_above A \<and> bdd_above B)"
+proof
+  assume "bdd_above (A \<union> B)"
+  thus "bdd_above A \<and> bdd_above B" unfolding bdd_above_def by auto
+next
+  assume "bdd_above A \<and> bdd_above B"
+  then obtain a b where "\<forall>x\<in>A. x \<le> a" "\<forall>x\<in>B. x \<le> b" unfolding bdd_above_def by auto
+  hence "\<forall>x \<in> A \<union> B. x \<le> sup a b" by (auto intro: Un_iff le_supI1 le_supI2)
+  thus "bdd_above (A \<union> B)" unfolding bdd_above_def ..
+qed
+
+lemma bdd_below_Un [simp]: "bdd_below (A \<union> B) = (bdd_below A \<and> bdd_below B)"
+proof
+  assume "bdd_below (A \<union> B)"
+  thus "bdd_below A \<and> bdd_below B" unfolding bdd_below_def by auto
+next
+  assume "bdd_below A \<and> bdd_below B"
+  then obtain a b where "\<forall>x\<in>A. a \<le> x" "\<forall>x\<in>B. b \<le> x" unfolding bdd_below_def by auto
+  hence "\<forall>x \<in> A \<union> B. inf a b \<le> x" by (auto intro: Un_iff le_infI1 le_infI2)
+  thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
+qed
+
+end
+
+
 text {*
 
 To avoid name classes with the @{class complete_lattice}-class we prefix @{const Sup} and
@@ -23,24 +136,22 @@
 *}
 
 class conditionally_complete_lattice = lattice + Sup + Inf +
-  assumes cInf_lower: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> z \<le> a) \<Longrightarrow> Inf X \<le> x"
+  assumes cInf_lower: "x \<in> X \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> x"
     and cInf_greatest: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> z \<le> Inf X"
-  assumes cSup_upper: "x \<in> X \<Longrightarrow> (\<And>a. a \<in> X \<Longrightarrow> a \<le> z) \<Longrightarrow> x \<le> Sup X"
+  assumes cSup_upper: "x \<in> X \<Longrightarrow> bdd_above X \<Longrightarrow> x \<le> Sup X"
     and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
 begin
 
-lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*)
-  "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
-  by (blast intro: antisym cSup_upper cSup_least)
+lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
+  by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
 
-lemma cInf_eq_minimum: (*REAL_INF_MIN in HOL4*)
-  "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
-  by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto
+lemma cInf_eq_minimum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X = z"
+  by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto
 
-lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> a \<le> z) \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
+lemma cSup_le_iff: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S \<le> a \<longleftrightarrow> (\<forall>x\<in>S. x \<le> a)"
   by (metis order_trans cSup_upper cSup_least)
 
-lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> (\<And>a. a \<in> S \<Longrightarrow> z \<le> a) \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
+lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
   by (metis order_trans cInf_lower cInf_greatest)
 
 lemma cSup_singleton [simp]: "Sup {x} = x"
@@ -49,20 +160,12 @@
 lemma cInf_singleton [simp]: "Inf {x} = x"
   by (intro cInf_eq_minimum) auto
 
-lemma cSup_upper2: (*REAL_IMP_LE_SUP in HOL4*)
-  "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y \<le> Sup X"
+lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
   by (metis cSup_upper order_trans)
  
-lemma cInf_lower2:
-  "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X \<le> y"
+lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
   by (metis cInf_lower order_trans)
 
-lemma cSup_upper_EX: "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> x \<le> z \<Longrightarrow> x \<le> Sup X"
-  by (blast intro: cSup_upper)
-
-lemma cInf_lower_EX:  "x \<in> X \<Longrightarrow> \<exists>z. \<forall>x. x \<in> X \<longrightarrow> z \<le> x \<Longrightarrow> Inf X \<le> x"
-  by (blast intro: cInf_lower)
-
 lemma cSup_eq_non_empty:
   assumes 1: "X \<noteq> {}"
   assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
@@ -77,67 +180,41 @@
   shows "Inf X = a"
   by (intro 3 1 antisym cInf_greatest) (auto intro: 2 1 cInf_lower)
 
-lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
-  by (rule cInf_eq_non_empty) (auto intro: cSup_upper cSup_least)
-
-lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
-  by (rule cSup_eq_non_empty) (auto intro: cInf_lower cInf_greatest)
+lemma cInf_cSup: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> Inf S = Sup {x. \<forall>s\<in>S. x \<le> s}"
+  by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def)
 
-lemma cSup_insert: 
-  assumes x: "X \<noteq> {}"
-      and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
-  shows "Sup (insert a X) = sup a (Sup X)"
-proof (intro cSup_eq_non_empty)
-  fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> x \<le> y" with x show "sup a (Sup X) \<le> y" by (auto intro: cSup_least)
-qed (auto intro: le_supI2 z cSup_upper)
+lemma cSup_cInf: "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> Sup S = Inf {x. \<forall>s\<in>S. s \<le> x}"
+  by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def)
+
+lemma cSup_insert: "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> Sup (insert a X) = sup a (Sup X)"
+  by (intro cSup_eq_non_empty) (auto intro: le_supI2 cSup_upper cSup_least)
 
-lemma cInf_insert: 
-  assumes x: "X \<noteq> {}"
-      and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
-  shows "Inf (insert a X) = inf a (Inf X)"
-proof (intro cInf_eq_non_empty)
-  fix y assume "\<And>x. x \<in> insert a X \<Longrightarrow> y \<le> x" with x show "y \<le> inf a (Inf X)" by (auto intro: cInf_greatest)
-qed (auto intro: le_infI2 z cInf_lower)
+lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
+  by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
 
-lemma cSup_insert_If: 
-  "(\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
-  using cSup_insert[of X z] by simp
+lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
+  using cSup_insert[of X] by simp
 
-lemma cInf_insert_if: 
-  "(\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
-  using cInf_insert[of X z] by simp
+lemma cInf_insert_if: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
+  using cInf_insert[of X] by simp
 
 lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
 proof (induct X arbitrary: x rule: finite_induct)
   case (insert x X y) then show ?case
-    apply (cases "X = {}")
-    apply simp
-    apply (subst cSup_insert[of _ "Sup X"])
-    apply (auto intro: le_supI2)
-    done
+    by (cases "X = {}") (auto simp: cSup_insert intro: le_supI2)
 qed simp
 
 lemma cInf_le_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> Inf X \<le> x"
 proof (induct X arbitrary: x rule: finite_induct)
   case (insert x X y) then show ?case
-    apply (cases "X = {}")
-    apply simp
-    apply (subst cInf_insert[of _ "Inf X"])
-    apply (auto intro: le_infI2)
-    done
+    by (cases "X = {}") (auto simp: cInf_insert intro: le_infI2)
 qed simp
 
 lemma cSup_eq_Sup_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Sup_fin X"
-proof (induct X rule: finite_ne_induct)
-  case (insert x X) then show ?case
-    using cSup_insert[of X "Sup_fin X" x] le_cSup_finite[of X] by simp
-qed simp
+  by (induct X rule: finite_ne_induct) (simp_all add: cSup_insert)
 
 lemma cInf_eq_Inf_fin: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Inf_fin X"
-proof (induct X rule: finite_ne_induct)
-  case (insert x X) then show ?case
-    using cInf_insert[of X "Inf_fin X" x] cInf_le_finite[of X] by simp
-qed simp
+  by (induct X rule: finite_ne_induct) (simp_all add: cInf_insert)
 
 lemma cSup_atMost[simp]: "Sup {..x} = x"
   by (auto intro!: cSup_eq_maximum)
@@ -165,7 +242,7 @@
 lemma isLub_cSup: 
   "(S::'a :: conditionally_complete_lattice set) \<noteq> {} \<Longrightarrow> (\<exists>b. S *<= b) \<Longrightarrow> isLub UNIV S (Sup S)"
   by  (auto simp add: isLub_def setle_def leastP_def isUb_def
-            intro!: setgeI intro: cSup_upper cSup_least)
+            intro!: setgeI cSup_upper cSup_least)
 
 lemma cSup_eq:
   fixes a :: "'a :: {conditionally_complete_lattice, no_bot}"
@@ -195,10 +272,10 @@
 begin
 
 lemma less_cSup_iff : (*REAL_SUP_LE in HOL4*)
-  "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
+  "X \<noteq> {} \<Longrightarrow> bdd_above X \<Longrightarrow> y < Sup X \<longleftrightarrow> (\<exists>x\<in>X. y < x)"
   by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans)
 
-lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> z \<le> x) \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
+lemma cInf_less_iff: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X < y \<longleftrightarrow> (\<exists>x\<in>X. x < y)"
   by (rule iffI) (metis cInf_greatest not_less, metis cInf_lower le_less_trans)
 
 lemma less_cSupE:
@@ -207,11 +284,11 @@
 
 lemma less_cSupD:
   "X \<noteq> {} \<Longrightarrow> z < Sup X \<Longrightarrow> \<exists>x\<in>X. z < x"
-  by (metis less_cSup_iff not_leE)
+  by (metis less_cSup_iff not_leE bdd_above_def)
 
 lemma cInf_lessD:
   "X \<noteq> {} \<Longrightarrow> Inf X < z \<Longrightarrow> \<exists>x\<in>X. x < z"
-  by (metis cInf_less_iff not_leE)
+  by (metis cInf_less_iff not_leE bdd_below_def)
 
 lemma complete_interval:
   assumes "a < b" and "P a" and "\<not> P b"
@@ -219,7 +296,7 @@
              (\<forall>d. (\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x) \<longrightarrow> d \<le> c)"
 proof (rule exI [where x = "Sup {d. \<forall>x. a \<le> x & x < d --> P x}"], auto)
   show "a \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
-    by (rule cSup_upper [where z=b], auto)
+    by (rule cSup_upper, auto simp: bdd_above_def)
        (metis `a < b` `\<not> P b` linear less_le)
 next
   show "Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c} \<le> b"
@@ -240,7 +317,7 @@
   fix d
     assume 0: "\<forall>x. a \<le> x \<and> x < d \<longrightarrow> P x"
     thus "d \<le> Sup {d. \<forall>c. a \<le> c \<and> c < d \<longrightarrow> P c}"
-      by (rule_tac z="b" in cSup_upper, auto) 
+      by (rule_tac cSup_upper, auto simp: bdd_above_def)
          (metis `a<b` `~ P b` linear less_le)
 qed
 
--- a/src/HOL/Library/FSet.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Library/FSet.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -101,19 +101,25 @@
 lemma finite_Sup: "\<exists>z. finite z \<and> (\<forall>a. a \<in> X \<longrightarrow> a \<le> z) \<Longrightarrow> finite (Sup X)"
 by (auto intro: finite_subset)
 
+lemma transfer_bdd_below[transfer_rule]: "(set_rel (pcr_fset op =) ===> op =) bdd_below bdd_below"
+  by auto
+
 instance
 proof 
   fix x z :: "'a fset"
   fix X :: "'a fset set"
   {
-    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> z |\<subseteq>| a)" 
+    assume "x \<in> X" "bdd_below X" 
     then show "Inf X |\<subseteq>| x"  by transfer auto
   next
     assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> z |\<subseteq>| x)"
     then show "z |\<subseteq>| Inf X" by transfer (clarsimp, blast)
   next
-    assume "x \<in> X" "(\<And>a. a \<in> X \<Longrightarrow> a |\<subseteq>| z)"
-    then show "x |\<subseteq>| Sup X" by transfer (auto intro!: finite_Sup)
+    assume "x \<in> X" "bdd_above X"
+    then obtain z where "x \<in> X" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
+      by (auto simp: bdd_above_def)
+    then show "x |\<subseteq>| Sup X"
+      by transfer (auto intro!: finite_Sup)
   next
     assume "X \<noteq> {}" "(\<And>x. x \<in> X \<Longrightarrow> x |\<subseteq>| z)"
     then show "Sup X |\<subseteq>| z" by transfer (clarsimp, blast)
--- a/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -8724,7 +8724,7 @@
     using interior_subset[of I] `x \<in> interior I` by auto
 
   have "Inf (?F x) \<le> (f x - f y) / (x - y)"
-  proof (rule cInf_lower2)
+  proof (intro bdd_belowI cInf_lower2)
     show "(f x - f t) / (x - t) \<in> ?F x"
       using `t \<in> I` `x < t` by auto
     show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -660,7 +660,7 @@
     assume "S \<noteq> {}"
     { assume ex: "\<exists>B. \<forall>x\<in>S. B \<le> x"
       then have *: "\<forall>x\<in>S. Inf S \<le> x"
-        using cInf_lower_EX[of _ S] ex by metis
+        using cInf_lower[of _ S] ex by (metis bdd_below_def)
       then have "Inf S \<in> S"
         apply (subst closed_contains_Inf)
         using ex `S \<noteq> {}` `closed S`
--- a/src/HOL/Multivariate_Analysis/Integration.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Integration.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -13,7 +13,7 @@
 lemma cSup_abs_le: (* TODO: is this really needed? *)
   fixes S :: "real set"
   shows "S \<noteq> {} \<Longrightarrow> (\<forall>x\<in>S. \<bar>x\<bar> \<le> a) \<Longrightarrow> \<bar>Sup S\<bar> \<le> a"
-  by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2)
+  by (auto simp add: abs_le_interval_iff intro: cSup_least) (metis cSup_upper2 bdd_aboveI)
 
 lemma cInf_abs_ge: (* TODO: is this really needed? *)
   fixes S :: "real set"
@@ -86,7 +86,7 @@
   apply (insert assms)
   apply (erule exE)
   apply (rule_tac x = b in exI)
-  apply (auto simp: isLb_def setge_def intro: cInf_lower cInf_greatest)
+  apply (auto simp: isLb_def setge_def intro!: cInf_lower cInf_greatest)
   done
 
 lemma real_ge_sup_subset:
@@ -100,7 +100,7 @@
   apply (insert assms)
   apply (erule exE)
   apply (rule_tac x = b in exI)
-  apply (auto simp: isUb_def setle_def intro: cSup_upper cSup_least)
+  apply (auto simp: isUb_def setle_def intro!: cSup_upper cSup_least)
   done
 
 (*declare not_less[simp] not_le[simp]*)
@@ -12728,8 +12728,8 @@
           assume x: "x \<in> s"
           have *: "\<And>x y::real. x \<ge> - y \<Longrightarrow> - x \<le> y" by auto
           show "Inf {f j x |j. n \<le> j} \<le> f n x"
-            apply (rule cInf_lower[where z="- h x"])
-            defer
+            apply (intro cInf_lower bdd_belowI)
+            apply auto []
             apply (rule *)
             using assms(3)[rule_format,OF x]
             unfolding real_norm_def abs_le_iff
@@ -12741,8 +12741,7 @@
           fix x
           assume x: "x \<in> s"
           show "f n x \<le> Sup {f j x |j. n \<le> j}"
-            apply (rule cSup_upper[where z="h x"])
-            defer
+            apply (rule cSup_upper)
             using assms(3)[rule_format,OF x]
             unfolding real_norm_def abs_le_iff
             apply auto
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -1909,17 +1909,17 @@
 
 lemma closure_contains_Inf:
   fixes S :: "real set"
-  assumes "S \<noteq> {}" "\<forall>x\<in>S. B \<le> x"
+  assumes "S \<noteq> {}" "bdd_below S"
   shows "Inf S \<in> closure S"
 proof -
   have *: "\<forall>x\<in>S. Inf S \<le> x"
-    using cInf_lower_EX[of _ S] assms by metis
+    using cInf_lower[of _ S] assms by metis
   {
     fix e :: real
     assume "e > 0"
     then have "Inf S < Inf S + e" by simp
     with assms obtain x where "x \<in> S" "x < Inf S + e"
-      by (subst (asm) cInf_less_iff[of _ B]) auto
+      by (subst (asm) cInf_less_iff) auto
     with * have "\<exists>x\<in>S. dist x (Inf S) < e"
       by (intro bexI[of _ x]) (auto simp add: dist_real_def)
   }
@@ -1928,12 +1928,9 @@
 
 lemma closed_contains_Inf:
   fixes S :: "real set"
-  assumes "S \<noteq> {}" "\<forall>x\<in>S. B \<le> x"
-    and "closed S"
-  shows "Inf S \<in> S"
+  shows "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> closed S \<Longrightarrow> Inf S \<in> S"
   by (metis closure_contains_Inf closure_closed assms)
 
-
 lemma not_trivial_limit_within_ball:
   "\<not> trivial_limit (at x within S) \<longleftrightarrow> (\<forall>e>0. S \<inter> ball x e - {x} \<noteq> {})"
   (is "?lhs = ?rhs")
@@ -1977,27 +1974,20 @@
 
 definition "infdist x A = (if A = {} then 0 else Inf {dist x a|a. a \<in> A})"
 
+lemma bdd_below_infdist[intro, simp]: "bdd_below {dist x a|a. a \<in> A}"
+  by (auto intro!: zero_le_dist)
+
 lemma infdist_notempty: "A \<noteq> {} \<Longrightarrow> infdist x A = Inf {dist x a|a. a \<in> A}"
   by (simp add: infdist_def)
 
 lemma infdist_nonneg: "0 \<le> infdist x A"
   by (auto simp add: infdist_def intro: cInf_greatest)
 
-lemma infdist_le:
-  assumes "a \<in> A"
-    and "d = dist x a"
-  shows "infdist x A \<le> d"
-  using assms by (auto intro!: cInf_lower[where z=0] simp add: infdist_def)
-
-lemma infdist_zero[simp]:
-  assumes "a \<in> A"
-  shows "infdist a A = 0"
-proof -
-  from infdist_le[OF assms, of "dist a a"] have "infdist a A \<le> 0"
-    by auto
-  with infdist_nonneg[of a A] assms show "infdist a A = 0"
-    by auto
-qed
+lemma infdist_le: "a \<in> A \<Longrightarrow> d = dist x a \<Longrightarrow> infdist x A \<le> d"
+  using assms by (auto intro: cInf_lower simp add: infdist_def)
+
+lemma infdist_zero[simp]: "a \<in> A \<Longrightarrow> infdist a A = 0"
+  by (auto intro!: antisym infdist_nonneg infdist_le)
 
 lemma infdist_triangle: "infdist x A \<le> infdist y A + dist x y"
 proof (cases "A = {}")
@@ -2021,13 +2011,7 @@
         using `a \<in> A` by auto
       show "dist x a \<le> d"
         unfolding d by (rule dist_triangle)
-      fix d
-      assume "d \<in> {dist x a |a. a \<in> A}"
-      then obtain a where "a \<in> A" "d = dist x a"
-        by auto
-      then show "infdist x A \<le> d"
-        by (rule infdist_le)
-    qed
+    qed simp
   qed
   also have "\<dots> = dist x y + infdist y A"
   proof (rule cInf_eq, safe)
@@ -2651,11 +2635,19 @@
 
 text{* Some theorems on sups and infs using the notion "bounded". *}
 
-lemma bounded_real:
-  fixes S :: "real set"
-  shows "bounded S \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. abs x \<le> a)"
+lemma bounded_real: "bounded (S::real set) \<longleftrightarrow> (\<exists>a. \<forall>x\<in>S. \<bar>x\<bar> \<le> a)"
   by (simp add: bounded_iff)
 
+lemma bounded_imp_bdd_above: "bounded S \<Longrightarrow> bdd_above (S :: real set)"
+  by (auto simp: bounded_def bdd_above_def dist_real_def)
+     (metis abs_le_D1 abs_minus_commute diff_le_eq)
+
+lemma bounded_imp_bdd_below: "bounded S \<Longrightarrow> bdd_below (S :: real set)"
+  by (auto simp: bounded_def bdd_below_def dist_real_def)
+     (metis abs_le_D1 add_commute diff_le_eq)
+
+(* TODO: remove the following lemmas about Inf and Sup, is now in conditionally complete lattice *)
+
 lemma bounded_has_Sup:
   fixes S :: "real set"
   assumes "bounded S"
@@ -2663,22 +2655,14 @@
   shows "\<forall>x\<in>S. x \<le> Sup S"
     and "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
 proof
-  fix x
-  assume "x\<in>S"
-  then show "x \<le> Sup S"
-    by (metis cSup_upper abs_le_D1 assms(1) bounded_real)
-next
   show "\<forall>b. (\<forall>x\<in>S. x \<le> b) \<longrightarrow> Sup S \<le> b"
     using assms by (metis cSup_least)
-qed
+qed (metis cSup_upper assms(1) bounded_imp_bdd_above)
 
 lemma Sup_insert:
   fixes S :: "real set"
   shows "bounded S \<Longrightarrow> Sup (insert x S) = (if S = {} then x else max x (Sup S))"
-  apply (subst cSup_insert_If)
-  apply (rule bounded_has_Sup(1)[of S, rule_format])
-  apply (auto simp: sup_max)
-  done
+  by (auto simp: bounded_imp_bdd_above sup_max cSup_insert_If)
 
 lemma Sup_insert_finite:
   fixes S :: "real set"
@@ -2695,24 +2679,14 @@
   shows "\<forall>x\<in>S. x \<ge> Inf S"
     and "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
 proof
-  fix x
-  assume "x \<in> S"
-  from assms(1) obtain a where a: "\<forall>x\<in>S. \<bar>x\<bar> \<le> a"
-    unfolding bounded_real by auto
-  then show "x \<ge> Inf S" using `x \<in> S`
-    by (metis cInf_lower_EX abs_le_D2 minus_le_iff)
-next
   show "\<forall>b. (\<forall>x\<in>S. x \<ge> b) \<longrightarrow> Inf S \<ge> b"
     using assms by (metis cInf_greatest)
-qed
+qed (metis cInf_lower assms(1) bounded_imp_bdd_below)
 
 lemma Inf_insert:
   fixes S :: "real set"
   shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
-  apply (subst cInf_insert_if)
-  apply (rule bounded_has_Inf(1)[of S, rule_format])
-  apply (auto simp: inf_min)
-  done
+  by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_if)
 
 lemma Inf_insert_finite:
   fixes S :: "real set"
@@ -5738,12 +5712,16 @@
   from s obtain z d where z: "\<And>x. x \<in> s \<Longrightarrow> dist z x \<le> d"
     unfolding bounded_def by auto
   have "dist x y \<le> Sup ?D"
-  proof (rule cSup_upper, safe)
-    fix a b
-    assume "a \<in> s" "b \<in> s"
-    with z[of a] z[of b] dist_triangle[of a b z]
-    show "dist a b \<le> 2 * d"
-      by (simp add: dist_commute)
+  proof (rule cSup_upper)
+    show "bdd_above ?D"
+      unfolding bdd_above_def
+    proof (safe intro!: exI)
+      fix a b
+      assume "a \<in> s" "b \<in> s"
+      with z[of a] z[of b] dist_triangle[of a b z]
+      show "dist a b \<le> 2 * d"
+        by (simp add: dist_commute)
+    qed
   qed (insert s, auto)
   with `x \<in> s` show ?thesis
     by (auto simp add: diameter_def)
--- a/src/HOL/Real.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Real.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -919,6 +919,12 @@
     using 1 2 3 by (rule_tac x="Real B" in exI, simp)
 qed
 
+(* TODO: generalize to ordered group *)
+lemma bdd_above_uminus[simp]: "bdd_above (uminus ` X) \<longleftrightarrow> bdd_below (X::real set)"
+  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
+
+lemma bdd_below_uminus[simp]: "bdd_below (uminus ` X) \<longleftrightarrow> bdd_above (X::real set)"
+  by (auto simp: bdd_above_def bdd_below_def intro: le_imp_neg_le) (metis le_imp_neg_le minus_minus)
 
 instantiation real :: linear_continuum
 begin
@@ -933,10 +939,10 @@
 
 instance
 proof
-  { fix z x :: real and X :: "real set"
-    assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> x \<le> z"
+  { fix x :: real and X :: "real set"
+    assume x: "x \<in> X" "bdd_above X"
     then obtain s where s: "\<forall>y\<in>X. y \<le> s" "\<And>z. \<forall>y\<in>X. y \<le> z \<Longrightarrow> s \<le> z"
-      using complete_real[of X] by blast
+      using complete_real[of X] unfolding bdd_above_def by blast
     then show "x \<le> Sup X"
       unfolding Sup_real_def by (rule LeastI2_order) (auto simp: x) }
   note Sup_upper = this
@@ -953,9 +959,9 @@
   note Sup_least = this
 
   { fix x z :: real and X :: "real set"
-    assume x: "x \<in> X" and z: "\<And>x. x \<in> X \<Longrightarrow> z \<le> x"
+    assume x: "x \<in> X" and [simp]: "bdd_below X"
     have "-x \<le> Sup (uminus ` X)"
-      by (rule Sup_upper[of _ _ "- z"]) (auto simp add: image_iff x z)
+      by (rule Sup_upper) (auto simp add: image_iff x)
     then show "Inf X \<le> x" 
       by (auto simp add: Inf_real_def) }
 
@@ -976,7 +982,7 @@
 *}
 
 lemma reals_complete: "\<exists>X. X \<in> S \<Longrightarrow> \<exists>Y. isUb (UNIV::real set) S Y \<Longrightarrow> \<exists>t. isLub (UNIV :: real set) S t"
-  by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro: cSup_upper)
+  by (intro exI[of _ "Sup S"] isLub_cSup) (auto simp: setle_def isUb_def intro!: cSup_upper)
 
 
 subsection {* Hiding implementation details *}
--- a/src/HOL/Topological_Spaces.thy	Tue Nov 05 09:44:57 2013 +0100
+++ b/src/HOL/Topological_Spaces.thy	Tue Nov 05 09:44:58 2013 +0100
@@ -2112,7 +2112,7 @@
   with dense[of b "Inf A"] obtain c where "c < Inf A" "c \<in> A"
     by (auto simp: subset_eq)
   then show False
-    using cInf_lower[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
+    using cInf_lower[OF `c \<in> A`] bnd by (metis not_le less_imp_le bdd_belowI)
 qed
 
 lemma Sup_notin_open:
@@ -2125,7 +2125,7 @@
   with dense[of "Sup A" b] obtain c where "Sup A < c" "c \<in> A"
     by (auto simp: subset_eq)
   then show False
-    using cSup_upper[OF `c \<in> A`, of x] bnd by (metis less_imp_le not_le)
+    using cSup_upper[OF `c \<in> A`] bnd by (metis less_imp_le not_le bdd_aboveI)
 qed
 
 end
@@ -2151,7 +2151,7 @@
     let ?z = "Inf (B \<inter> {x <..})"
 
     have "x \<le> ?z" "?z \<le> y"
-      using `y \<in> B` `x < y` by (auto intro: cInf_lower[where z=x] cInf_greatest)
+      using `y \<in> B` `x < y` by (auto intro: cInf_lower cInf_greatest)
     with `x \<in> U` `y \<in> U` have "?z \<in> U"
       by (rule *)
     moreover have "?z \<notin> B \<inter> {x <..}"
@@ -2163,11 +2163,11 @@
       obtain a where "?z < a" "{?z ..< a} \<subseteq> A"
         using open_right[OF `open A` `?z \<in> A` `?z < y`] by auto
       moreover obtain b where "b \<in> B" "x < b" "b < min a y"
-        using cInf_less_iff[of "B \<inter> {x <..}" x "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
+        using cInf_less_iff[of "B \<inter> {x <..}" "min a y"] `?z < a` `?z < y` `x < y` `y \<in> B`
         by (auto intro: less_imp_le)
       moreover have "?z \<le> b"
         using `b \<in> B` `x < b`
-        by (intro cInf_lower[where z=x]) auto
+        by (intro cInf_lower) auto
       moreover have "b \<in> U"
         using `x \<le> ?z` `?z \<le> b` `b < min a y`
         by (intro *[OF `x \<in> U` `y \<in> U`]) (auto simp: less_imp_le)