diff -r 1fad3160d873 -r 3b7e2dbbd684 src/HOL/SupInf.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SupInf.thy Tue Oct 27 12:59:57 2009 +0000 @@ -0,0 +1,467 @@ +(* Author: Amine Chaieb and L C Paulson, University of Cambridge *) + +header {*Sup and Inf Operators on Sets of Reals.*} + +theory SupInf +imports RComplete +begin + +lemma minus_max_eq_min: + fixes x :: "'a::{lordered_ab_group_add, linorder}" + shows "- (max x y) = min (-x) (-y)" +by (metis le_imp_neg_le linorder_linear min_max.inf_absorb2 min_max.le_iff_inf min_max.le_iff_sup min_max.sup_absorb1) + +lemma minus_min_eq_max: + fixes x :: "'a::{lordered_ab_group_add, linorder}" + shows "- (min x y) = max (-x) (-y)" +by (metis minus_max_eq_min minus_minus) + +lemma minus_Max_eq_Min [simp]: + fixes S :: "'a::{lordered_ab_group_add, linorder} set" + shows "finite S \ S \ {} \ - (Max S) = Min (uminus ` S)" +proof (induct S rule: finite_ne_induct) + case (singleton x) + thus ?case by simp +next + case (insert x S) + thus ?case by (simp add: minus_max_eq_min) +qed + +lemma minus_Min_eq_Max [simp]: + fixes S :: "'a::{lordered_ab_group_add, linorder} set" + shows "finite S \ S \ {} \ - (Min S) = Max (uminus ` S)" +proof (induct S rule: finite_ne_induct) + case (singleton x) + thus ?case by simp +next + case (insert x S) + thus ?case by (simp add: minus_min_eq_max) +qed + +instantiation real :: Sup +begin +definition + Sup_real_def [code del]: "Sup X == (LEAST z::real. \x\X. x\z)" + +instance .. +end + +instantiation real :: Inf +begin +definition + Inf_real_def [code del]: "Inf (X::real set) == - (Sup (uminus ` X))" + +instance .. +end + +subsection{*Supremum of a set of reals*} + +lemma Sup_upper [intro]: (*REAL_SUP_UBOUND in HOL4*) + fixes x :: real + assumes x: "x \ X" + and z: "!!x. x \ X \ x \ z" + shows "x \ Sup X" +proof (auto simp add: Sup_real_def) + from reals_complete2 + obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" + by (blast intro: x z) + hence "x \ s" + by (blast intro: x z) + also with s have "... = (LEAST z. \x\X. x \ z)" + by (fast intro: Least_equality [symmetric]) + finally show "x \ (LEAST z. \x\X. x \ z)" . +qed + +lemma Sup_least [intro]: (*REAL_IMP_SUP_LE in HOL4*) + fixes z :: real + assumes x: "X \ {}" + and z: "\x. x \ X \ x \ z" + shows "Sup X \ z" +proof (auto simp add: Sup_real_def) + from reals_complete2 x + obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" + by (blast intro: z) + hence "(LEAST z. \x\X. x \ z) = s" + by (best intro: Least_equality) + also with s z have "... \ z" + by blast + finally show "(LEAST z. \x\X. x \ z) \ z" . +qed + +lemma Sup_singleton [simp]: "Sup {x::real} = x" + by (force intro: Least_equality simp add: Sup_real_def) + +lemma Sup_eq_maximum: (*REAL_SUP_MAX in HOL4*) + fixes z :: real + assumes X: "z \ X" and z: "!!x. x \ X \ x \ z" + shows "Sup X = z" + by (force intro: Least_equality X z simp add: Sup_real_def) + +lemma Sup_upper2: (*REAL_IMP_LE_SUP in HOL4*) + fixes x :: real + shows "x \ X \ y \ x \ (!!x. x \ X \ x \ z) \ y \ Sup X" + by (metis Sup_upper real_le_trans) + +lemma Sup_real_iff : (*REAL_SUP_LE in HOL4*) + fixes z :: real + shows "X ~= {} ==> (!!x. x \ X ==> x \ z) ==> (\x\X. y y < Sup X" + by (metis Sup_least Sup_upper linorder_not_le le_less_trans) + +lemma Sup_eq: + fixes a :: real + shows "(!!x. x \ X \ x \ a) + \ (!!y. (!!x. x \ X \ x \ y) \ a \ y) \ Sup X = a" + by (metis Sup_least Sup_upper add_le_cancel_left diff_add_cancel insert_absorb + insert_not_empty real_le_anti_sym) + +lemma Sup_le: + fixes S :: "real set" + shows "S \ {} \ S *<= b \ Sup S \ b" +by (metis SupInf.Sup_least setle_def) + +lemma Sup_upper_EX: + fixes x :: real + shows "x \ X \ \z. \x. x \ X \ x \ z \ x \ Sup X" + by blast + +lemma Sup_insert_nonempty: + fixes x :: real + assumes x: "x \ X" + and z: "!!x. x \ X \ x \ z" + shows "Sup (insert a X) = max a (Sup X)" +proof (cases "Sup X \ a") + case True + thus ?thesis + apply (simp add: max_def) + apply (rule Sup_eq_maximum) + apply (metis insertCI) + apply (metis Sup_upper insertE le_iff_sup real_le_linear real_le_trans sup_absorb1 z) + done +next + case False + hence 1:"a < Sup X" by simp + have "Sup X \ Sup (insert a X)" + apply (rule Sup_least) + apply (metis empty_psubset_nonempty psubset_eq x) + apply (rule Sup_upper_EX) + apply blast + apply (metis insert_iff real_le_linear real_le_refl real_le_trans z) + done + moreover + have "Sup (insert a X) \ Sup X" + apply (rule Sup_least) + apply blast + apply (metis False Sup_upper insertE real_le_linear z) + done + ultimately have "Sup (insert a X) = Sup X" + by (blast intro: antisym ) + thus ?thesis + by (metis 1 min_max.le_iff_sup real_less_def) +qed + +lemma Sup_insert_if: + fixes X :: "real set" + assumes z: "!!x. x \ X \ x \ z" + shows "Sup (insert a X) = (if X={} then a else max a (Sup X))" +by auto (metis Sup_insert_nonempty z) + +lemma Sup: + fixes S :: "real set" + shows "S \ {} \ (\b. S *<= b) \ isLub UNIV S (Sup S)" +by (auto simp add: isLub_def setle_def leastP_def isUb_def intro!: setgeI) + +lemma Sup_finite_Max: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "Sup S = Max S" +using fS Se +proof- + let ?m = "Max S" + from Max_ge[OF fS] have Sm: "\ x\ S. x \ ?m" by metis + with Sup[OF Se] have lub: "isLub UNIV S (Sup S)" by (metis setle_def) + from Max_in[OF fS Se] lub have mrS: "?m \ Sup S" + by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) + moreover + have "Sup S \ ?m" using Sm lub + by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) + ultimately show ?thesis by arith +qed + +lemma Sup_finite_in: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "Sup S \ S" + using Sup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis + +lemma Sup_finite_ge_iff: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "a \ Sup S \ (\ x \ S. a \ x)" +by (metis Max_ge Se Sup_finite_Max Sup_finite_in fS linorder_not_le less_le_trans) + +lemma Sup_finite_le_iff: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "a \ Sup S \ (\ x \ S. a \ x)" +by (metis Max_ge Se Sup_finite_Max Sup_finite_in fS le_iff_sup real_le_trans) + +lemma Sup_finite_gt_iff: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "a < Sup S \ (\ x \ S. a < x)" +by (metis Se Sup_finite_le_iff fS linorder_not_less) + +lemma Sup_finite_lt_iff: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "a > Sup S \ (\ x \ S. a > x)" +by (metis Se Sup_finite_ge_iff fS linorder_not_less) + +lemma Sup_unique: + fixes S :: "real set" + shows "S *<= b \ (\b' < b. \x \ S. b' < x) \ Sup S = b" +unfolding setle_def +apply (rule Sup_eq, auto) +apply (metis linorder_not_less) +done + +lemma Sup_abs_le: + fixes S :: "real set" + shows "S \ {} \ (\x\S. \x\ \ a) \ \Sup S\ \ a" +by (auto simp add: abs_le_interval_iff) (metis Sup_upper2) + +lemma Sup_bounds: + fixes S :: "real set" + assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" + shows "a \ Sup S \ Sup S \ b" +proof- + from Sup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast + hence b: "Sup S \ b" using u + by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) + from Se obtain y where y: "y \ S" by blast + from lub l have "a \ Sup S" + by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) + (metis le_iff_sup le_sup_iff y) + with b show ?thesis by blast +qed + +lemma Sup_asclose: + fixes S :: "real set" + assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\Sup S - l\ \ e" +proof- + have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith + thus ?thesis using S b Sup_bounds[of S "l - e" "l+e"] unfolding th + by (auto simp add: setge_def setle_def) +qed + + +subsection{*Infimum of a set of reals*} + +lemma Inf_lower [intro]: + fixes z :: real + assumes x: "x \ X" + and z: "!!x. x \ X \ z \ x" + shows "Inf X \ x" +proof - + have "-x \ Sup (uminus ` X)" + by (rule Sup_upper [where z = "-z"]) (auto simp add: image_iff x z) + thus ?thesis + by (auto simp add: Inf_real_def) +qed + +lemma Inf_greatest [intro]: + fixes z :: real + assumes x: "X \ {}" + and z: "\x. x \ X \ z \ x" + shows "z \ Inf X" +proof - + have "Sup (uminus ` X) \ -z" using x z by (force intro: Sup_least) + hence "z \ - Sup (uminus ` X)" + by simp + thus ?thesis + by (auto simp add: Inf_real_def) +qed + +lemma Inf_singleton [simp]: "Inf {x::real} = x" + by (simp add: Inf_real_def) + +lemma Inf_eq_minimum: (*REAL_INF_MIN in HOL4*) + fixes z :: real + assumes x: "z \ X" and z: "!!x. x \ X \ z \ x" + shows "Inf X = z" +proof - + have "Sup (uminus ` X) = -z" using x z + by (force intro: Sup_eq_maximum x z) + thus ?thesis + by (simp add: Inf_real_def) +qed + +lemma Inf_lower2: + fixes x :: real + shows "x \ X \ x \ y \ (!!x. x \ X \ z \ x) \ Inf X \ y" + by (metis Inf_lower real_le_trans) + +lemma Inf_real_iff: + fixes z :: real + shows "X \ {} \ (!!x. x \ X \ z \ x) \ (\x\X. x Inf X < y" + by (metis Inf_greatest Inf_lower less_le_not_le real_le_linear + order_less_le_trans) + +lemma Inf_eq: + fixes a :: real + shows "(!!x. x \ X \ a \ x) \ (!!y. (!!x. x \ X \ y \ x) \ y \ a) \ Inf X = a" + by (metis Inf_greatest Inf_lower add_le_cancel_left diff_add_cancel + insert_absorb insert_not_empty real_le_anti_sym) + +lemma Inf_ge: + fixes S :: "real set" + shows "S \ {} \ b <=* S \ Inf S \ b" +by (metis SupInf.Inf_greatest setge_def) + +lemma Inf_lower_EX: + fixes x :: real + shows "x \ X \ \z. \x. x \ X \ z \ x \ Inf X \ x" + by blast + +lemma Inf_insert_nonempty: + fixes x :: real + assumes x: "x \ X" + and z: "!!x. x \ X \ z \ x" + shows "Inf (insert a X) = min a (Inf X)" +proof (cases "a \ Inf X") + case True + thus ?thesis + by (simp add: min_def) + (blast intro: Inf_eq_minimum Inf_lower real_le_refl real_le_trans z) +next + case False + hence 1:"Inf X < a" by simp + have "Inf (insert a X) \ Inf X" + apply (rule Inf_greatest) + apply (metis empty_psubset_nonempty psubset_eq x) + apply (rule Inf_lower_EX) + apply (blast intro: elim:) + apply (metis insert_iff real_le_linear real_le_refl real_le_trans z) + done + moreover + have "Inf X \ Inf (insert a X)" + apply (rule Inf_greatest) + apply blast + apply (metis False Inf_lower insertE real_le_linear z) + done + ultimately have "Inf (insert a X) = Inf X" + by (blast intro: antisym ) + thus ?thesis + by (metis False min_max.inf_absorb2 real_le_linear) +qed + +lemma Inf_insert_if: + fixes X :: "real set" + assumes z: "!!x. x \ X \ z \ x" + shows "Inf (insert a X) = (if X={} then a else min a (Inf X))" +by auto (metis Inf_insert_nonempty z) + +text{*We could prove the analogous result for the supremum, and also generalise it to the union operator.*} +lemma Inf_binary: + "Inf{a, b::real} = min a b" + by (subst Inf_insert_nonempty, auto) + +lemma Inf_greater: + fixes z :: real + shows "X \ {} \ Inf X < z \ \x \ X. x < z" + by (metis Inf_real_iff mem_def not_leE) + +lemma Inf_close: + fixes e :: real + shows "X \ {} \ 0 < e \ \x \ X. x < Inf X + e" + by (metis add_strict_increasing comm_monoid_add.mult_commute Inf_greater linorder_not_le pos_add_strict) + +lemma Inf_finite_Min: + fixes S :: "real set" + shows "finite S \ S \ {} \ Inf S = Min S" +by (simp add: Inf_real_def Sup_finite_Max image_image) + +lemma Inf_finite_in: + fixes S :: "real set" + assumes fS: "finite S" and Se: "S \ {}" + shows "Inf S \ S" + using Inf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis + +lemma Inf_finite_ge_iff: + fixes S :: "real set" + shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" +by (metis Inf_finite_Min Inf_finite_in Min_le real_le_trans) + +lemma Inf_finite_le_iff: + fixes S :: "real set" + shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" +by (metis Inf_finite_Min Inf_finite_ge_iff Inf_finite_in Min_le + real_le_anti_sym real_le_linear) + +lemma Inf_finite_gt_iff: + fixes S :: "real set" + shows "finite S \ S \ {} \ a < Inf S \ (\ x \ S. a < x)" +by (metis Inf_finite_le_iff linorder_not_less) + +lemma Inf_finite_lt_iff: + fixes S :: "real set" + shows "finite S \ S \ {} \ a > Inf S \ (\ x \ S. a > x)" +by (metis Inf_finite_ge_iff linorder_not_less) + +lemma Inf_unique: + fixes S :: "real set" + shows "b <=* S \ (\b' > b. \x \ S. b' > x) \ Inf S = b" +unfolding setge_def +apply (rule Inf_eq, auto) +apply (metis less_le_not_le linorder_not_less) +done + +lemma Inf_abs_ge: + fixes S :: "real set" + shows "S \ {} \ (\x\S. \x\ \ a) \ \Inf S\ \ a" +by (simp add: Inf_real_def) (rule Sup_abs_le, auto) + +lemma Inf_asclose: + fixes S :: "real set" + assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\Inf S - l\ \ e" +proof - + have "\- Sup (uminus ` S) - l\ = \Sup (uminus ` S) - (-l)\" + by auto + also have "... \ e" + apply (rule Sup_asclose) + apply (auto simp add: S) + apply (metis abs_minus_add_cancel b comm_monoid_add.mult_commute real_diff_def) + done + finally have "\- Sup (uminus ` S) - l\ \ e" . + thus ?thesis + by (simp add: Inf_real_def) +qed + +subsection{*Relate max and min to sup and inf.*} + +lemma real_max_Sup: + fixes x :: real + shows "max x y = Sup {x,y}" +proof- + have f: "finite {x, y}" "{x,y} \ {}" by simp_all + from Sup_finite_le_iff[OF f, of "max x y"] have "Sup {x,y} \ max x y" by simp + moreover + have "max x y \ Sup {x,y}" using Sup_finite_ge_iff[OF f, of "max x y"] + by (simp add: linorder_linear) + ultimately show ?thesis by arith +qed + +lemma real_min_Inf: + fixes x :: real + shows "min x y = Inf {x,y}" +proof- + have f: "finite {x, y}" "{x,y} \ {}" by simp_all + from Inf_finite_le_iff[OF f, of "min x y"] have "Inf {x,y} \ min x y" + by (simp add: linorder_linear) + moreover + have "min x y \ Inf {x,y}" using Inf_finite_ge_iff[OF f, of "min x y"] + by simp + ultimately show ?thesis by arith +qed + +end