add SUP and INF for conditionally complete lattices
authorhoelzl
Tue, 05 Nov 2013 09:44:59 +0100
changeset 5571171c701dc5bf9
parent 55710 adfc759263ab
child 55712 6a967667fd45
add SUP and INF for conditionally complete lattices
src/HOL/Complete_Lattices.thy
src/HOL/Conditionally_Complete_Lattices.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
     1.1 --- a/src/HOL/Complete_Lattices.thy	Tue Nov 05 09:44:58 2013 +0100
     1.2 +++ b/src/HOL/Complete_Lattices.thy	Tue Nov 05 09:44:59 2013 +0100
     1.3 @@ -20,6 +20,12 @@
     1.4  definition INFI :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
     1.5    INF_def: "INFI A f = \<Sqinter>(f ` A)"
     1.6  
     1.7 +lemma INF_image [simp]: "INFI (f`A) g = INFI A (\<lambda>x. g (f x))"
     1.8 +  by (simp add: INF_def image_image)
     1.9 +
    1.10 +lemma INF_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> INFI A C = INFI B D"
    1.11 +  by (simp add: INF_def image_def)
    1.12 +
    1.13  end
    1.14  
    1.15  class Sup =
    1.16 @@ -29,6 +35,12 @@
    1.17  definition SUPR :: "'b set \<Rightarrow> ('b \<Rightarrow> 'a) \<Rightarrow> 'a" where
    1.18    SUP_def: "SUPR A f = \<Squnion>(f ` A)"
    1.19  
    1.20 +lemma SUP_image [simp]: "SUPR (f`A) g = SUPR A (%x. g (f x))"
    1.21 +  by (simp add: SUP_def image_image)
    1.22 +
    1.23 +lemma SUP_cong: "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> SUPR A C = SUPR B D"
    1.24 +  by (simp add: SUP_def image_def)
    1.25 +
    1.26  end
    1.27  
    1.28  text {*
    1.29 @@ -183,12 +195,6 @@
    1.30    "\<Squnion>UNIV = \<top>"
    1.31    by (auto intro!: antisym Sup_upper)
    1.32  
    1.33 -lemma INF_image [simp]: "(\<Sqinter>x\<in>f`A. g x) = (\<Sqinter>x\<in>A. g (f x))"
    1.34 -  by (simp add: INF_def image_image)
    1.35 -
    1.36 -lemma SUP_image [simp]: "(\<Squnion>x\<in>f`A. g x) = (\<Squnion>x\<in>A. g (f x))"
    1.37 -  by (simp add: SUP_def image_image)
    1.38 -
    1.39  lemma Inf_Sup: "\<Sqinter>A = \<Squnion>{b. \<forall>a \<in> A. b \<sqsubseteq> a}"
    1.40    by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least)
    1.41  
    1.42 @@ -201,14 +207,6 @@
    1.43  lemma Sup_subset_mono: "A \<subseteq> B \<Longrightarrow> \<Squnion>A \<sqsubseteq> \<Squnion>B"
    1.44    by (auto intro: Sup_least Sup_upper)
    1.45  
    1.46 -lemma INF_cong:
    1.47 -  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Sqinter>x\<in>A. C x) = (\<Sqinter>x\<in>B. D x)"
    1.48 -  by (simp add: INF_def image_def)
    1.49 -
    1.50 -lemma SUP_cong:
    1.51 -  "A = B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> C x = D x) \<Longrightarrow> (\<Squnion>x\<in>A. C x) = (\<Squnion>x\<in>B. D x)"
    1.52 -  by (simp add: SUP_def image_def)
    1.53 -
    1.54  lemma Inf_mono:
    1.55    assumes "\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<sqsubseteq> b"
    1.56    shows "\<Sqinter>A \<sqsubseteq> \<Sqinter>B"
     2.1 --- a/src/HOL/Conditionally_Complete_Lattices.thy	Tue Nov 05 09:44:58 2013 +0100
     2.2 +++ b/src/HOL/Conditionally_Complete_Lattices.thy	Tue Nov 05 09:44:59 2013 +0100
     2.3 @@ -10,10 +10,10 @@
     2.4  imports Main Lubs
     2.5  begin
     2.6  
     2.7 -lemma Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
     2.8 +lemma (in linorder) Sup_fin_eq_Max: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup_fin X = Max X"
     2.9    by (induct X rule: finite_ne_induct) (simp_all add: sup_max)
    2.10  
    2.11 -lemma Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
    2.12 +lemma (in linorder) Inf_fin_eq_Min: "finite X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf_fin X = Min X"
    2.13    by (induct X rule: finite_ne_induct) (simp_all add: inf_min)
    2.14  
    2.15  context preorder
    2.16 @@ -125,6 +125,12 @@
    2.17    thus "bdd_below (A \<union> B)" unfolding bdd_below_def ..
    2.18  qed
    2.19  
    2.20 +lemma bdd_above_sup[simp]: "bdd_above ((\<lambda>x. sup (f x) (g x)) ` A) \<longleftrightarrow> bdd_above (f`A) \<and> bdd_above (g`A)"
    2.21 +  by (auto simp: bdd_above_def intro: le_supI1 le_supI2)
    2.22 +
    2.23 +lemma bdd_below_inf[simp]: "bdd_below ((\<lambda>x. inf (f x) (g x)) ` A) \<longleftrightarrow> bdd_below (f`A) \<and> bdd_below (g`A)"
    2.24 +  by (auto simp: bdd_below_def intro: le_infI1 le_infI2)
    2.25 +
    2.26  end
    2.27  
    2.28  
    2.29 @@ -142,6 +148,24 @@
    2.30      and cSup_least: "X \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X \<le> z"
    2.31  begin
    2.32  
    2.33 +lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
    2.34 +  by (metis cSup_upper order_trans)
    2.35 +
    2.36 +lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
    2.37 +  by (metis cInf_lower order_trans)
    2.38 +
    2.39 +lemma cSup_mono: "B \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. b \<le> a) \<Longrightarrow> Sup B \<le> Sup A"
    2.40 +  by (metis cSup_least cSup_upper2)
    2.41 +
    2.42 +lemma cInf_mono: "B \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> \<exists>a\<in>A. a \<le> b) \<Longrightarrow> Inf A \<le> Inf B"
    2.43 +  by (metis cInf_greatest cInf_lower2)
    2.44 +
    2.45 +lemma cSup_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Sup A \<le> Sup B"
    2.46 +  by (metis cSup_least cSup_upper subsetD)
    2.47 +
    2.48 +lemma cInf_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> A \<subseteq> B \<Longrightarrow> Inf B \<le> Inf A"
    2.49 +  by (metis cInf_greatest cInf_lower subsetD)
    2.50 +
    2.51  lemma cSup_eq_maximum: "z \<in> X \<Longrightarrow> (\<And>x. x \<in> X \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup X = z"
    2.52    by (intro antisym cSup_upper[of z X] cSup_least[of X z]) auto
    2.53  
    2.54 @@ -154,18 +178,6 @@
    2.55  lemma le_cInf_iff: "S \<noteq> {} \<Longrightarrow> bdd_below S \<Longrightarrow> a \<le> Inf S \<longleftrightarrow> (\<forall>x\<in>S. a \<le> x)"
    2.56    by (metis order_trans cInf_lower cInf_greatest)
    2.57  
    2.58 -lemma cSup_singleton [simp]: "Sup {x} = x"
    2.59 -  by (intro cSup_eq_maximum) auto
    2.60 -
    2.61 -lemma cInf_singleton [simp]: "Inf {x} = x"
    2.62 -  by (intro cInf_eq_minimum) auto
    2.63 -
    2.64 -lemma cSup_upper2: "x \<in> X \<Longrightarrow> y \<le> x \<Longrightarrow> bdd_above X \<Longrightarrow> y \<le> Sup X"
    2.65 -  by (metis cSup_upper order_trans)
    2.66 - 
    2.67 -lemma cInf_lower2: "x \<in> X \<Longrightarrow> x \<le> y \<Longrightarrow> bdd_below X \<Longrightarrow> Inf X \<le> y"
    2.68 -  by (metis cInf_lower order_trans)
    2.69 -
    2.70  lemma cSup_eq_non_empty:
    2.71    assumes 1: "X \<noteq> {}"
    2.72    assumes 2: "\<And>x. x \<in> X \<Longrightarrow> x \<le> a"
    2.73 @@ -192,10 +204,16 @@
    2.74  lemma cInf_insert: "X \<noteq> {} \<Longrightarrow> bdd_below X \<Longrightarrow> Inf (insert a X) = inf a (Inf X)"
    2.75    by (intro cInf_eq_non_empty) (auto intro: le_infI2 cInf_lower cInf_greatest)
    2.76  
    2.77 +lemma cSup_singleton [simp]: "Sup {x} = x"
    2.78 +  by (intro cSup_eq_maximum) auto
    2.79 +
    2.80 +lemma cInf_singleton [simp]: "Inf {x} = x"
    2.81 +  by (intro cInf_eq_minimum) auto
    2.82 +
    2.83  lemma cSup_insert_If:  "bdd_above X \<Longrightarrow> Sup (insert a X) = (if X = {} then a else sup a (Sup X))"
    2.84    using cSup_insert[of X] by simp
    2.85  
    2.86 -lemma cInf_insert_if: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
    2.87 +lemma cInf_insert_If: "bdd_below X \<Longrightarrow> Inf (insert a X) = (if X = {} then a else inf a (Inf X))"
    2.88    using cInf_insert[of X] by simp
    2.89  
    2.90  lemma le_cSup_finite: "finite X \<Longrightarrow> x \<in> X \<Longrightarrow> x \<le> Sup X"
    2.91 @@ -234,6 +252,74 @@
    2.92  lemma cInf_atLeastAtMost[simp]: "y \<le> x \<Longrightarrow> Inf {y..x} = y"
    2.93    by (auto intro!: cInf_eq_minimum)
    2.94  
    2.95 +lemma cINF_lower: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> INFI A f \<le> f x"
    2.96 +  unfolding INF_def by (rule cInf_lower) auto
    2.97 +
    2.98 +lemma cINF_greatest: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> m \<le> f x) \<Longrightarrow> m \<le> INFI A f"
    2.99 +  unfolding INF_def by (rule cInf_greatest) auto
   2.100 +
   2.101 +lemma cSUP_upper: "x \<in> A \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> f x \<le> SUPR A f"
   2.102 +  unfolding SUP_def by (rule cSup_upper) auto
   2.103 +
   2.104 +lemma cSUP_least: "A \<noteq> {} \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> f x \<le> M) \<Longrightarrow> SUPR A f \<le> M"
   2.105 +  unfolding SUP_def by (rule cSup_least) auto
   2.106 +
   2.107 +lemma cINF_lower2: "bdd_below (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> f x \<le> u \<Longrightarrow> INFI A f \<le> u"
   2.108 +  by (auto intro: cINF_lower assms order_trans)
   2.109 +
   2.110 +lemma cSUP_upper2: "bdd_above (f ` A) \<Longrightarrow> x \<in> A \<Longrightarrow> u \<le> f x \<Longrightarrow> u \<le> SUPR A f"
   2.111 +  by (auto intro: cSUP_upper assms order_trans)
   2.112 +
   2.113 +lemma le_cINF_iff: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> u \<le> INFI A f \<longleftrightarrow> (\<forall>x\<in>A. u \<le> f x)"
   2.114 +  by (metis cINF_greatest cINF_lower assms order_trans)
   2.115 +
   2.116 +lemma cSUP_le_iff: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR A f \<le> u \<longleftrightarrow> (\<forall>x\<in>A. f x \<le> u)"
   2.117 +  by (metis cSUP_least cSUP_upper assms order_trans)
   2.118 +
   2.119 +lemma cINF_insert: "A \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> INFI (insert a A) f = inf (f a) (INFI A f)"
   2.120 +  by (metis INF_def cInf_insert assms empty_is_image image_insert)
   2.121 +
   2.122 +lemma cSUP_insert: "A \<noteq> {} \<Longrightarrow> bdd_above (f ` A) \<Longrightarrow> SUPR (insert a A) f = sup (f a) (SUPR A f)"
   2.123 +  by (metis SUP_def cSup_insert assms empty_is_image image_insert)
   2.124 +
   2.125 +lemma cINF_mono: "B \<noteq> {} \<Longrightarrow> bdd_below (f ` A) \<Longrightarrow> (\<And>m. m \<in> B \<Longrightarrow> \<exists>n\<in>A. f n \<le> g m) \<Longrightarrow> INFI A f \<le> INFI B g"
   2.126 +  unfolding INF_def by (auto intro: cInf_mono)
   2.127 +
   2.128 +lemma cSUP_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> (\<And>n. n \<in> A \<Longrightarrow> \<exists>m\<in>B. f n \<le> g m) \<Longrightarrow> SUPR A f \<le> SUPR B g"
   2.129 +  unfolding SUP_def by (auto intro: cSup_mono)
   2.130 +
   2.131 +lemma cINF_superset_mono: "A \<noteq> {} \<Longrightarrow> bdd_below (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> g x \<le> f x) \<Longrightarrow> INFI B g \<le> INFI A f"
   2.132 +  by (rule cINF_mono) auto
   2.133 +
   2.134 +lemma cSUP_subset_mono: "A \<noteq> {} \<Longrightarrow> bdd_above (g ` B) \<Longrightarrow> A \<subseteq> B \<Longrightarrow> (\<And>x. x \<in> B \<Longrightarrow> f x \<le> g x) \<Longrightarrow> SUPR A f \<le> SUPR B g"
   2.135 +  by (rule cSUP_mono) auto
   2.136 +
   2.137 +lemma less_eq_cInf_inter: "bdd_below A \<Longrightarrow> bdd_below B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> inf (Inf A) (Inf B) \<le> Inf (A \<inter> B)"
   2.138 +  by (metis cInf_superset_mono lattice_class.inf_sup_ord(1) le_infI1)
   2.139 +
   2.140 +lemma cSup_inter_less_eq: "bdd_above A \<Longrightarrow> bdd_above B \<Longrightarrow> A \<inter> B \<noteq> {} \<Longrightarrow> Sup (A \<inter> B) \<le> sup (Sup A) (Sup B) "
   2.141 +  by (metis cSup_subset_mono lattice_class.inf_sup_ord(1) le_supI1)
   2.142 +
   2.143 +lemma cInf_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below B \<Longrightarrow> Inf (A \<union> B) = inf (Inf A) (Inf B)"
   2.144 +  by (intro antisym le_infI cInf_greatest cInf_lower) (auto intro: le_infI1 le_infI2 cInf_lower)
   2.145 +
   2.146 +lemma cINF_union: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_below (f`B) \<Longrightarrow> INFI (A \<union> B) f = inf (INFI A f) (INFI B f)"
   2.147 +  unfolding INF_def using assms by (auto simp add: image_Un intro: cInf_union_distrib)
   2.148 +
   2.149 +lemma cSup_union_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above A \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above B \<Longrightarrow> Sup (A \<union> B) = sup (Sup A) (Sup B)"
   2.150 +  by (intro antisym le_supI cSup_least cSup_upper) (auto intro: le_supI1 le_supI2 cSup_upper)
   2.151 +
   2.152 +lemma cSUP_union: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> B \<noteq> {} \<Longrightarrow> bdd_above (f`B) \<Longrightarrow> SUPR (A \<union> B) f = sup (SUPR A f) (SUPR B f)"
   2.153 +  unfolding SUP_def by (auto simp add: image_Un intro: cSup_union_distrib)
   2.154 +
   2.155 +lemma cINF_inf_distrib: "A \<noteq> {} \<Longrightarrow> bdd_below (f`A) \<Longrightarrow> bdd_below (g`A) \<Longrightarrow> inf (INFI A f) (INFI A g) = (INF a:A. inf (f a) (g a))"
   2.156 +  by (intro antisym le_infI cINF_greatest cINF_lower2)
   2.157 +     (auto intro: le_infI1 le_infI2 cINF_greatest cINF_lower le_infI)
   2.158 +
   2.159 +lemma SUP_sup_distrib: "A \<noteq> {} \<Longrightarrow> bdd_above (f`A) \<Longrightarrow> bdd_above (g`A) \<Longrightarrow> sup (SUPR A f) (SUPR A g) = (SUP a:A. sup (f a) (g a))"
   2.160 +  by (intro antisym le_supI cSUP_least cSUP_upper2)
   2.161 +     (auto intro: le_supI1 le_supI2 cSUP_least cSUP_upper le_supI)
   2.162 +
   2.163  end
   2.164  
   2.165  instance complete_lattice \<subseteq> conditionally_complete_lattice
   2.166 @@ -323,14 +409,11 @@
   2.167  
   2.168  end
   2.169  
   2.170 -class linear_continuum = conditionally_complete_linorder + dense_linorder +
   2.171 -  assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
   2.172 -begin
   2.173 +lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
   2.174 +  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
   2.175  
   2.176 -lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
   2.177 -  by (metis UNIV_not_singleton neq_iff)
   2.178 -
   2.179 -end
   2.180 +lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
   2.181 +  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
   2.182  
   2.183  lemma cSup_bounds:
   2.184    fixes S :: "'a :: conditionally_complete_lattice set"
   2.185 @@ -347,19 +430,12 @@
   2.186    with b show ?thesis by blast
   2.187  qed
   2.188  
   2.189 -
   2.190  lemma cSup_unique: "(S::'a :: {conditionally_complete_linorder, no_bot} set) *<= b \<Longrightarrow> (\<forall>b'<b. \<exists>x\<in>S. b' < x) \<Longrightarrow> Sup S = b"
   2.191    by (rule cSup_eq) (auto simp: not_le[symmetric] setle_def)
   2.192  
   2.193  lemma cInf_unique: "b <=* (S::'a :: {conditionally_complete_linorder, no_top} set) \<Longrightarrow> (\<forall>b'>b. \<exists>x\<in>S. b' > x) \<Longrightarrow> Inf S = b"
   2.194    by (rule cInf_eq) (auto simp: not_le[symmetric] setge_def)
   2.195  
   2.196 -lemma cSup_eq_Max: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Sup X = Max X"
   2.197 -  using cSup_eq_Sup_fin[of X] Sup_fin_eq_Max[of X] by simp
   2.198 -
   2.199 -lemma cInf_eq_Min: "finite (X::'a::conditionally_complete_linorder set) \<Longrightarrow> X \<noteq> {} \<Longrightarrow> Inf X = Min X"
   2.200 -  using cInf_eq_Inf_fin[of X] Inf_fin_eq_Min[of X] by simp
   2.201 -
   2.202  lemma cSup_lessThan[simp]: "Sup {..<x::'a::{conditionally_complete_linorder, no_bot, dense_linorder}} = x"
   2.203    by (auto intro!: cSup_eq_non_empty intro: dense_le)
   2.204  
   2.205 @@ -378,4 +454,13 @@
   2.206  lemma cInf_greaterThanLessThan[simp]: "y < x \<Longrightarrow> Inf {y<..<x::'a::{conditionally_complete_linorder, no_top, dense_linorder}} = y"
   2.207    by (auto intro!: cInf_eq intro: dense_ge_bounded)
   2.208  
   2.209 +class linear_continuum = conditionally_complete_linorder + dense_linorder +
   2.210 +  assumes UNIV_not_singleton: "\<exists>a b::'a. a \<noteq> b"
   2.211 +begin
   2.212 +
   2.213 +lemma ex_gt_or_lt: "\<exists>b. a < b \<or> b < a"
   2.214 +  by (metis UNIV_not_singleton neq_iff)
   2.215 +
   2.216  end
   2.217 +
   2.218 +end
     3.1 --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Tue Nov 05 09:44:58 2013 +0100
     3.2 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Tue Nov 05 09:44:59 2013 +0100
     3.3 @@ -2686,7 +2686,7 @@
     3.4  lemma Inf_insert:
     3.5    fixes S :: "real set"
     3.6    shows "bounded S \<Longrightarrow> Inf (insert x S) = (if S = {} then x else min x (Inf S))"
     3.7 -  by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_if)
     3.8 +  by (auto simp: bounded_imp_bdd_below inf_min cInf_insert_If)
     3.9  
    3.10  lemma Inf_insert_finite:
    3.11    fixes S :: "real set"