# HG changeset patch # User hoelzl # Date 1383641098 -3600 # Node ID adfc759263ab69120c69e90e59f842c49fc31c61 # Parent 5c7a3b6b05a9cf4e4f14748f6be80bf2d52e4876 use bdd_above and bdd_below for conditionally complete lattices diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Conditionally_Complete_Lattices.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 \ X \ {} \ Inf_fin X = Min X" by (induct X rule: finite_ne_induct) (simp_all add: inf_min) +context preorder +begin + +definition "bdd_above A \ (\M. \x \ A. x \ M)" +definition "bdd_below A \ (\m. \x \ A. m \ x)" + +lemma bdd_aboveI[intro]: "(\x. x \ A \ x \ M) \ bdd_above A" + by (auto simp: bdd_above_def) + +lemma bdd_belowI[intro]: "(\x. x \ A \ m \ x) \ 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 \ A \ B \ bdd_above A" + by (metis (full_types) bdd_above_def order_class.le_neq_trans psubsetD) + +lemma bdd_below_mono: "bdd_below B \ A \ B \ bdd_below A" + by (metis bdd_below_def order_class.le_neq_trans psubsetD) + +lemma bdd_above_Int1 [simp]: "bdd_above A \ bdd_above (A \ B)" + using bdd_above_mono by auto + +lemma bdd_above_Int2 [simp]: "bdd_above B \ bdd_above (A \ B)" + using bdd_above_mono by auto + +lemma bdd_below_Int1 [simp]: "bdd_below A \ bdd_below (A \ B)" + using bdd_below_mono by auto + +lemma bdd_below_Int2 [simp]: "bdd_below B \ bdd_below (A \ 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 \ B) = (bdd_above A \ bdd_above B)" +proof + assume "bdd_above (A \ B)" + thus "bdd_above A \ bdd_above B" unfolding bdd_above_def by auto +next + assume "bdd_above A \ bdd_above B" + then obtain a b where "\x\A. x \ a" "\x\B. x \ b" unfolding bdd_above_def by auto + hence "\x \ A \ B. x \ sup a b" by (auto intro: Un_iff le_supI1 le_supI2) + thus "bdd_above (A \ B)" unfolding bdd_above_def .. +qed + +lemma bdd_below_Un [simp]: "bdd_below (A \ B) = (bdd_below A \ bdd_below B)" +proof + assume "bdd_below (A \ B)" + thus "bdd_below A \ bdd_below B" unfolding bdd_below_def by auto +next + assume "bdd_below A \ bdd_below B" + then obtain a b where "\x\A. a \ x" "\x\B. b \ x" unfolding bdd_below_def by auto + hence "\x \ A \ B. inf a b \ x" by (auto intro: Un_iff le_infI1 le_infI2) + thus "bdd_below (A \ 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 \ X \ (\a. a \ X \ z \ a) \ Inf X \ x" + assumes cInf_lower: "x \ X \ bdd_below X \ Inf X \ x" and cInf_greatest: "X \ {} \ (\x. x \ X \ z \ x) \ z \ Inf X" - assumes cSup_upper: "x \ X \ (\a. a \ X \ a \ z) \ x \ Sup X" + assumes cSup_upper: "x \ X \ bdd_above X \ x \ Sup X" and cSup_least: "X \ {} \ (\x. x \ X \ x \ z) \ Sup X \ z" begin -lemma cSup_eq_maximum: (*REAL_SUP_MAX in HOL4*) - "z \ X \ (\x. x \ X \ x \ z) \ Sup X = z" - by (blast intro: antisym cSup_upper cSup_least) +lemma cSup_eq_maximum: "z \ X \ (\x. x \ X \ x \ z) \ 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 \ X \ (\x. x \ X \ z \ x) \ Inf X = z" - by (intro antisym cInf_lower[of z X z] cInf_greatest[of X z]) auto +lemma cInf_eq_minimum: "z \ X \ (\x. x \ X \ z \ x) \ Inf X = z" + by (intro antisym cInf_lower[of z X] cInf_greatest[of X z]) auto -lemma cSup_le_iff: "S \ {} \ (\a. a \ S \ a \ z) \ Sup S \ a \ (\x\S. x \ a)" +lemma cSup_le_iff: "S \ {} \ bdd_above S \ Sup S \ a \ (\x\S. x \ a)" by (metis order_trans cSup_upper cSup_least) -lemma le_cInf_iff: "S \ {} \ (\a. a \ S \ z \ a) \ a \ Inf S \ (\x\S. a \ x)" +lemma le_cInf_iff: "S \ {} \ bdd_below S \ a \ Inf S \ (\x\S. a \ 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 \ X \ y \ x \ (\x. x \ X \ x \ z) \ y \ Sup X" +lemma cSup_upper2: "x \ X \ y \ x \ bdd_above X \ y \ Sup X" by (metis cSup_upper order_trans) -lemma cInf_lower2: - "x \ X \ x \ y \ (\x. x \ X \ z \ x) \ Inf X \ y" +lemma cInf_lower2: "x \ X \ x \ y \ bdd_below X \ Inf X \ y" by (metis cInf_lower order_trans) -lemma cSup_upper_EX: "x \ X \ \z. \x. x \ X \ x \ z \ x \ Sup X" - by (blast intro: cSup_upper) - -lemma cInf_lower_EX: "x \ X \ \z. \x. x \ X \ z \ x \ Inf X \ x" - by (blast intro: cInf_lower) - lemma cSup_eq_non_empty: assumes 1: "X \ {}" assumes 2: "\x. x \ X \ x \ 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 \ {} \ (\x. x \ S \ z \ x) \ Inf S = Sup {x. \s\S. x \ s}" - by (rule cInf_eq_non_empty) (auto intro: cSup_upper cSup_least) - -lemma cSup_cInf: "S \ {} \ (\x. x \ S \ x \ z) \ Sup S = Inf {x. \s\S. s \ x}" - by (rule cSup_eq_non_empty) (auto intro: cInf_lower cInf_greatest) +lemma cInf_cSup: "S \ {} \ bdd_below S \ Inf S = Sup {x. \s\S. x \ s}" + by (rule cInf_eq_non_empty) (auto intro!: cSup_upper cSup_least simp: bdd_below_def) -lemma cSup_insert: - assumes x: "X \ {}" - and z: "\x. x \ X \ x \ z" - shows "Sup (insert a X) = sup a (Sup X)" -proof (intro cSup_eq_non_empty) - fix y assume "\x. x \ insert a X \ x \ y" with x show "sup a (Sup X) \ y" by (auto intro: cSup_least) -qed (auto intro: le_supI2 z cSup_upper) +lemma cSup_cInf: "S \ {} \ bdd_above S \ Sup S = Inf {x. \s\S. s \ x}" + by (rule cSup_eq_non_empty) (auto intro!: cInf_lower cInf_greatest simp: bdd_above_def) + +lemma cSup_insert: "X \ {} \ bdd_above X \ 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 \ {}" - and z: "\x. x \ X \ z \ x" - shows "Inf (insert a X) = inf a (Inf X)" -proof (intro cInf_eq_non_empty) - fix y assume "\x. x \ insert a X \ y \ x" with x show "y \ inf a (Inf X)" by (auto intro: cInf_greatest) -qed (auto intro: le_infI2 z cInf_lower) +lemma cInf_insert: "X \ {} \ bdd_below X \ 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: - "(\x. x \ X \ x \ z) \ 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 \ Sup (insert a X) = (if X = {} then a else sup a (Sup X))" + using cSup_insert[of X] by simp -lemma cInf_insert_if: - "(\x. x \ X \ z \ x) \ 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 \ 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 \ x \ X \ x \ 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 \ x \ X \ Inf X \ 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 \ X \ {} \ 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 \ X \ {} \ 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) \ {} \ (\b. S *<= b) \ 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 \ {} \ (\x. x \ X \ x \ z) \ y < Sup X \ (\x\X. y < x)" + "X \ {} \ bdd_above X \ y < Sup X \ (\x\X. y < x)" by (rule iffI) (metis cSup_least not_less, metis cSup_upper less_le_trans) -lemma cInf_less_iff: "X \ {} \ (\x. x \ X \ z \ x) \ Inf X < y \ (\x\X. x < y)" +lemma cInf_less_iff: "X \ {} \ bdd_below X \ Inf X < y \ (\x\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 \ {} \ z < Sup X \ \x\X. z < x" - by (metis less_cSup_iff not_leE) + by (metis less_cSup_iff not_leE bdd_above_def) lemma cInf_lessD: "X \ {} \ Inf X < z \ \x\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 "\ P b" @@ -219,7 +296,7 @@ (\d. (\x. a \ x \ x < d \ P x) \ d \ c)" proof (rule exI [where x = "Sup {d. \x. a \ x & x < d --> P x}"], auto) show "a \ Sup {d. \c. a \ c \ c < d \ P c}" - by (rule cSup_upper [where z=b], auto) + by (rule cSup_upper, auto simp: bdd_above_def) (metis `a < b` `\ P b` linear less_le) next show "Sup {d. \c. a \ c \ c < d \ P c} \ b" @@ -240,7 +317,7 @@ fix d assume 0: "\x. a \ x \ x < d \ P x" thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" - by (rule_tac z="b" in cSup_upper, auto) + by (rule_tac cSup_upper, auto simp: bdd_above_def) (metis `az. finite z \ (\a. a \ X \ a \ z) \ 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 \ X" "(\a. a \ X \ z |\| a)" + assume "x \ X" "bdd_below X" then show "Inf X |\| x" by transfer auto next assume "X \ {}" "(\x. x \ X \ z |\| x)" then show "z |\| Inf X" by transfer (clarsimp, blast) next - assume "x \ X" "(\a. a \ X \ a |\| z)" - then show "x |\| Sup X" by transfer (auto intro!: finite_Sup) + assume "x \ X" "bdd_above X" + then obtain z where "x \ X" "(\x. x \ X \ x |\| z)" + by (auto simp: bdd_above_def) + then show "x |\| Sup X" + by transfer (auto intro!: finite_Sup) next assume "X \ {}" "(\x. x \ X \ x |\| z)" then show "Sup X |\| z" by transfer (clarsimp, blast) diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Multivariate_Analysis/Convex_Euclidean_Space.thy --- 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 \ interior I` by auto have "Inf (?F x) \ (f x - f y) / (x - y)" - proof (rule cInf_lower2) + proof (intro bdd_belowI cInf_lower2) show "(f x - f t) / (x - t) \ ?F x" using `t \ I` `x < t` by auto show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy --- 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 \ {}" { assume ex: "\B. \x\S. B \ x" then have *: "\x\S. Inf S \ 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 \ S" apply (subst closed_contains_Inf) using ex `S \ {}` `closed S` diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Multivariate_Analysis/Integration.thy --- 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 \ {} \ (\x\S. \x\ \ a) \ \Sup S\ \ 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 \ s" have *: "\x y::real. x \ - y \ - x \ y" by auto show "Inf {f j x |j. n \ j} \ 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 \ s" show "f n x \ Sup {f j x |j. n \ 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 diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- 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 \ {}" "\x\S. B \ x" + assumes "S \ {}" "bdd_below S" shows "Inf S \ closure S" proof - have *: "\x\S. Inf S \ 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 \ S" "x < Inf S + e" - by (subst (asm) cInf_less_iff[of _ B]) auto + by (subst (asm) cInf_less_iff) auto with * have "\x\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 \ {}" "\x\S. B \ x" - and "closed S" - shows "Inf S \ S" + shows "S \ {} \ bdd_below S \ closed S \ Inf S \ S" by (metis closure_contains_Inf closure_closed assms) - lemma not_trivial_limit_within_ball: "\ trivial_limit (at x within S) \ (\e>0. S \ ball x e - {x} \ {})" (is "?lhs = ?rhs") @@ -1977,27 +1974,20 @@ definition "infdist x A = (if A = {} then 0 else Inf {dist x a|a. a \ A})" +lemma bdd_below_infdist[intro, simp]: "bdd_below {dist x a|a. a \ A}" + by (auto intro!: zero_le_dist) + lemma infdist_notempty: "A \ {} \ infdist x A = Inf {dist x a|a. a \ A}" by (simp add: infdist_def) lemma infdist_nonneg: "0 \ infdist x A" by (auto simp add: infdist_def intro: cInf_greatest) -lemma infdist_le: - assumes "a \ A" - and "d = dist x a" - shows "infdist x A \ d" - using assms by (auto intro!: cInf_lower[where z=0] simp add: infdist_def) - -lemma infdist_zero[simp]: - assumes "a \ A" - shows "infdist a A = 0" -proof - - from infdist_le[OF assms, of "dist a a"] have "infdist a A \ 0" - by auto - with infdist_nonneg[of a A] assms show "infdist a A = 0" - by auto -qed +lemma infdist_le: "a \ A \ d = dist x a \ infdist x A \ d" + using assms by (auto intro: cInf_lower simp add: infdist_def) + +lemma infdist_zero[simp]: "a \ A \ infdist a A = 0" + by (auto intro!: antisym infdist_nonneg infdist_le) lemma infdist_triangle: "infdist x A \ infdist y A + dist x y" proof (cases "A = {}") @@ -2021,13 +2011,7 @@ using `a \ A` by auto show "dist x a \ d" unfolding d by (rule dist_triangle) - fix d - assume "d \ {dist x a |a. a \ A}" - then obtain a where "a \ A" "d = dist x a" - by auto - then show "infdist x A \ d" - by (rule infdist_le) - qed + qed simp qed also have "\ = 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 \ (\a. \x\S. abs x \ a)" +lemma bounded_real: "bounded (S::real set) \ (\a. \x\S. \x\ \ a)" by (simp add: bounded_iff) +lemma bounded_imp_bdd_above: "bounded S \ 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 \ 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 "\x\S. x \ Sup S" and "\b. (\x\S. x \ b) \ Sup S \ b" proof - fix x - assume "x\S" - then show "x \ Sup S" - by (metis cSup_upper abs_le_D1 assms(1) bounded_real) -next show "\b. (\x\S. x \ b) \ Sup S \ 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 \ 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 "\x\S. x \ Inf S" and "\b. (\x\S. x \ b) \ Inf S \ b" proof - fix x - assume "x \ S" - from assms(1) obtain a where a: "\x\S. \x\ \ a" - unfolding bounded_real by auto - then show "x \ Inf S" using `x \ S` - by (metis cInf_lower_EX abs_le_D2 minus_le_iff) -next show "\b. (\x\S. x \ b) \ Inf S \ 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 \ 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: "\x. x \ s \ dist z x \ d" unfolding bounded_def by auto have "dist x y \ Sup ?D" - proof (rule cSup_upper, safe) - fix a b - assume "a \ s" "b \ s" - with z[of a] z[of b] dist_triangle[of a b z] - show "dist a b \ 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 \ s" "b \ s" + with z[of a] z[of b] dist_triangle[of a b z] + show "dist a b \ 2 * d" + by (simp add: dist_commute) + qed qed (insert s, auto) with `x \ s` show ?thesis by (auto simp add: diameter_def) diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Real.thy --- 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) \ 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) \ 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 \ X" and z: "\x. x \ X \ x \ z" + { fix x :: real and X :: "real set" + assume x: "x \ X" "bdd_above X" then obtain s where s: "\y\X. y \ s" "\z. \y\X. y \ z \ s \ z" - using complete_real[of X] by blast + using complete_real[of X] unfolding bdd_above_def by blast then show "x \ 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 \ X" and z: "\x. x \ X \ z \ x" + assume x: "x \ X" and [simp]: "bdd_below X" have "-x \ 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 \ x" by (auto simp add: Inf_real_def) } @@ -976,7 +982,7 @@ *} lemma reals_complete: "\X. X \ S \ \Y. isUb (UNIV::real set) S Y \ \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 *} diff -r 5c7a3b6b05a9 -r adfc759263ab src/HOL/Topological_Spaces.thy --- 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 \ A" by (auto simp: subset_eq) then show False - using cInf_lower[OF `c \ A`, of x] bnd by (metis less_imp_le not_le) + using cInf_lower[OF `c \ 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 \ A" by (auto simp: subset_eq) then show False - using cSup_upper[OF `c \ A`, of x] bnd by (metis less_imp_le not_le) + using cSup_upper[OF `c \ A`] bnd by (metis less_imp_le not_le bdd_aboveI) qed end @@ -2151,7 +2151,7 @@ let ?z = "Inf (B \ {x <..})" have "x \ ?z" "?z \ y" - using `y \ B` `x < y` by (auto intro: cInf_lower[where z=x] cInf_greatest) + using `y \ B` `x < y` by (auto intro: cInf_lower cInf_greatest) with `x \ U` `y \ U` have "?z \ U" by (rule *) moreover have "?z \ B \ {x <..}" @@ -2163,11 +2163,11 @@ obtain a where "?z < a" "{?z ..< a} \ A" using open_right[OF `open A` `?z \ A` `?z < y`] by auto moreover obtain b where "b \ B" "x < b" "b < min a y" - using cInf_less_iff[of "B \ {x <..}" x "min a y"] `?z < a` `?z < y` `x < y` `y \ B` + using cInf_less_iff[of "B \ {x <..}" "min a y"] `?z < a` `?z < y` `x < y` `y \ B` by (auto intro: less_imp_le) moreover have "?z \ b" using `b \ B` `x < b` - by (intro cInf_lower[where z=x]) auto + by (intro cInf_lower) auto moreover have "b \ U" using `x \ ?z` `?z \ b` `b < min a y` by (intro *[OF `x \ U` `y \ U`]) (auto simp: less_imp_le)