paulson@33269: (* Author: Amine Chaieb and L C Paulson, University of Cambridge *) paulson@33269: paulson@33269: header {*Sup and Inf Operators on Sets of Reals.*} paulson@33269: paulson@33269: theory SupInf paulson@33269: imports RComplete paulson@33269: begin paulson@33269: paulson@33269: instantiation real :: Sup paulson@33269: begin paulson@33269: definition paulson@33269: Sup_real_def [code del]: "Sup X == (LEAST z::real. \x\X. x\z)" paulson@33269: paulson@33269: instance .. paulson@33269: end paulson@33269: paulson@33269: instantiation real :: Inf paulson@33269: begin paulson@33269: definition paulson@33269: Inf_real_def [code del]: "Inf (X::real set) == - (Sup (uminus ` X))" paulson@33269: paulson@33269: instance .. paulson@33269: end paulson@33269: paulson@33269: subsection{*Supremum of a set of reals*} paulson@33269: paulson@33269: lemma Sup_upper [intro]: (*REAL_SUP_UBOUND in HOL4*) paulson@33269: fixes x :: real paulson@33269: assumes x: "x \ X" paulson@33269: and z: "!!x. x \ X \ x \ z" paulson@33269: shows "x \ Sup X" paulson@33269: proof (auto simp add: Sup_real_def) paulson@33269: from reals_complete2 paulson@33269: obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" paulson@33269: by (blast intro: x z) paulson@33269: hence "x \ s" paulson@33269: by (blast intro: x z) paulson@33269: also with s have "... = (LEAST z. \x\X. x \ z)" paulson@33269: by (fast intro: Least_equality [symmetric]) paulson@33269: finally show "x \ (LEAST z. \x\X. x \ z)" . paulson@33269: qed paulson@33269: paulson@33269: lemma Sup_least [intro]: (*REAL_IMP_SUP_LE in HOL4*) paulson@33269: fixes z :: real paulson@33269: assumes x: "X \ {}" paulson@33269: and z: "\x. x \ X \ x \ z" paulson@33269: shows "Sup X \ z" paulson@33269: proof (auto simp add: Sup_real_def) paulson@33269: from reals_complete2 x paulson@33269: obtain s where s: "(\y\X. y \ s) & (\z. ((\y\X. y \ z) --> s \ z))" paulson@33269: by (blast intro: z) paulson@33269: hence "(LEAST z. \x\X. x \ z) = s" paulson@33269: by (best intro: Least_equality) paulson@33269: also with s z have "... \ z" paulson@33269: by blast paulson@33269: finally show "(LEAST z. \x\X. x \ z) \ z" . paulson@33269: qed paulson@33269: paulson@33609: lemma less_SupE: paulson@33609: fixes y :: real paulson@33609: assumes "y < Sup X" "X \ {}" paulson@33609: obtains x where "x\X" "y < x" paulson@33609: by (metis SupInf.Sup_least assms linorder_not_less that) paulson@33609: paulson@33269: lemma Sup_singleton [simp]: "Sup {x::real} = x" paulson@33269: by (force intro: Least_equality simp add: Sup_real_def) paulson@33269: paulson@33269: lemma Sup_eq_maximum: (*REAL_SUP_MAX in HOL4*) paulson@33269: fixes z :: real paulson@33269: assumes X: "z \ X" and z: "!!x. x \ X \ x \ z" paulson@33269: shows "Sup X = z" paulson@33269: by (force intro: Least_equality X z simp add: Sup_real_def) paulson@33269: paulson@33269: lemma Sup_upper2: (*REAL_IMP_LE_SUP in HOL4*) paulson@33269: fixes x :: real paulson@33269: shows "x \ X \ y \ x \ (!!x. x \ X \ x \ z) \ y \ Sup X" paulson@33269: by (metis Sup_upper real_le_trans) paulson@33269: paulson@33269: lemma Sup_real_iff : (*REAL_SUP_LE in HOL4*) paulson@33269: fixes z :: real paulson@33269: shows "X ~= {} ==> (!!x. x \ X ==> x \ z) ==> (\x\X. y y < Sup X" paulson@33269: by (metis Sup_least Sup_upper linorder_not_le le_less_trans) paulson@33269: paulson@33269: lemma Sup_eq: paulson@33269: fixes a :: real paulson@33269: shows "(!!x. x \ X \ x \ a) paulson@33269: \ (!!y. (!!x. x \ X \ x \ y) \ a \ y) \ Sup X = a" paulson@33269: by (metis Sup_least Sup_upper add_le_cancel_left diff_add_cancel insert_absorb nipkow@33657: insert_not_empty real_le_antisym) paulson@33269: paulson@33269: lemma Sup_le: paulson@33269: fixes S :: "real set" paulson@33269: shows "S \ {} \ S *<= b \ Sup S \ b" paulson@33269: by (metis SupInf.Sup_least setle_def) paulson@33269: paulson@33269: lemma Sup_upper_EX: paulson@33269: fixes x :: real paulson@33269: shows "x \ X \ \z. \x. x \ X \ x \ z \ x \ Sup X" paulson@33269: by blast paulson@33269: paulson@33269: lemma Sup_insert_nonempty: paulson@33269: fixes x :: real paulson@33269: assumes x: "x \ X" paulson@33269: and z: "!!x. x \ X \ x \ z" paulson@33269: shows "Sup (insert a X) = max a (Sup X)" paulson@33269: proof (cases "Sup X \ a") paulson@33269: case True paulson@33269: thus ?thesis paulson@33269: apply (simp add: max_def) paulson@33269: apply (rule Sup_eq_maximum) paulson@33269: apply (metis insertCI) paulson@33269: apply (metis Sup_upper insertE le_iff_sup real_le_linear real_le_trans sup_absorb1 z) paulson@33269: done paulson@33269: next paulson@33269: case False paulson@33269: hence 1:"a < Sup X" by simp paulson@33269: have "Sup X \ Sup (insert a X)" paulson@33269: apply (rule Sup_least) paulson@33269: apply (metis empty_psubset_nonempty psubset_eq x) paulson@33269: apply (rule Sup_upper_EX) paulson@33269: apply blast paulson@33269: apply (metis insert_iff real_le_linear real_le_refl real_le_trans z) paulson@33269: done paulson@33269: moreover paulson@33269: have "Sup (insert a X) \ Sup X" paulson@33269: apply (rule Sup_least) paulson@33269: apply blast paulson@33269: apply (metis False Sup_upper insertE real_le_linear z) paulson@33269: done paulson@33269: ultimately have "Sup (insert a X) = Sup X" paulson@33269: by (blast intro: antisym ) paulson@33269: thus ?thesis paulson@33269: by (metis 1 min_max.le_iff_sup real_less_def) paulson@33269: qed paulson@33269: paulson@33269: lemma Sup_insert_if: paulson@33269: fixes X :: "real set" paulson@33269: assumes z: "!!x. x \ X \ x \ z" paulson@33269: shows "Sup (insert a X) = (if X={} then a else max a (Sup X))" paulson@33269: by auto (metis Sup_insert_nonempty z) paulson@33269: paulson@33269: lemma Sup: paulson@33269: fixes S :: "real set" paulson@33269: shows "S \ {} \ (\b. S *<= b) \ isLub UNIV S (Sup S)" paulson@33269: by (auto simp add: isLub_def setle_def leastP_def isUb_def intro!: setgeI) paulson@33269: paulson@33269: lemma Sup_finite_Max: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "Sup S = Max S" paulson@33269: using fS Se paulson@33269: proof- paulson@33269: let ?m = "Max S" paulson@33269: from Max_ge[OF fS] have Sm: "\ x\ S. x \ ?m" by metis paulson@33269: with Sup[OF Se] have lub: "isLub UNIV S (Sup S)" by (metis setle_def) paulson@33269: from Max_in[OF fS Se] lub have mrS: "?m \ Sup S" paulson@33269: by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) paulson@33269: moreover paulson@33269: have "Sup S \ ?m" using Sm lub paulson@33269: by (auto simp add: isLub_def leastP_def isUb_def setle_def setge_def) paulson@33269: ultimately show ?thesis by arith paulson@33269: qed paulson@33269: paulson@33269: lemma Sup_finite_in: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "Sup S \ S" paulson@33269: using Sup_finite_Max[OF fS Se] Max_in[OF fS Se] by metis paulson@33269: paulson@33269: lemma Sup_finite_ge_iff: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "a \ Sup S \ (\ x \ S. a \ x)" paulson@33269: by (metis Max_ge Se Sup_finite_Max Sup_finite_in fS linorder_not_le less_le_trans) paulson@33269: paulson@33269: lemma Sup_finite_le_iff: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "a \ Sup S \ (\ x \ S. a \ x)" paulson@33269: by (metis Max_ge Se Sup_finite_Max Sup_finite_in fS le_iff_sup real_le_trans) paulson@33269: paulson@33269: lemma Sup_finite_gt_iff: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "a < Sup S \ (\ x \ S. a < x)" paulson@33269: by (metis Se Sup_finite_le_iff fS linorder_not_less) paulson@33269: paulson@33269: lemma Sup_finite_lt_iff: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "a > Sup S \ (\ x \ S. a > x)" paulson@33269: by (metis Se Sup_finite_ge_iff fS linorder_not_less) paulson@33269: paulson@33269: lemma Sup_unique: paulson@33269: fixes S :: "real set" paulson@33269: shows "S *<= b \ (\b' < b. \x \ S. b' < x) \ Sup S = b" paulson@33269: unfolding setle_def paulson@33269: apply (rule Sup_eq, auto) paulson@33269: apply (metis linorder_not_less) paulson@33269: done paulson@33269: paulson@33269: lemma Sup_abs_le: paulson@33269: fixes S :: "real set" paulson@33269: shows "S \ {} \ (\x\S. \x\ \ a) \ \Sup S\ \ a" paulson@33269: by (auto simp add: abs_le_interval_iff) (metis Sup_upper2) paulson@33269: paulson@33269: lemma Sup_bounds: paulson@33269: fixes S :: "real set" paulson@33269: assumes Se: "S \ {}" and l: "a <=* S" and u: "S *<= b" paulson@33269: shows "a \ Sup S \ Sup S \ b" paulson@33269: proof- paulson@33269: from Sup[OF Se] u have lub: "isLub UNIV S (Sup S)" by blast paulson@33269: hence b: "Sup S \ b" using u paulson@33269: by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) paulson@33269: from Se obtain y where y: "y \ S" by blast paulson@33269: from lub l have "a \ Sup S" paulson@33269: by (auto simp add: isLub_def leastP_def setle_def setge_def isUb_def) paulson@33269: (metis le_iff_sup le_sup_iff y) paulson@33269: with b show ?thesis by blast paulson@33269: qed paulson@33269: paulson@33269: lemma Sup_asclose: paulson@33269: fixes S :: "real set" paulson@33269: assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\Sup S - l\ \ e" paulson@33269: proof- paulson@33269: have th: "\(x::real) l e. \x - l\ \ e \ l - e \ x \ x \ l + e" by arith paulson@33269: thus ?thesis using S b Sup_bounds[of S "l - e" "l+e"] unfolding th paulson@33269: by (auto simp add: setge_def setle_def) paulson@33269: qed paulson@33269: paulson@33269: paulson@33269: subsection{*Infimum of a set of reals*} paulson@33269: paulson@33269: lemma Inf_lower [intro]: paulson@33269: fixes z :: real paulson@33269: assumes x: "x \ X" paulson@33269: and z: "!!x. x \ X \ z \ x" paulson@33269: shows "Inf X \ x" paulson@33269: proof - paulson@33269: have "-x \ Sup (uminus ` X)" paulson@33269: by (rule Sup_upper [where z = "-z"]) (auto simp add: image_iff x z) paulson@33269: thus ?thesis paulson@33269: by (auto simp add: Inf_real_def) paulson@33269: qed paulson@33269: paulson@33269: lemma Inf_greatest [intro]: paulson@33269: fixes z :: real paulson@33269: assumes x: "X \ {}" paulson@33269: and z: "\x. x \ X \ z \ x" paulson@33269: shows "z \ Inf X" paulson@33269: proof - huffman@35216: have "Sup (uminus ` X) \ -z" using x z by force paulson@33269: hence "z \ - Sup (uminus ` X)" paulson@33269: by simp paulson@33269: thus ?thesis paulson@33269: by (auto simp add: Inf_real_def) paulson@33269: qed paulson@33269: paulson@33269: lemma Inf_singleton [simp]: "Inf {x::real} = x" paulson@33269: by (simp add: Inf_real_def) paulson@33269: paulson@33269: lemma Inf_eq_minimum: (*REAL_INF_MIN in HOL4*) paulson@33269: fixes z :: real paulson@33269: assumes x: "z \ X" and z: "!!x. x \ X \ z \ x" paulson@33269: shows "Inf X = z" paulson@33269: proof - paulson@33269: have "Sup (uminus ` X) = -z" using x z paulson@33269: by (force intro: Sup_eq_maximum x z) paulson@33269: thus ?thesis paulson@33269: by (simp add: Inf_real_def) paulson@33269: qed paulson@33269: paulson@33269: lemma Inf_lower2: paulson@33269: fixes x :: real paulson@33269: shows "x \ X \ x \ y \ (!!x. x \ X \ z \ x) \ Inf X \ y" paulson@33269: by (metis Inf_lower real_le_trans) paulson@33269: paulson@33269: lemma Inf_real_iff: paulson@33269: fixes z :: real paulson@33269: shows "X \ {} \ (!!x. x \ X \ z \ x) \ (\x\X. x Inf X < y" paulson@33269: by (metis Inf_greatest Inf_lower less_le_not_le real_le_linear paulson@33269: order_less_le_trans) paulson@33269: paulson@33269: lemma Inf_eq: paulson@33269: fixes a :: real paulson@33269: shows "(!!x. x \ X \ a \ x) \ (!!y. (!!x. x \ X \ y \ x) \ y \ a) \ Inf X = a" paulson@33269: by (metis Inf_greatest Inf_lower add_le_cancel_left diff_add_cancel nipkow@33657: insert_absorb insert_not_empty real_le_antisym) paulson@33269: paulson@33269: lemma Inf_ge: paulson@33269: fixes S :: "real set" paulson@33269: shows "S \ {} \ b <=* S \ Inf S \ b" paulson@33269: by (metis SupInf.Inf_greatest setge_def) paulson@33269: paulson@33269: lemma Inf_lower_EX: paulson@33269: fixes x :: real paulson@33269: shows "x \ X \ \z. \x. x \ X \ z \ x \ Inf X \ x" paulson@33269: by blast paulson@33269: paulson@33269: lemma Inf_insert_nonempty: paulson@33269: fixes x :: real paulson@33269: assumes x: "x \ X" paulson@33269: and z: "!!x. x \ X \ z \ x" paulson@33269: shows "Inf (insert a X) = min a (Inf X)" paulson@33269: proof (cases "a \ Inf X") paulson@33269: case True paulson@33269: thus ?thesis paulson@33269: by (simp add: min_def) huffman@35216: (blast intro: Inf_eq_minimum real_le_refl real_le_trans z) paulson@33269: next paulson@33269: case False paulson@33269: hence 1:"Inf X < a" by simp paulson@33269: have "Inf (insert a X) \ Inf X" paulson@33269: apply (rule Inf_greatest) paulson@33269: apply (metis empty_psubset_nonempty psubset_eq x) paulson@33269: apply (rule Inf_lower_EX) paulson@33269: apply (blast intro: elim:) paulson@33269: apply (metis insert_iff real_le_linear real_le_refl real_le_trans z) paulson@33269: done paulson@33269: moreover paulson@33269: have "Inf X \ Inf (insert a X)" paulson@33269: apply (rule Inf_greatest) paulson@33269: apply blast paulson@33269: apply (metis False Inf_lower insertE real_le_linear z) paulson@33269: done paulson@33269: ultimately have "Inf (insert a X) = Inf X" paulson@33269: by (blast intro: antisym ) paulson@33269: thus ?thesis paulson@33269: by (metis False min_max.inf_absorb2 real_le_linear) paulson@33269: qed paulson@33269: paulson@33269: lemma Inf_insert_if: paulson@33269: fixes X :: "real set" paulson@33269: assumes z: "!!x. x \ X \ z \ x" paulson@33269: shows "Inf (insert a X) = (if X={} then a else min a (Inf X))" paulson@33269: by auto (metis Inf_insert_nonempty z) paulson@33269: paulson@33269: lemma Inf_greater: paulson@33269: fixes z :: real paulson@33269: shows "X \ {} \ Inf X < z \ \x \ X. x < z" paulson@33269: by (metis Inf_real_iff mem_def not_leE) paulson@33269: paulson@33269: lemma Inf_close: paulson@33269: fixes e :: real paulson@33269: shows "X \ {} \ 0 < e \ \x \ X. x < Inf X + e" paulson@33269: by (metis add_strict_increasing comm_monoid_add.mult_commute Inf_greater linorder_not_le pos_add_strict) paulson@33269: paulson@33269: lemma Inf_finite_Min: paulson@33269: fixes S :: "real set" paulson@33269: shows "finite S \ S \ {} \ Inf S = Min S" paulson@33269: by (simp add: Inf_real_def Sup_finite_Max image_image) paulson@33269: paulson@33269: lemma Inf_finite_in: paulson@33269: fixes S :: "real set" paulson@33269: assumes fS: "finite S" and Se: "S \ {}" paulson@33269: shows "Inf S \ S" paulson@33269: using Inf_finite_Min[OF fS Se] Min_in[OF fS Se] by metis paulson@33269: paulson@33269: lemma Inf_finite_ge_iff: paulson@33269: fixes S :: "real set" paulson@33269: shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" paulson@33269: by (metis Inf_finite_Min Inf_finite_in Min_le real_le_trans) paulson@33269: paulson@33269: lemma Inf_finite_le_iff: paulson@33269: fixes S :: "real set" paulson@33269: shows "finite S \ S \ {} \ a \ Inf S \ (\ x \ S. a \ x)" paulson@33269: by (metis Inf_finite_Min Inf_finite_ge_iff Inf_finite_in Min_le nipkow@33657: real_le_antisym real_le_linear) paulson@33269: paulson@33269: lemma Inf_finite_gt_iff: paulson@33269: fixes S :: "real set" paulson@33269: shows "finite S \ S \ {} \ a < Inf S \ (\ x \ S. a < x)" paulson@33269: by (metis Inf_finite_le_iff linorder_not_less) paulson@33269: paulson@33269: lemma Inf_finite_lt_iff: paulson@33269: fixes S :: "real set" paulson@33269: shows "finite S \ S \ {} \ a > Inf S \ (\ x \ S. a > x)" paulson@33269: by (metis Inf_finite_ge_iff linorder_not_less) paulson@33269: paulson@33269: lemma Inf_unique: paulson@33269: fixes S :: "real set" paulson@33269: shows "b <=* S \ (\b' > b. \x \ S. b' > x) \ Inf S = b" paulson@33269: unfolding setge_def paulson@33269: apply (rule Inf_eq, auto) paulson@33269: apply (metis less_le_not_le linorder_not_less) paulson@33269: done paulson@33269: paulson@33269: lemma Inf_abs_ge: paulson@33269: fixes S :: "real set" paulson@33269: shows "S \ {} \ (\x\S. \x\ \ a) \ \Inf S\ \ a" paulson@33269: by (simp add: Inf_real_def) (rule Sup_abs_le, auto) paulson@33269: paulson@33269: lemma Inf_asclose: paulson@33269: fixes S :: "real set" paulson@33269: assumes S:"S \ {}" and b: "\x\S. \x - l\ \ e" shows "\Inf S - l\ \ e" paulson@33269: proof - paulson@33269: have "\- Sup (uminus ` S) - l\ = \Sup (uminus ` S) - (-l)\" paulson@33269: by auto paulson@33269: also have "... \ e" paulson@33269: apply (rule Sup_asclose) paulson@33269: apply (auto simp add: S) paulson@33269: apply (metis abs_minus_add_cancel b comm_monoid_add.mult_commute real_diff_def) paulson@33269: done paulson@33269: finally have "\- Sup (uminus ` S) - l\ \ e" . paulson@33269: thus ?thesis paulson@33269: by (simp add: Inf_real_def) paulson@33269: qed paulson@33269: paulson@33271: subsection{*Relate max and min to Sup and Inf.*} paulson@33269: paulson@33269: lemma real_max_Sup: paulson@33269: fixes x :: real paulson@33269: shows "max x y = Sup {x,y}" paulson@33269: proof- paulson@33269: have f: "finite {x, y}" "{x,y} \ {}" by simp_all paulson@33269: from Sup_finite_le_iff[OF f, of "max x y"] have "Sup {x,y} \ max x y" by simp paulson@33269: moreover paulson@33269: have "max x y \ Sup {x,y}" using Sup_finite_ge_iff[OF f, of "max x y"] paulson@33269: by (simp add: linorder_linear) paulson@33269: ultimately show ?thesis by arith paulson@33269: qed paulson@33269: paulson@33269: lemma real_min_Inf: paulson@33269: fixes x :: real paulson@33269: shows "min x y = Inf {x,y}" paulson@33269: proof- paulson@33269: have f: "finite {x, y}" "{x,y} \ {}" by simp_all paulson@33269: from Inf_finite_le_iff[OF f, of "min x y"] have "Inf {x,y} \ min x y" paulson@33269: by (simp add: linorder_linear) paulson@33269: moreover paulson@33269: have "min x y \ Inf {x,y}" using Inf_finite_ge_iff[OF f, of "min x y"] paulson@33269: by simp paulson@33269: ultimately show ?thesis by arith paulson@33269: qed paulson@33269: paulson@33609: lemma reals_complete_interval: paulson@33609: fixes a::real and b::real paulson@33609: assumes "a < b" and "P a" and "~P b" paulson@33609: shows "\c. a \ c & c \ b & (\x. a \ x & x < c --> P x) & paulson@33609: (\d. (\x. a \ x & x < d --> P x) --> d \ c)" paulson@33609: proof (rule exI [where x = "Sup {d. \x. a \ x & x < d --> P x}"], auto) paulson@33609: show "a \ Sup {d. \c. a \ c \ c < d \ P c}" paulson@33609: by (rule SupInf.Sup_upper [where z=b], auto) huffman@35216: (metis `a < b` `\ P b` real_le_linear real_less_def) paulson@33609: next paulson@33609: show "Sup {d. \c. a \ c \ c < d \ P c} \ b" paulson@33609: apply (rule SupInf.Sup_least) paulson@33609: apply (auto simp add: ) paulson@33609: apply (metis less_le_not_le) paulson@33609: apply (metis `a x" and lt: "x < Sup {d. \c. a \ c \ c < d \ P c}" paulson@33609: show "P x" paulson@33609: apply (rule less_SupE [OF lt], auto) paulson@33609: apply (metis less_le_not_le) paulson@33609: apply (metis x) paulson@33609: done paulson@33609: next paulson@33609: fix d paulson@33609: assume 0: "\x. a \ x \ x < d \ P x" paulson@33609: thus "d \ Sup {d. \c. a \ c \ c < d \ P c}" paulson@33609: by (rule_tac z="b" in SupInf.Sup_upper, auto) paulson@33609: (metis `a