diff -r 6ca79a354c51 -r f651cb053927 src/HOL/Library/Extended_Reals.thy --- a/src/HOL/Library/Extended_Reals.thy Wed Jul 20 10:11:08 2011 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2535 +0,0 @@ -(* Title: HOL/Library/Extended_Reals.thy - Author: Johannes Hölzl, TU München - Author: Robert Himmelmann, TU München - Author: Armin Heller, TU München - Author: Bogdan Grechuk, University of Edinburgh -*) - -header {* Extended real number line *} - -theory Extended_Reals - imports Complex_Main -begin - -text {* - -For more lemmas about the extended real numbers go to - @{text "src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"} - -*} - -lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" -proof - assume "{x..} = UNIV" - show "x = bot" - proof (rule ccontr) - assume "x \ bot" then have "bot \ {x..}" by (simp add: le_less) - then show False using `{x..} = UNIV` by simp - qed -qed auto - -lemma SUPR_pair: - "(SUP i : A. SUP j : B. f i j) = (SUP p : A \ B. f (fst p) (snd p))" - by (rule antisym) (auto intro!: SUP_leI le_SUPI2) - -lemma INFI_pair: - "(INF i : A. INF j : B. f i j) = (INF p : A \ B. f (fst p) (snd p))" - by (rule antisym) (auto intro!: le_INFI INF_leI2) - -subsection {* Definition and basic properties *} - -datatype extreal = extreal real | PInfty | MInfty - -notation (xsymbols) - PInfty ("\") - -notation (HTML output) - PInfty ("\") - -declare [[coercion "extreal :: real \ extreal"]] - -instantiation extreal :: uminus -begin - fun uminus_extreal where - "- (extreal r) = extreal (- r)" - | "- \ = MInfty" - | "- MInfty = \" - instance .. -end - -lemma inj_extreal[simp]: "inj_on extreal A" - unfolding inj_on_def by auto - -lemma MInfty_neq_PInfty[simp]: - "\ \ - \" "- \ \ \" by simp_all - -lemma MInfty_neq_extreal[simp]: - "extreal r \ - \" "- \ \ extreal r" by simp_all - -lemma MInfinity_cases[simp]: - "(case - \ of extreal r \ f r | \ \ y | MInfinity \ z) = z" - by simp - -lemma extreal_uminus_uminus[simp]: - fixes a :: extreal shows "- (- a) = a" - by (cases a) simp_all - -lemma MInfty_eq[simp, code_post]: - "MInfty = - \" by simp - -declare uminus_extreal.simps(2)[code_inline, simp del] - -lemma extreal_cases[case_names real PInf MInf, cases type: extreal]: - assumes "\r. x = extreal r \ P" - assumes "x = \ \ P" - assumes "x = -\ \ P" - shows P - using assms by (cases x) auto - -lemmas extreal2_cases = extreal_cases[case_product extreal_cases] -lemmas extreal3_cases = extreal2_cases[case_product extreal_cases] - -lemma extreal_uminus_eq_iff[simp]: - fixes a b :: extreal shows "-a = -b \ a = b" - by (cases rule: extreal2_cases[of a b]) simp_all - -function of_extreal :: "extreal \ real" where -"of_extreal (extreal r) = r" | -"of_extreal \ = 0" | -"of_extreal (-\) = 0" - by (auto intro: extreal_cases) -termination proof qed (rule wf_empty) - -defs (overloaded) - real_of_extreal_def [code_unfold]: "real \ of_extreal" - -lemma real_of_extreal[simp]: - "real (- x :: extreal) = - (real x)" - "real (extreal r) = r" - "real \ = 0" - by (cases x) (simp_all add: real_of_extreal_def) - -lemma range_extreal[simp]: "range extreal = UNIV - {\, -\}" -proof safe - fix x assume "x \ range extreal" "x \ \" - then show "x = -\" by (cases x) auto -qed auto - -lemma extreal_range_uminus[simp]: "range uminus = (UNIV::extreal set)" -proof safe - fix x :: extreal show "x \ range uminus" by (intro image_eqI[of _ _ "-x"]) auto -qed auto - -instantiation extreal :: number -begin -definition [simp]: "number_of x = extreal (number_of x)" -instance proof qed -end - -instantiation extreal :: abs -begin - function abs_extreal where - "\extreal r\ = extreal \r\" - | "\-\\ = \" - | "\\\ = \" - by (auto intro: extreal_cases) - termination proof qed (rule wf_empty) - instance .. -end - -lemma abs_eq_infinity_cases[elim!]: "\ \x\ = \ ; x = \ \ P ; x = -\ \ P \ \ P" - by (cases x) auto - -lemma abs_neq_infinity_cases[elim!]: "\ \x\ \ \ ; \r. x = extreal r \ P \ \ P" - by (cases x) auto - -lemma abs_extreal_uminus[simp]: "\- x\ = \x::extreal\" - by (cases x) auto - -subsubsection "Addition" - -instantiation extreal :: comm_monoid_add -begin - -definition "0 = extreal 0" - -function plus_extreal where -"extreal r + extreal p = extreal (r + p)" | -"\ + a = \" | -"a + \ = \" | -"extreal r + -\ = - \" | -"-\ + extreal p = -\" | -"-\ + -\ = -\" -proof - - case (goal1 P x) - moreover then obtain a b where "x = (a, b)" by (cases x) auto - ultimately show P - by (cases rule: extreal2_cases[of a b]) auto -qed auto -termination proof qed (rule wf_empty) - -lemma Infty_neq_0[simp]: - "\ \ 0" "0 \ \" - "-\ \ 0" "0 \ -\" - by (simp_all add: zero_extreal_def) - -lemma extreal_eq_0[simp]: - "extreal r = 0 \ r = 0" - "0 = extreal r \ r = 0" - unfolding zero_extreal_def by simp_all - -instance -proof - fix a :: extreal show "0 + a = a" - by (cases a) (simp_all add: zero_extreal_def) - fix b :: extreal show "a + b = b + a" - by (cases rule: extreal2_cases[of a b]) simp_all - fix c :: extreal show "a + b + c = a + (b + c)" - by (cases rule: extreal3_cases[of a b c]) simp_all -qed -end - -lemma real_of_extreal_0[simp]: "real (0::extreal) = 0" - unfolding real_of_extreal_def zero_extreal_def by simp - -lemma abs_extreal_zero[simp]: "\0\ = (0::extreal)" - unfolding zero_extreal_def abs_extreal.simps by simp - -lemma extreal_uminus_zero[simp]: - "- 0 = (0::extreal)" - by (simp add: zero_extreal_def) - -lemma extreal_uminus_zero_iff[simp]: - fixes a :: extreal shows "-a = 0 \ a = 0" - by (cases a) simp_all - -lemma extreal_plus_eq_PInfty[simp]: - shows "a + b = \ \ a = \ \ b = \" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_plus_eq_MInfty[simp]: - shows "a + b = -\ \ - (a = -\ \ b = -\) \ a \ \ \ b \ \" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_add_cancel_left: - assumes "a \ -\" - shows "a + b = a + c \ (a = \ \ b = c)" - using assms by (cases rule: extreal3_cases[of a b c]) auto - -lemma extreal_add_cancel_right: - assumes "a \ -\" - shows "b + a = c + a \ (a = \ \ b = c)" - using assms by (cases rule: extreal3_cases[of a b c]) auto - -lemma extreal_real: - "extreal (real x) = (if \x\ = \ then 0 else x)" - by (cases x) simp_all - -lemma real_of_extreal_add: - fixes a b :: extreal - shows "real (a + b) = (if (\a\ = \) \ (\b\ = \) \ (\a\ \ \) \ (\b\ \ \) then real a + real b else 0)" - by (cases rule: extreal2_cases[of a b]) auto - -subsubsection "Linear order on @{typ extreal}" - -instantiation extreal :: linorder -begin - -function less_extreal where -"extreal x < extreal y \ x < y" | -" \ < a \ False" | -" a < -\ \ False" | -"extreal x < \ \ True" | -" -\ < extreal r \ True" | -" -\ < \ \ True" -proof - - case (goal1 P x) - moreover then obtain a b where "x = (a,b)" by (cases x) auto - ultimately show P by (cases rule: extreal2_cases[of a b]) auto -qed simp_all -termination by (relation "{}") simp - -definition "x \ (y::extreal) \ x < y \ x = y" - -lemma extreal_infty_less[simp]: - "x < \ \ (x \ \)" - "-\ < x \ (x \ -\)" - by (cases x, simp_all) (cases x, simp_all) - -lemma extreal_infty_less_eq[simp]: - "\ \ x \ x = \" - "x \ -\ \ x = -\" - by (auto simp add: less_eq_extreal_def) - -lemma extreal_less[simp]: - "extreal r < 0 \ (r < 0)" - "0 < extreal r \ (0 < r)" - "0 < \" - "-\ < 0" - by (simp_all add: zero_extreal_def) - -lemma extreal_less_eq[simp]: - "x \ \" - "-\ \ x" - "extreal r \ extreal p \ r \ p" - "extreal r \ 0 \ r \ 0" - "0 \ extreal r \ 0 \ r" - by (auto simp add: less_eq_extreal_def zero_extreal_def) - -lemma extreal_infty_less_eq2: - "a \ b \ a = \ \ b = \" - "a \ b \ b = -\ \ a = -\" - by simp_all - -instance -proof - fix x :: extreal show "x \ x" - by (cases x) simp_all - fix y :: extreal show "x < y \ x \ y \ \ y \ x" - by (cases rule: extreal2_cases[of x y]) auto - show "x \ y \ y \ x " - by (cases rule: extreal2_cases[of x y]) auto - { assume "x \ y" "y \ x" then show "x = y" - by (cases rule: extreal2_cases[of x y]) auto } - { fix z assume "x \ y" "y \ z" then show "x \ z" - by (cases rule: extreal3_cases[of x y z]) auto } -qed -end - -instance extreal :: ordered_ab_semigroup_add -proof - fix a b c :: extreal assume "a \ b" then show "c + a \ c + b" - by (cases rule: extreal3_cases[of a b c]) auto -qed - -lemma real_of_extreal_positive_mono: - "\0 \ x; x \ y; y \ \\ \ real x \ real y" - by (cases rule: extreal2_cases[of x y]) auto - -lemma extreal_MInfty_lessI[intro, simp]: - "a \ -\ \ -\ < a" - by (cases a) auto - -lemma extreal_less_PInfty[intro, simp]: - "a \ \ \ a < \" - by (cases a) auto - -lemma extreal_less_extreal_Ex: - fixes a b :: extreal - shows "x < extreal r \ x = -\ \ (\p. p < r \ x = extreal p)" - by (cases x) auto - -lemma less_PInf_Ex_of_nat: "x \ \ \ (\n::nat. x < extreal (real n))" -proof (cases x) - case (real r) then show ?thesis - using reals_Archimedean2[of r] by simp -qed simp_all - -lemma extreal_add_mono: - fixes a b c d :: extreal assumes "a \ b" "c \ d" shows "a + c \ b + d" - using assms - apply (cases a) - apply (cases rule: extreal3_cases[of b c d], auto) - apply (cases rule: extreal3_cases[of b c d], auto) - done - -lemma extreal_minus_le_minus[simp]: - fixes a b :: extreal shows "- a \ - b \ b \ a" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_minus_less_minus[simp]: - fixes a b :: extreal shows "- a < - b \ b < a" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_le_real_iff: - "x \ real y \ ((\y\ \ \ \ extreal x \ y) \ (\y\ = \ \ x \ 0))" - by (cases y) auto - -lemma real_le_extreal_iff: - "real y \ x \ ((\y\ \ \ \ y \ extreal x) \ (\y\ = \ \ 0 \ x))" - by (cases y) auto - -lemma extreal_less_real_iff: - "x < real y \ ((\y\ \ \ \ extreal x < y) \ (\y\ = \ \ x < 0))" - by (cases y) auto - -lemma real_less_extreal_iff: - "real y < x \ ((\y\ \ \ \ y < extreal x) \ (\y\ = \ \ 0 < x))" - by (cases y) auto - -lemma real_of_extreal_pos: - fixes x :: extreal shows "0 \ x \ 0 \ real x" by (cases x) auto - -lemmas real_of_extreal_ord_simps = - extreal_le_real_iff real_le_extreal_iff extreal_less_real_iff real_less_extreal_iff - -lemma abs_extreal_ge0[simp]: "0 \ x \ \x :: extreal\ = x" - by (cases x) auto - -lemma abs_extreal_less0[simp]: "x < 0 \ \x :: extreal\ = -x" - by (cases x) auto - -lemma abs_extreal_pos[simp]: "0 \ \x :: extreal\" - by (cases x) auto - -lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \ 0 \ (X \ 0 \ X = \)" - by (cases X) auto - -lemma abs_real_of_extreal[simp]: "\real (X :: extreal)\ = real \X\" - by (cases X) auto - -lemma zero_less_real_of_extreal: "0 < real X \ (0 < X \ X \ \)" - by (cases X) auto - -lemma extreal_0_le_uminus_iff[simp]: - fixes a :: extreal shows "0 \ -a \ a \ 0" - by (cases rule: extreal2_cases[of a]) auto - -lemma extreal_uminus_le_0_iff[simp]: - fixes a :: extreal shows "-a \ 0 \ 0 \ a" - by (cases rule: extreal2_cases[of a]) auto - -lemma extreal_dense: - fixes x y :: extreal assumes "x < y" - shows "EX z. x < z & z < y" -proof - -{ assume a: "x = (-\)" - { assume "y = \" hence ?thesis using a by (auto intro!: exI[of _ "0"]) } - moreover - { assume "y ~= \" - with `x < y` obtain r where r: "y = extreal r" by (cases y) auto - hence ?thesis using `x < y` a by (auto intro!: exI[of _ "extreal (r - 1)"]) - } ultimately have ?thesis by auto -} -moreover -{ assume "x ~= (-\)" - with `x < y` obtain p where p: "x = extreal p" by (cases x) auto - { assume "y = \" hence ?thesis using `x < y` p - by (auto intro!: exI[of _ "extreal (p + 1)"]) } - moreover - { assume "y ~= \" - with `x < y` obtain r where r: "y = extreal r" by (cases y) auto - with p `x < y` have "p < r" by auto - with dense obtain z where "p < z" "z < r" by auto - hence ?thesis using r p by (auto intro!: exI[of _ "extreal z"]) - } ultimately have ?thesis by auto -} ultimately show ?thesis by auto -qed - -lemma extreal_dense2: - fixes x y :: extreal assumes "x < y" - shows "EX z. x < extreal z & extreal z < y" - by (metis extreal_dense[OF `x < y`] extreal_cases less_extreal.simps(2,3)) - -lemma extreal_add_strict_mono: - fixes a b c d :: extreal - assumes "a = b" "0 \ a" "a \ \" "c < d" - shows "a + c < b + d" - using assms by (cases rule: extreal3_cases[case_product extreal_cases, of a b c d]) auto - -lemma extreal_less_add: "\a\ \ \ \ c < b \ a + c < a + b" - by (cases rule: extreal2_cases[of b c]) auto - -lemma extreal_uminus_eq_reorder: "- a = b \ a = (-b::extreal)" by auto - -lemma extreal_uminus_less_reorder: "- a < b \ -b < (a::extreal)" - by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_less_minus) - -lemma extreal_uminus_le_reorder: "- a \ b \ -b \ (a::extreal)" - by (subst (3) extreal_uminus_uminus[symmetric]) (simp only: extreal_minus_le_minus) - -lemmas extreal_uminus_reorder = - extreal_uminus_eq_reorder extreal_uminus_less_reorder extreal_uminus_le_reorder - -lemma extreal_bot: - fixes x :: extreal assumes "\B. x \ extreal B" shows "x = - \" -proof (cases x) - case (real r) with assms[of "r - 1"] show ?thesis by auto -next case PInf with assms[of 0] show ?thesis by auto -next case MInf then show ?thesis by simp -qed - -lemma extreal_top: - fixes x :: extreal assumes "\B. x \ extreal B" shows "x = \" -proof (cases x) - case (real r) with assms[of "r + 1"] show ?thesis by auto -next case MInf with assms[of 0] show ?thesis by auto -next case PInf then show ?thesis by simp -qed - -lemma - shows extreal_max[simp]: "extreal (max x y) = max (extreal x) (extreal y)" - and extreal_min[simp]: "extreal (min x y) = min (extreal x) (extreal y)" - by (simp_all add: min_def max_def) - -lemma extreal_max_0: "max 0 (extreal r) = extreal (max 0 r)" - by (auto simp: zero_extreal_def) - -lemma - fixes f :: "nat \ extreal" - shows incseq_uminus[simp]: "incseq (\x. - f x) \ decseq f" - and decseq_uminus[simp]: "decseq (\x. - f x) \ incseq f" - unfolding decseq_def incseq_def by auto - -lemma incseq_extreal: "incseq f \ incseq (\x. extreal (f x))" - unfolding incseq_def by auto - -lemma extreal_add_nonneg_nonneg: - fixes a b :: extreal shows "0 \ a \ 0 \ b \ 0 \ a + b" - using add_mono[of 0 a 0 b] by simp - -lemma image_eqD: "f ` A = B \ (\x\A. f x \ B)" - by auto - -lemma incseq_setsumI: - fixes f :: "nat \ 'a::{comm_monoid_add, ordered_ab_semigroup_add}" - assumes "\i. 0 \ f i" - shows "incseq (\i. setsum f {..< i})" -proof (intro incseq_SucI) - fix n have "setsum f {..< n} + 0 \ setsum f {.. setsum f {..< Suc n}" - by auto -qed - -lemma incseq_setsumI2: - fixes f :: "'i \ nat \ 'a::{comm_monoid_add, ordered_ab_semigroup_add}" - assumes "\n. n \ A \ incseq (f n)" - shows "incseq (\i. \n\A. f n i)" - using assms unfolding incseq_def by (auto intro: setsum_mono) - -subsubsection "Multiplication" - -instantiation extreal :: "{comm_monoid_mult, sgn}" -begin - -definition "1 = extreal 1" - -function sgn_extreal where - "sgn (extreal r) = extreal (sgn r)" -| "sgn \ = 1" -| "sgn (-\) = -1" -by (auto intro: extreal_cases) -termination proof qed (rule wf_empty) - -function times_extreal where -"extreal r * extreal p = extreal (r * p)" | -"extreal r * \ = (if r = 0 then 0 else if r > 0 then \ else -\)" | -"\ * extreal r = (if r = 0 then 0 else if r > 0 then \ else -\)" | -"extreal r * -\ = (if r = 0 then 0 else if r > 0 then -\ else \)" | -"-\ * extreal r = (if r = 0 then 0 else if r > 0 then -\ else \)" | -"\ * \ = \" | -"-\ * \ = -\" | -"\ * -\ = -\" | -"-\ * -\ = \" -proof - - case (goal1 P x) - moreover then obtain a b where "x = (a, b)" by (cases x) auto - ultimately show P by (cases rule: extreal2_cases[of a b]) auto -qed simp_all -termination by (relation "{}") simp - -instance -proof - fix a :: extreal show "1 * a = a" - by (cases a) (simp_all add: one_extreal_def) - fix b :: extreal show "a * b = b * a" - by (cases rule: extreal2_cases[of a b]) simp_all - fix c :: extreal show "a * b * c = a * (b * c)" - by (cases rule: extreal3_cases[of a b c]) - (simp_all add: zero_extreal_def zero_less_mult_iff) -qed -end - -lemma real_of_extreal_le_1: - fixes a :: extreal shows "a \ 1 \ real a \ 1" - by (cases a) (auto simp: one_extreal_def) - -lemma abs_extreal_one[simp]: "\1\ = (1::extreal)" - unfolding one_extreal_def by simp - -lemma extreal_mult_zero[simp]: - fixes a :: extreal shows "a * 0 = 0" - by (cases a) (simp_all add: zero_extreal_def) - -lemma extreal_zero_mult[simp]: - fixes a :: extreal shows "0 * a = 0" - by (cases a) (simp_all add: zero_extreal_def) - -lemma extreal_m1_less_0[simp]: - "-(1::extreal) < 0" - by (simp add: zero_extreal_def one_extreal_def) - -lemma extreal_zero_m1[simp]: - "1 \ (0::extreal)" - by (simp add: zero_extreal_def one_extreal_def) - -lemma extreal_times_0[simp]: - fixes x :: extreal shows "0 * x = 0" - by (cases x) (auto simp: zero_extreal_def) - -lemma extreal_times[simp]: - "1 \ \" "\ \ 1" - "1 \ -\" "-\ \ 1" - by (auto simp add: times_extreal_def one_extreal_def) - -lemma extreal_plus_1[simp]: - "1 + extreal r = extreal (r + 1)" "extreal r + 1 = extreal (r + 1)" - "1 + -\ = -\" "-\ + 1 = -\" - unfolding one_extreal_def by auto - -lemma extreal_zero_times[simp]: - fixes a b :: extreal shows "a * b = 0 \ a = 0 \ b = 0" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_mult_eq_PInfty[simp]: - shows "a * b = \ \ - (a = \ \ b > 0) \ (a > 0 \ b = \) \ (a = -\ \ b < 0) \ (a < 0 \ b = -\)" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_mult_eq_MInfty[simp]: - shows "a * b = -\ \ - (a = \ \ b < 0) \ (a < 0 \ b = \) \ (a = -\ \ b > 0) \ (a > 0 \ b = -\)" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_0_less_1[simp]: "0 < (1::extreal)" - by (simp_all add: zero_extreal_def one_extreal_def) - -lemma extreal_zero_one[simp]: "0 \ (1::extreal)" - by (simp_all add: zero_extreal_def one_extreal_def) - -lemma extreal_mult_minus_left[simp]: - fixes a b :: extreal shows "-a * b = - (a * b)" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_mult_minus_right[simp]: - fixes a b :: extreal shows "a * -b = - (a * b)" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_mult_infty[simp]: - "a * \ = (if a = 0 then 0 else if 0 < a then \ else - \)" - by (cases a) auto - -lemma extreal_infty_mult[simp]: - "\ * a = (if a = 0 then 0 else if 0 < a then \ else - \)" - by (cases a) auto - -lemma extreal_mult_strict_right_mono: - assumes "a < b" and "0 < c" "c < \" - shows "a * c < b * c" - using assms - by (cases rule: extreal3_cases[of a b c]) - (auto simp: zero_le_mult_iff extreal_less_PInfty) - -lemma extreal_mult_strict_left_mono: - "\ a < b ; 0 < c ; c < \\ \ c * a < c * b" - using extreal_mult_strict_right_mono by (simp add: mult_commute[of c]) - -lemma extreal_mult_right_mono: - fixes a b c :: extreal shows "\a \ b; 0 \ c\ \ a*c \ b*c" - using assms - apply (cases "c = 0") apply simp - by (cases rule: extreal3_cases[of a b c]) - (auto simp: zero_le_mult_iff extreal_less_PInfty) - -lemma extreal_mult_left_mono: - fixes a b c :: extreal shows "\a \ b; 0 \ c\ \ c * a \ c * b" - using extreal_mult_right_mono by (simp add: mult_commute[of c]) - -lemma zero_less_one_extreal[simp]: "0 \ (1::extreal)" - by (simp add: one_extreal_def zero_extreal_def) - -lemma extreal_0_le_mult[simp]: "0 \ a \ 0 \ b \ 0 \ a * (b :: extreal)" - by (cases rule: extreal2_cases[of a b]) (auto simp: mult_nonneg_nonneg) - -lemma extreal_right_distrib: - fixes r a b :: extreal shows "0 \ a \ 0 \ b \ r * (a + b) = r * a + r * b" - by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps) - -lemma extreal_left_distrib: - fixes r a b :: extreal shows "0 \ a \ 0 \ b \ (a + b) * r = a * r + b * r" - by (cases rule: extreal3_cases[of r a b]) (simp_all add: field_simps) - -lemma extreal_mult_le_0_iff: - fixes a b :: extreal - shows "a * b \ 0 \ (0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b)" - by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_le_0_iff) - -lemma extreal_zero_le_0_iff: - fixes a b :: extreal - shows "0 \ a * b \ (0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0)" - by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_le_mult_iff) - -lemma extreal_mult_less_0_iff: - fixes a b :: extreal - shows "a * b < 0 \ (0 < a \ b < 0) \ (a < 0 \ 0 < b)" - by (cases rule: extreal2_cases[of a b]) (simp_all add: mult_less_0_iff) - -lemma extreal_zero_less_0_iff: - fixes a b :: extreal - shows "0 < a * b \ (0 < a \ 0 < b) \ (a < 0 \ b < 0)" - by (cases rule: extreal2_cases[of a b]) (simp_all add: zero_less_mult_iff) - -lemma extreal_distrib: - fixes a b c :: extreal - assumes "a \ \ \ b \ -\" "a \ -\ \ b \ \" "\c\ \ \" - shows "(a + b) * c = a * c + b * c" - using assms - by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps) - -lemma extreal_le_epsilon: - fixes x y :: extreal - assumes "ALL e. 0 < e --> x <= y + e" - shows "x <= y" -proof- -{ assume a: "EX r. y = extreal r" - from this obtain r where r_def: "y = extreal r" by auto - { assume "x=(-\)" hence ?thesis by auto } - moreover - { assume "~(x=(-\))" - from this obtain p where p_def: "x = extreal p" - using a assms[rule_format, of 1] by (cases x) auto - { fix e have "0 < e --> p <= r + e" - using assms[rule_format, of "extreal e"] p_def r_def by auto } - hence "p <= r" apply (subst field_le_epsilon) by auto - hence ?thesis using r_def p_def by auto - } ultimately have ?thesis by blast -} -moreover -{ assume "y=(-\) | y=\" hence ?thesis - using assms[rule_format, of 1] by (cases x) auto -} ultimately show ?thesis by (cases y) auto -qed - - -lemma extreal_le_epsilon2: - fixes x y :: extreal - assumes "ALL e. 0 < e --> x <= y + extreal e" - shows "x <= y" -proof- -{ fix e :: extreal assume "e>0" - { assume "e=\" hence "x<=y+e" by auto } - moreover - { assume "e~=\" - from this obtain r where "e = extreal r" using `e>0` apply (cases e) by auto - hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto - } ultimately have "x<=y+e" by blast -} from this show ?thesis using extreal_le_epsilon by auto -qed - -lemma extreal_le_real: - fixes x y :: extreal - assumes "ALL z. x <= extreal z --> y <= extreal z" - shows "y <= x" -by (metis assms extreal.exhaust extreal_bot extreal_less_eq(1) - extreal_less_eq(2) order_refl uminus_extreal.simps(2)) - -lemma extreal_le_extreal: - fixes x y :: extreal - assumes "\B. B < x \ B <= y" - shows "x <= y" -by (metis assms extreal_dense leD linorder_le_less_linear) - -lemma extreal_ge_extreal: - fixes x y :: extreal - assumes "ALL B. B>x --> B >= y" - shows "x >= y" -by (metis assms extreal_dense leD linorder_le_less_linear) - -lemma setprod_extreal_0: - fixes f :: "'a \ extreal" - shows "(\i\A. f i) = 0 \ (finite A \ (\i\A. f i = 0))" -proof cases - assume "finite A" - then show ?thesis by (induct A) auto -qed auto - -lemma setprod_extreal_pos: - fixes f :: "'a \ extreal" assumes pos: "\i. i \ I \ 0 \ f i" shows "0 \ (\i\I. f i)" -proof cases - assume "finite I" from this pos show ?thesis by induct auto -qed simp - -lemma setprod_PInf: - assumes "\i. i \ I \ 0 \ f i" - shows "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" -proof cases - assume "finite I" from this assms show ?thesis - proof (induct I) - case (insert i I) - then have pos: "0 \ f i" "0 \ setprod f I" by (auto intro!: setprod_extreal_pos) - from insert have "(\j\insert i I. f j) = \ \ setprod f I * f i = \" by auto - also have "\ \ (setprod f I = \ \ f i = \) \ f i \ 0 \ setprod f I \ 0" - using setprod_extreal_pos[of I f] pos - by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto - also have "\ \ finite (insert i I) \ (\j\insert i I. f j = \) \ (\j\insert i I. f j \ 0)" - using insert by (auto simp: setprod_extreal_0) - finally show ?case . - qed simp -qed simp - -lemma setprod_extreal: "(\i\A. extreal (f i)) = extreal (setprod f A)" -proof cases - assume "finite A" then show ?thesis - by induct (auto simp: one_extreal_def) -qed (simp add: one_extreal_def) - -subsubsection {* Power *} - -instantiation extreal :: power -begin -primrec power_extreal where - "power_extreal x 0 = 1" | - "power_extreal x (Suc n) = x * x ^ n" -instance .. -end - -lemma extreal_power[simp]: "(extreal x) ^ n = extreal (x^n)" - by (induct n) (auto simp: one_extreal_def) - -lemma extreal_power_PInf[simp]: "\ ^ n = (if n = 0 then 1 else \)" - by (induct n) (auto simp: one_extreal_def) - -lemma extreal_power_uminus[simp]: - fixes x :: extreal - shows "(- x) ^ n = (if even n then x ^ n else - (x^n))" - by (induct n) (auto simp: one_extreal_def) - -lemma extreal_power_number_of[simp]: - "(number_of num :: extreal) ^ n = extreal (number_of num ^ n)" - by (induct n) (auto simp: one_extreal_def) - -lemma zero_le_power_extreal[simp]: - fixes a :: extreal assumes "0 \ a" - shows "0 \ a ^ n" - using assms by (induct n) (auto simp: extreal_zero_le_0_iff) - -subsubsection {* Subtraction *} - -lemma extreal_minus_minus_image[simp]: - fixes S :: "extreal set" - shows "uminus ` uminus ` S = S" - by (auto simp: image_iff) - -lemma extreal_uminus_lessThan[simp]: - fixes a :: extreal shows "uminus ` {.. - extreal r = -\" - "extreal r - \ = -\" - "\ - x = \" - "-\ - \ = -\" - "x - -y = x + y" - "x - 0 = x" - "0 - x = -x" - by (simp_all add: minus_extreal_def) - -lemma extreal_x_minus_x[simp]: - "x - x = (if \x\ = \ then \ else 0)" - by (cases x) simp_all - -lemma extreal_eq_minus_iff: - fixes x y z :: extreal - shows "x = z - y \ - (\y\ \ \ \ x + y = z) \ - (y = -\ \ x = \) \ - (y = \ \ z = \ \ x = \) \ - (y = \ \ z \ \ \ x = -\)" - by (cases rule: extreal3_cases[of x y z]) auto - -lemma extreal_eq_minus: - fixes x y z :: extreal - shows "\y\ \ \ \ x = z - y \ x + y = z" - by (auto simp: extreal_eq_minus_iff) - -lemma extreal_less_minus_iff: - fixes x y z :: extreal - shows "x < z - y \ - (y = \ \ z = \ \ x \ \) \ - (y = -\ \ x \ \) \ - (\y\ \ \\ x + y < z)" - by (cases rule: extreal3_cases[of x y z]) auto - -lemma extreal_less_minus: - fixes x y z :: extreal - shows "\y\ \ \ \ x < z - y \ x + y < z" - by (auto simp: extreal_less_minus_iff) - -lemma extreal_le_minus_iff: - fixes x y z :: extreal - shows "x \ z - y \ - (y = \ \ z \ \ \ x = -\) \ - (\y\ \ \ \ x + y \ z)" - by (cases rule: extreal3_cases[of x y z]) auto - -lemma extreal_le_minus: - fixes x y z :: extreal - shows "\y\ \ \ \ x \ z - y \ x + y \ z" - by (auto simp: extreal_le_minus_iff) - -lemma extreal_minus_less_iff: - fixes x y z :: extreal - shows "x - y < z \ - y \ -\ \ (y = \ \ x \ \ \ z \ -\) \ - (y \ \ \ x < z + y)" - by (cases rule: extreal3_cases[of x y z]) auto - -lemma extreal_minus_less: - fixes x y z :: extreal - shows "\y\ \ \ \ x - y < z \ x < z + y" - by (auto simp: extreal_minus_less_iff) - -lemma extreal_minus_le_iff: - fixes x y z :: extreal - shows "x - y \ z \ - (y = -\ \ z = \) \ - (y = \ \ x = \ \ z = \) \ - (\y\ \ \ \ x \ z + y)" - by (cases rule: extreal3_cases[of x y z]) auto - -lemma extreal_minus_le: - fixes x y z :: extreal - shows "\y\ \ \ \ x - y \ z \ x \ z + y" - by (auto simp: extreal_minus_le_iff) - -lemma extreal_minus_eq_minus_iff: - fixes a b c :: extreal - shows "a - b = a - c \ - b = c \ a = \ \ (a = -\ \ b \ -\ \ c \ -\)" - by (cases rule: extreal3_cases[of a b c]) auto - -lemma extreal_add_le_add_iff: - "c + a \ c + b \ - a \ b \ c = \ \ (c = -\ \ a \ \ \ b \ \)" - by (cases rule: extreal3_cases[of a b c]) (simp_all add: field_simps) - -lemma extreal_mult_le_mult_iff: - "\c\ \ \ \ c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" - by (cases rule: extreal3_cases[of a b c]) (simp_all add: mult_le_cancel_left) - -lemma extreal_minus_mono: - fixes A B C D :: extreal assumes "A \ B" "D \ C" - shows "A - C \ B - D" - using assms - by (cases rule: extreal3_cases[case_product extreal_cases, of A B C D]) simp_all - -lemma real_of_extreal_minus: - "real (a - b) = (if \a\ = \ \ \b\ = \ then 0 else real a - real b)" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_diff_positive: - fixes a b :: extreal shows "a \ b \ 0 \ b - a" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_between: - fixes x e :: extreal - assumes "\x\ \ \" "0 < e" - shows "x - e < x" "x < x + e" -using assms apply (cases x, cases e) apply auto -using assms by (cases x, cases e) auto - -subsubsection {* Division *} - -instantiation extreal :: inverse -begin - -function inverse_extreal where -"inverse (extreal r) = (if r = 0 then \ else extreal (inverse r))" | -"inverse \ = 0" | -"inverse (-\) = 0" - by (auto intro: extreal_cases) -termination by (relation "{}") simp - -definition "x / y = x * inverse (y :: extreal)" - -instance proof qed -end - -lemma real_of_extreal_inverse[simp]: - fixes a :: extreal - shows "real (inverse a) = 1 / real a" - by (cases a) (auto simp: inverse_eq_divide) - -lemma extreal_inverse[simp]: - "inverse 0 = \" - "inverse (1::extreal) = 1" - by (simp_all add: one_extreal_def zero_extreal_def) - -lemma extreal_divide[simp]: - "extreal r / extreal p = (if p = 0 then extreal r * \ else extreal (r / p))" - unfolding divide_extreal_def by (auto simp: divide_real_def) - -lemma extreal_divide_same[simp]: - "x / x = (if \x\ = \ \ x = 0 then 0 else 1)" - by (cases x) - (simp_all add: divide_real_def divide_extreal_def one_extreal_def) - -lemma extreal_inv_inv[simp]: - "inverse (inverse x) = (if x \ -\ then x else \)" - by (cases x) auto - -lemma extreal_inverse_minus[simp]: - "inverse (- x) = (if x = 0 then \ else -inverse x)" - by (cases x) simp_all - -lemma extreal_uminus_divide[simp]: - fixes x y :: extreal shows "- x / y = - (x / y)" - unfolding divide_extreal_def by simp - -lemma extreal_divide_Infty[simp]: - "x / \ = 0" "x / -\ = 0" - unfolding divide_extreal_def by simp_all - -lemma extreal_divide_one[simp]: - "x / 1 = (x::extreal)" - unfolding divide_extreal_def by simp - -lemma extreal_divide_extreal[simp]: - "\ / extreal r = (if 0 \ r then \ else -\)" - unfolding divide_extreal_def by simp - -lemma zero_le_divide_extreal[simp]: - fixes a :: extreal assumes "0 \ a" "0 \ b" - shows "0 \ a / b" - using assms by (cases rule: extreal2_cases[of a b]) (auto simp: zero_le_divide_iff) - -lemma extreal_le_divide_pos: - "x > 0 \ x \ \ \ y \ z / x \ x * y \ z" - by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps) - -lemma extreal_divide_le_pos: - "x > 0 \ x \ \ \ z / x \ y \ z \ x * y" - by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps) - -lemma extreal_le_divide_neg: - "x < 0 \ x \ -\ \ y \ z / x \ z \ x * y" - by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps) - -lemma extreal_divide_le_neg: - "x < 0 \ x \ -\ \ z / x \ y \ x * y \ z" - by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps) - -lemma extreal_inverse_antimono_strict: - fixes x y :: extreal - shows "0 \ x \ x < y \ inverse y < inverse x" - by (cases rule: extreal2_cases[of x y]) auto - -lemma extreal_inverse_antimono: - fixes x y :: extreal - shows "0 \ x \ x <= y \ inverse y <= inverse x" - by (cases rule: extreal2_cases[of x y]) auto - -lemma inverse_inverse_Pinfty_iff[simp]: - "inverse x = \ \ x = 0" - by (cases x) auto - -lemma extreal_inverse_eq_0: - "inverse x = 0 \ x = \ \ x = -\" - by (cases x) auto - -lemma extreal_0_gt_inverse: - fixes x :: extreal shows "0 < inverse x \ x \ \ \ 0 \ x" - by (cases x) auto - -lemma extreal_mult_less_right: - assumes "b * a < c * a" "0 < a" "a < \" - shows "b < c" - using assms - by (cases rule: extreal3_cases[of a b c]) - (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff) - -lemma extreal_power_divide: - "y \ 0 \ (x / y :: extreal) ^ n = x^n / y^n" - by (cases rule: extreal2_cases[of x y]) - (auto simp: one_extreal_def zero_extreal_def power_divide not_le - power_less_zero_eq zero_le_power_iff) - -lemma extreal_le_mult_one_interval: - fixes x y :: extreal - assumes y: "y \ -\" - assumes z: "\z. \ 0 < z ; z < 1 \ \ z * x \ y" - shows "x \ y" -proof (cases x) - case PInf with z[of "1 / 2"] show "x \ y" by (simp add: one_extreal_def) -next - case (real r) note r = this - show "x \ y" - proof (cases y) - case (real p) note p = this - have "r \ p" - proof (rule field_le_mult_one_interval) - fix z :: real assume "0 < z" and "z < 1" - with z[of "extreal z"] - show "z * r \ p" using p r by (auto simp: zero_le_mult_iff one_extreal_def) - qed - then show "x \ y" using p r by simp - qed (insert y, simp_all) -qed simp - -subsection "Complete lattice" - -instantiation extreal :: lattice -begin -definition [simp]: "sup x y = (max x y :: extreal)" -definition [simp]: "inf x y = (min x y :: extreal)" -instance proof qed simp_all -end - -instantiation extreal :: complete_lattice -begin - -definition "bot = -\" -definition "top = \" - -definition "Sup S = (LEAST z. ALL x:S. x <= z :: extreal)" -definition "Inf S = (GREATEST z. ALL x:S. z <= x :: extreal)" - -lemma extreal_complete_Sup: - fixes S :: "extreal set" assumes "S \ {}" - shows "\x. (\y\S. y \ x) \ (\z. (\y\S. y \ z) \ x \ z)" -proof cases - assume "\x. \a\S. a \ extreal x" - then obtain y where y: "\a. a\S \ a \ extreal y" by auto - then have "\ \ S" by force - show ?thesis - proof cases - assume "S = {-\}" - then show ?thesis by (auto intro!: exI[of _ "-\"]) - next - assume "S \ {-\}" - with `S \ {}` `\ \ S` obtain x where "x \ S - {-\}" "x \ \" by auto - with y `\ \ S` have "\z\real ` (S - {-\}). z \ y" - by (auto simp: real_of_extreal_ord_simps) - with reals_complete2[of "real ` (S - {-\})"] `x \ S - {-\}` - obtain s where s: - "\y\S - {-\}. real y \ s" "\z. (\y\S - {-\}. real y \ z) \ s \ z" - by auto - show ?thesis - proof (safe intro!: exI[of _ "extreal s"]) - fix z assume "z \ S" with `\ \ S` show "z \ extreal s" - proof (cases z) - case (real r) - then show ?thesis - using s(1)[rule_format, of z] `z \ S` `z = extreal r` by auto - qed auto - next - fix z assume *: "\y\S. y \ z" - with `S \ {-\}` `S \ {}` show "extreal s \ z" - proof (cases z) - case (real u) - with * have "s \ u" - by (intro s(2)[of u]) (auto simp: real_of_extreal_ord_simps) - then show ?thesis using real by simp - qed auto - qed - qed -next - assume *: "\ (\x. \a\S. a \ extreal x)" - show ?thesis - proof (safe intro!: exI[of _ \]) - fix y assume **: "\z\S. z \ y" - with * show "\ \ y" - proof (cases y) - case MInf with * ** show ?thesis by (force simp: not_le) - qed auto - qed simp -qed - -lemma extreal_complete_Inf: - fixes S :: "extreal set" assumes "S ~= {}" - shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)" -proof- -def S1 == "uminus ` S" -hence "S1 ~= {}" using assms by auto -from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)" - using extreal_complete_Sup[of S1] by auto -{ fix z assume "ALL y:S. z <= y" - hence "ALL y:S1. y <= -z" unfolding S1_def by auto - hence "x <= -z" using x_def by auto - hence "z <= -x" - apply (subst extreal_uminus_uminus[symmetric]) - unfolding extreal_minus_le_minus . } -moreover have "(ALL y:S. -x <= y)" - using x_def unfolding S1_def - apply simp - apply (subst (3) extreal_uminus_uminus[symmetric]) - unfolding extreal_minus_le_minus by simp -ultimately show ?thesis by auto -qed - -lemma extreal_complete_uminus_eq: - fixes S :: "extreal set" - shows "(\y\uminus`S. y \ x) \ (\z. (\y\uminus`S. y \ z) \ x \ z) - \ (\y\S. -x \ y) \ (\z. (\y\S. z \ y) \ z \ -x)" - by simp (metis extreal_minus_le_minus extreal_uminus_uminus) - -lemma extreal_Sup_uminus_image_eq: - fixes S :: "extreal set" - shows "Sup (uminus ` S) = - Inf S" -proof cases - assume "S = {}" - moreover have "(THE x. All (op \ x)) = (-\::extreal)" - by (rule the_equality) (auto intro!: extreal_bot) - moreover have "(SOME x. \y. y \ x) = (\::extreal)" - by (rule some_equality) (auto intro!: extreal_top) - ultimately show ?thesis unfolding Inf_extreal_def Sup_extreal_def - Least_def Greatest_def GreatestM_def by simp -next - assume "S \ {}" - with extreal_complete_Sup[of "uminus`S"] - obtain x where x: "(\y\S. -x \ y) \ (\z. (\y\S. z \ y) \ z \ -x)" - unfolding extreal_complete_uminus_eq by auto - show "Sup (uminus ` S) = - Inf S" - unfolding Inf_extreal_def Greatest_def GreatestM_def - proof (intro someI2[of _ _ "\x. Sup (uminus`S) = - x"]) - show "(\y\S. -x \ y) \ (\y. (\z\S. y \ z) \ y \ -x)" - using x . - fix x' assume "(\y\S. x' \ y) \ (\y. (\z\S. y \ z) \ y \ x')" - then have "(\y\uminus`S. y \ - x') \ (\y. (\z\uminus`S. z \ y) \ - x' \ y)" - unfolding extreal_complete_uminus_eq by simp - then show "Sup (uminus ` S) = -x'" - unfolding Sup_extreal_def extreal_uminus_eq_iff - by (intro Least_equality) auto - qed -qed - -instance -proof - { fix x :: extreal and A - show "bot <= x" by (cases x) (simp_all add: bot_extreal_def) - show "x <= top" by (simp add: top_extreal_def) } - - { fix x :: extreal and A assume "x : A" - with extreal_complete_Sup[of A] - obtain s where s: "\y\A. y <= s" "\z. (\y\A. y <= z) \ s <= z" by auto - hence "x <= s" using `x : A` by auto - also have "... = Sup A" using s unfolding Sup_extreal_def - by (auto intro!: Least_equality[symmetric]) - finally show "x <= Sup A" . } - note le_Sup = this - - { fix x :: extreal and A assume *: "!!z. (z : A ==> z <= x)" - show "Sup A <= x" - proof (cases "A = {}") - case True - hence "Sup A = -\" unfolding Sup_extreal_def - by (auto intro!: Least_equality) - thus "Sup A <= x" by simp - next - case False - with extreal_complete_Sup[of A] - obtain s where s: "\y\A. y <= s" "\z. (\y\A. y <= z) \ s <= z" by auto - hence "Sup A = s" - unfolding Sup_extreal_def by (auto intro!: Least_equality) - also have "s <= x" using * s by auto - finally show "Sup A <= x" . - qed } - note Sup_le = this - - { fix x :: extreal and A assume "x \ A" - with le_Sup[of "-x" "uminus`A"] show "Inf A \ x" - unfolding extreal_Sup_uminus_image_eq by simp } - - { fix x :: extreal and A assume *: "!!z. (z : A ==> x <= z)" - with Sup_le[of "uminus`A" "-x"] show "x \ Inf A" - unfolding extreal_Sup_uminus_image_eq by force } -qed -end - -lemma extreal_SUPR_uminus: - fixes f :: "'a => extreal" - shows "(SUP i : R. -(f i)) = -(INF i : R. f i)" - unfolding SUPR_def INFI_def - using extreal_Sup_uminus_image_eq[of "f`R"] - by (simp add: image_image) - -lemma extreal_INFI_uminus: - fixes f :: "'a => extreal" - shows "(INF i : R. -(f i)) = -(SUP i : R. f i)" - using extreal_SUPR_uminus[of _ "\x. - f x"] by simp - -lemma extreal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::extreal set)" - using extreal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image) - -lemma extreal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: extreal set)" - by (auto intro!: inj_onI) - -lemma extreal_image_uminus_shift: - fixes X Y :: "extreal set" shows "uminus ` X = Y \ X = uminus ` Y" -proof - assume "uminus ` X = Y" - then have "uminus ` uminus ` X = uminus ` Y" - by (simp add: inj_image_eq_iff) - then show "X = uminus ` Y" by (simp add: image_image) -qed (simp add: image_image) - -lemma Inf_extreal_iff: - fixes z :: extreal - shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x Inf X < y" - by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear - order_less_le_trans) - -lemma Sup_eq_MInfty: - fixes S :: "extreal set" shows "Sup S = -\ \ S = {} \ S = {-\}" -proof - assume a: "Sup S = -\" - with complete_lattice_class.Sup_upper[of _ S] - show "S={} \ S={-\}" by auto -next - assume "S={} \ S={-\}" then show "Sup S = -\" - unfolding Sup_extreal_def by (auto intro!: Least_equality) -qed - -lemma Inf_eq_PInfty: - fixes S :: "extreal set" shows "Inf S = \ \ S = {} \ S = {\}" - using Sup_eq_MInfty[of "uminus`S"] - unfolding extreal_Sup_uminus_image_eq extreal_image_uminus_shift by simp - -lemma Inf_eq_MInfty: "-\ : S ==> Inf S = -\" - unfolding Inf_extreal_def - by (auto intro!: Greatest_equality) - -lemma Sup_eq_PInfty: "\ : S ==> Sup S = \" - unfolding Sup_extreal_def - by (auto intro!: Least_equality) - -lemma extreal_SUPI: - fixes x :: extreal - assumes "!!i. i : A ==> f i <= x" - assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y" - shows "(SUP i:A. f i) = x" - unfolding SUPR_def Sup_extreal_def - using assms by (auto intro!: Least_equality) - -lemma extreal_INFI: - fixes x :: extreal - assumes "!!i. i : A ==> f i >= x" - assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y" - shows "(INF i:A. f i) = x" - unfolding INFI_def Inf_extreal_def - using assms by (auto intro!: Greatest_equality) - -lemma Sup_extreal_close: - fixes e :: extreal - assumes "0 < e" and S: "\Sup S\ \ \" "S \ {}" - shows "\x\S. Sup S - e < x" - using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1]) - -lemma Inf_extreal_close: - fixes e :: extreal assumes "\Inf X\ \ \" "0 < e" - shows "\x\X. x < Inf X + e" -proof (rule Inf_less_iff[THEN iffD1]) - show "Inf X < Inf X + e" using assms - by (cases e) auto -qed - -lemma Sup_eq_top_iff: - fixes A :: "'a::{complete_lattice, linorder} set" - shows "Sup A = top \ (\xi\A. x < i)" -proof - assume *: "Sup A = top" - show "(\xi\A. x < i)" unfolding *[symmetric] - proof (intro allI impI) - fix x assume "x < Sup A" then show "\i\A. x < i" - unfolding less_Sup_iff by auto - qed -next - assume *: "\xi\A. x < i" - show "Sup A = top" - proof (rule ccontr) - assume "Sup A \ top" - with top_greatest[of "Sup A"] - have "Sup A < top" unfolding le_less by auto - then have "Sup A < Sup A" - using * unfolding less_Sup_iff by auto - then show False by auto - qed -qed - -lemma SUP_eq_top_iff: - fixes f :: "'a \ 'b::{complete_lattice, linorder}" - shows "(SUP i:A. f i) = top \ (\xi\A. x < f i)" - unfolding SUPR_def Sup_eq_top_iff by auto - -lemma SUP_nat_Infty: "(SUP i::nat. extreal (real i)) = \" -proof - - { fix x assume "x \ \" - then have "\k::nat. x < extreal (real k)" - proof (cases x) - case MInf then show ?thesis by (intro exI[of _ 0]) auto - next - case (real r) - moreover obtain k :: nat where "r < real k" - using ex_less_of_nat by (auto simp: real_eq_of_nat) - ultimately show ?thesis by auto - qed simp } - then show ?thesis - using SUP_eq_top_iff[of UNIV "\n::nat. extreal (real n)"] - by (auto simp: top_extreal_def) -qed - -lemma extreal_le_Sup: - fixes x :: extreal - shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))" -(is "?lhs <-> ?rhs") -proof- -{ assume "?rhs" - { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i) (ALL y. x < y --> (EX i. i : A & f i <= y))" -(is "?lhs <-> ?rhs") -proof- -{ assume "?rhs" - { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le) - from this obtain y where y_def: "x x" by auto - hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto - thus False using assms by auto -qed - -lemma same_INF: - assumes "ALL e:A. f e = g e" - shows "(INF e:A. f e) = (INF e:A. g e)" -proof- -have "f ` A = g ` A" unfolding image_def using assms by auto -thus ?thesis unfolding INFI_def by auto -qed - -lemma same_SUP: - assumes "ALL e:A. f e = g e" - shows "(SUP e:A. f e) = (SUP e:A. g e)" -proof- -have "f ` A = g ` A" unfolding image_def using assms by auto -thus ?thesis unfolding SUPR_def by auto -qed - -lemma SUPR_eq: - assumes "\i\A. \j\B. f i \ g j" - assumes "\j\B. \i\A. g j \ f i" - shows "(SUP i:A. f i) = (SUP j:B. g j)" -proof (intro antisym) - show "(SUP i:A. f i) \ (SUP j:B. g j)" - using assms by (metis SUP_leI le_SUPI2) - show "(SUP i:B. g i) \ (SUP j:A. f j)" - using assms by (metis SUP_leI le_SUPI2) -qed - -lemma SUP_extreal_le_addI: - assumes "\i. f i + y \ z" and "y \ -\" - shows "SUPR UNIV f + y \ z" -proof (cases y) - case (real r) - then have "\i. f i \ z - y" using assms by (simp add: extreal_le_minus_iff) - then have "SUPR UNIV f \ z - y" by (rule SUP_leI) - then show ?thesis using real by (simp add: extreal_le_minus_iff) -qed (insert assms, auto) - -lemma SUPR_extreal_add: - fixes f g :: "nat \ extreal" - assumes "incseq f" "incseq g" and pos: "\i. f i \ -\" "\i. g i \ -\" - shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g" -proof (rule extreal_SUPI) - fix y assume *: "\i. i \ UNIV \ f i + g i \ y" - have f: "SUPR UNIV f \ -\" using pos - unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD) - { fix j - { fix i - have "f i + g j \ f i + g (max i j)" - using `incseq g`[THEN incseqD] by (rule add_left_mono) auto - also have "\ \ f (max i j) + g (max i j)" - using `incseq f`[THEN incseqD] by (rule add_right_mono) auto - also have "\ \ y" using * by auto - finally have "f i + g j \ y" . } - then have "SUPR UNIV f + g j \ y" - using assms(4)[of j] by (intro SUP_extreal_le_addI) auto - then have "g j + SUPR UNIV f \ y" by (simp add: ac_simps) } - then have "SUPR UNIV g + SUPR UNIV f \ y" - using f by (rule SUP_extreal_le_addI) - then show "SUPR UNIV f + SUPR UNIV g \ y" by (simp add: ac_simps) -qed (auto intro!: add_mono le_SUPI) - -lemma SUPR_extreal_add_pos: - fixes f g :: "nat \ extreal" - assumes inc: "incseq f" "incseq g" and pos: "\i. 0 \ f i" "\i. 0 \ g i" - shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g" -proof (intro SUPR_extreal_add inc) - fix i show "f i \ -\" "g i \ -\" using pos[of i] by auto -qed - -lemma SUPR_extreal_setsum: - fixes f g :: "'a \ nat \ extreal" - assumes "\n. n \ A \ incseq (f n)" and pos: "\n i. n \ A \ 0 \ f n i" - shows "(SUP i. \n\A. f n i) = (\n\A. SUPR UNIV (f n))" -proof cases - assume "finite A" then show ?thesis using assms - by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_extreal_add_pos) -qed simp - -lemma SUPR_extreal_cmult: - fixes f :: "nat \ extreal" assumes "\i. 0 \ f i" "0 \ c" - shows "(SUP i. c * f i) = c * SUPR UNIV f" -proof (rule extreal_SUPI) - fix i have "f i \ SUPR UNIV f" by (rule le_SUPI) auto - then show "c * f i \ c * SUPR UNIV f" - using `0 \ c` by (rule extreal_mult_left_mono) -next - fix y assume *: "\i. i \ UNIV \ c * f i \ y" - show "c * SUPR UNIV f \ y" - proof cases - assume c: "0 < c \ c \ \" - with * have "SUPR UNIV f \ y / c" - by (intro SUP_leI) (auto simp: extreal_le_divide_pos) - with c show ?thesis - by (auto simp: extreal_le_divide_pos) - next - { assume "c = \" have ?thesis - proof cases - assume "\i. f i = 0" - moreover then have "range f = {0}" by auto - ultimately show "c * SUPR UNIV f \ y" using * by (auto simp: SUPR_def) - next - assume "\ (\i. f i = 0)" - then obtain i where "f i \ 0" by auto - with *[of i] `c = \` `0 \ f i` show ?thesis by (auto split: split_if_asm) - qed } - moreover assume "\ (0 < c \ c \ \)" - ultimately show ?thesis using * `0 \ c` by auto - qed -qed - -lemma SUP_PInfty: - fixes f :: "'a \ extreal" - assumes "\n::nat. \i\A. extreal (real n) \ f i" - shows "(SUP i:A. f i) = \" - unfolding SUPR_def Sup_eq_top_iff[where 'a=extreal, unfolded top_extreal_def] - apply simp -proof safe - fix x assume "x \ \" - show "\i\A. x < f i" - proof (cases x) - case PInf with `x \ \` show ?thesis by simp - next - case MInf with assms[of "0"] show ?thesis by force - next - case (real r) - with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < extreal (real n)" by auto - moreover from assms[of n] guess i .. - ultimately show ?thesis - by (auto intro!: bexI[of _ i]) - qed -qed - -lemma Sup_countable_SUPR: - assumes "A \ {}" - shows "\f::nat \ extreal. range f \ A \ Sup A = SUPR UNIV f" -proof (cases "Sup A") - case (real r) - have "\n::nat. \x. x \ A \ Sup A < x + 1 / extreal (real n)" - proof - fix n ::nat have "\x\A. Sup A - 1 / extreal (real n) < x" - using assms real by (intro Sup_extreal_close) (auto simp: one_extreal_def) - then guess x .. - then show "\x. x \ A \ Sup A < x + 1 / extreal (real n)" - by (auto intro!: exI[of _ x] simp: extreal_minus_less_iff) - qed - from choice[OF this] guess f .. note f = this - have "SUPR UNIV f = Sup A" - proof (rule extreal_SUPI) - fix i show "f i \ Sup A" using f - by (auto intro!: complete_lattice_class.Sup_upper) - next - fix y assume bound: "\i. i \ UNIV \ f i \ y" - show "Sup A \ y" - proof (rule extreal_le_epsilon, intro allI impI) - fix e :: extreal assume "0 < e" - show "Sup A \ y + e" - proof (cases e) - case (real r) - hence "0 < r" using `0 < e` by auto - then obtain n ::nat where *: "1 / real n < r" "0 < n" - using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide) - have "Sup A \ f n + 1 / extreal (real n)" using f[THEN spec, of n] by auto - also have "1 / extreal (real n) \ e" using real * by (auto simp: one_extreal_def ) - with bound have "f n + 1 / extreal (real n) \ y + e" by (rule add_mono) simp - finally show "Sup A \ y + e" . - qed (insert `0 < e`, auto) - qed - qed - with f show ?thesis by (auto intro!: exI[of _ f]) -next - case PInf - from `A \ {}` obtain x where "x \ A" by auto - show ?thesis - proof cases - assume "\ \ A" - moreover then have "\ \ Sup A" by (intro complete_lattice_class.Sup_upper) - ultimately show ?thesis by (auto intro!: exI[of _ "\x. \"]) - next - assume "\ \ A" - have "\x\A. 0 \ x" - by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least extreal_infty_less_eq2 linorder_linear) - then obtain x where "x \ A" "0 \ x" by auto - have "\n::nat. \f. f \ A \ x + extreal (real n) \ f" - proof (rule ccontr) - assume "\ ?thesis" - then have "\n::nat. Sup A \ x + extreal (real n)" - by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le) - then show False using `x \ A` `\ \ A` PInf - by(cases x) auto - qed - from choice[OF this] guess f .. note f = this - have "SUPR UNIV f = \" - proof (rule SUP_PInfty) - fix n :: nat show "\i\UNIV. extreal (real n) \ f i" - using f[THEN spec, of n] `0 \ x` - by (cases rule: extreal2_cases[of "f n" x]) (auto intro!: exI[of _ n]) - qed - then show ?thesis using f PInf by (auto intro!: exI[of _ f]) - qed -next - case MInf - with `A \ {}` have "A = {-\}" by (auto simp: Sup_eq_MInfty) - then show ?thesis using MInf by (auto intro!: exI[of _ "\x. -\"]) -qed - -lemma SUPR_countable_SUPR: - "A \ {} \ \f::nat \ extreal. range f \ g`A \ SUPR A g = SUPR UNIV f" - using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def) - - -lemma Sup_extreal_cadd: - fixes A :: "extreal set" assumes "A \ {}" and "a \ -\" - shows "Sup ((\x. a + x) ` A) = a + Sup A" -proof (rule antisym) - have *: "\a::extreal. \A. Sup ((\x. a + x) ` A) \ a + Sup A" - by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper) - then show "Sup ((\x. a + x) ` A) \ a + Sup A" . - show "a + Sup A \ Sup ((\x. a + x) ` A)" - proof (cases a) - case PInf with `A \ {}` show ?thesis by (auto simp: image_constant) - next - case (real r) - then have **: "op + (- a) ` op + a ` A = A" - by (auto simp: image_iff ac_simps zero_extreal_def[symmetric]) - from *[of "-a" "(\x. a + x) ` A"] real show ?thesis unfolding ** - by (cases rule: extreal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto - qed (insert `a \ -\`, auto) -qed - -lemma Sup_extreal_cminus: - fixes A :: "extreal set" assumes "A \ {}" and "a \ -\" - shows "Sup ((\x. a - x) ` A) = a - Inf A" - using Sup_extreal_cadd[of "uminus ` A" a] assms - by (simp add: comp_def image_image minus_extreal_def - extreal_Sup_uminus_image_eq) - -lemma SUPR_extreal_cminus: - fixes A assumes "A \ {}" and "a \ -\" - shows "(SUP x:A. a - f x) = a - (INF x:A. f x)" - using Sup_extreal_cminus[of "f`A" a] assms - unfolding SUPR_def INFI_def image_image by auto - -lemma Inf_extreal_cminus: - fixes A :: "extreal set" assumes "A \ {}" and "\a\ \ \" - shows "Inf ((\x. a - x) ` A) = a - Sup A" -proof - - { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto } - moreover then have "(\x. -a - x)`uminus`A = uminus ` (\x. a - x) ` A" - by (auto simp: image_image) - ultimately show ?thesis - using Sup_extreal_cminus[of "uminus ` A" "-a"] assms - by (auto simp add: extreal_Sup_uminus_image_eq extreal_Inf_uminus_image_eq) -qed - -lemma INFI_extreal_cminus: - fixes A assumes "A \ {}" and "\a\ \ \" - shows "(INF x:A. a - f x) = a - (SUP x:A. f x)" - using Inf_extreal_cminus[of "f`A" a] assms - unfolding SUPR_def INFI_def image_image - by auto - -lemma uminus_extreal_add_uminus_uminus: - fixes a b :: extreal shows "a \ \ \ b \ \ \ - (- a + - b) = a + b" - by (cases rule: extreal2_cases[of a b]) auto - -lemma INFI_extreal_add: - assumes "decseq f" "decseq g" and fin: "\i. f i \ \" "\i. g i \ \" - shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g" -proof - - have INF_less: "(INF i. f i) < \" "(INF i. g i) < \" - using assms unfolding INF_less_iff by auto - { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i" - by (rule uminus_extreal_add_uminus_uminus) } - then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))" - by simp - also have "\ = INFI UNIV f + INFI UNIV g" - unfolding extreal_INFI_uminus - using assms INF_less - by (subst SUPR_extreal_add) - (auto simp: extreal_SUPR_uminus intro!: uminus_extreal_add_uminus_uminus) - finally show ?thesis . -qed - -subsection "Limits on @{typ extreal}" - -subsubsection "Topological space" - -instantiation extreal :: topological_space -begin - -definition "open A \ open (extreal -` A) - \ (\ \ A \ (\x. {extreal x <..} \ A)) - \ (-\ \ A \ (\x. {.. A))" - -lemma open_PInfty: "open A \ \ \ A \ (\x. {extreal x<..} \ A)" - unfolding open_extreal_def by auto - -lemma open_MInfty: "open A \ -\ \ A \ (\x. {.. A)" - unfolding open_extreal_def by auto - -lemma open_PInfty2: assumes "open A" "\ \ A" obtains x where "{extreal x<..} \ A" - using open_PInfty[OF assms] by auto - -lemma open_MInfty2: assumes "open A" "-\ \ A" obtains x where "{.. A" - using open_MInfty[OF assms] by auto - -lemma extreal_openE: assumes "open A" obtains x y where - "open (extreal -` A)" - "\ \ A \ {extreal x<..} \ A" - "-\ \ A \ {.. A" - using assms open_extreal_def by auto - -instance -proof - let ?U = "UNIV::extreal set" - show "open ?U" unfolding open_extreal_def - by (auto intro!: exI[of _ 0]) -next - fix S T::"extreal set" assume "open S" and "open T" - from `open S`[THEN extreal_openE] guess xS yS . - moreover from `open T`[THEN extreal_openE] guess xT yT . - ultimately have - "open (extreal -` (S \ T))" - "\ \ S \ T \ {extreal (max xS xT) <..} \ S \ T" - "-\ \ S \ T \ {..< extreal (min yS yT)} \ S \ T" - by auto - then show "open (S Int T)" unfolding open_extreal_def by blast -next - fix K :: "extreal set set" assume "\S\K. open S" - then have *: "\S. \x y. S \ K \ open (extreal -` S) \ - (\ \ S \ {extreal x <..} \ S) \ (-\ \ S \ {..< extreal y} \ S)" - by (auto simp: open_extreal_def) - then show "open (Union K)" unfolding open_extreal_def - proof (intro conjI impI) - show "open (extreal -` \K)" - using *[THEN choice] by (auto simp: vimage_Union) - qed ((metis UnionE Union_upper subset_trans *)+) -qed -end - -lemma open_extreal: "open S \ open (extreal ` S)" - by (auto simp: inj_vimage_image_eq open_extreal_def) - -lemma open_extreal_vimage: "open S \ open (extreal -` S)" - unfolding open_extreal_def by auto - -lemma open_extreal_lessThan[intro, simp]: "open {..< a :: extreal}" -proof - - have "\x. extreal -` {..} = UNIV" "extreal -` {..< -\} = {}" by auto - then show ?thesis by (cases a) (auto simp: open_extreal_def) -qed - -lemma open_extreal_greaterThan[intro, simp]: - "open {a :: extreal <..}" -proof - - have "\x. extreal -` {extreal x<..} = {x<..}" - "extreal -` {\<..} = {}" "extreal -` {-\<..} = UNIV" by auto - then show ?thesis by (cases a) (auto simp: open_extreal_def) -qed - -lemma extreal_open_greaterThanLessThan[intro, simp]: "open {a::extreal <..< b}" - unfolding greaterThanLessThan_def by auto - -lemma closed_extreal_atLeast[simp, intro]: "closed {a :: extreal ..}" -proof - - have "- {a ..} = {..< a}" by auto - then show "closed {a ..}" - unfolding closed_def using open_extreal_lessThan by auto -qed - -lemma closed_extreal_atMost[simp, intro]: "closed {.. b :: extreal}" -proof - - have "- {.. b} = {b <..}" by auto - then show "closed {.. b}" - unfolding closed_def using open_extreal_greaterThan by auto -qed - -lemma closed_extreal_atLeastAtMost[simp, intro]: - shows "closed {a :: extreal .. b}" - unfolding atLeastAtMost_def by auto - -lemma closed_extreal_singleton: - "closed {a :: extreal}" -by (metis atLeastAtMost_singleton closed_extreal_atLeastAtMost) - -lemma extreal_open_cont_interval: - assumes "open S" "x \ S" "\x\ \ \" - obtains e where "e>0" "{x-e <..< x+e} \ S" -proof- - from `open S` have "open (extreal -` S)" by (rule extreal_openE) - then obtain e where "0 < e" and e: "\y. dist y (real x) < e \ extreal y \ S" - using assms unfolding open_dist by force - show thesis - proof (intro that subsetI) - show "0 < extreal e" using `0 < e` by auto - fix y assume "y \ {x - extreal e<.. S" using e[of t] by auto - qed -qed - -lemma extreal_open_cont_interval2: - assumes "open S" "x \ S" and x: "\x\ \ \" - obtains a b where "a < x" "x < b" "{a <..< b} \ S" -proof- - guess e using extreal_open_cont_interval[OF assms] . - with that[of "x-e" "x+e"] extreal_between[OF x, of e] - show thesis by auto -qed - -instance extreal :: t2_space -proof - fix x y :: extreal assume "x ~= y" - let "?P x (y::extreal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}" - - { fix x y :: extreal assume "x < y" - from extreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto - have "?P x y" - apply (rule exI[of _ "{..n. extreal (f n)) ---> extreal x) net \ (f ---> x) net" (is "?l = ?r") -proof (intro iffI topological_tendstoI) - fix S assume "?l" "open S" "x \ S" - then show "eventually (\x. f x \ S) net" - using `?l`[THEN topological_tendstoD, OF open_extreal, OF `open S`] - by (simp add: inj_image_mem_iff) -next - fix S assume "?r" "open S" "extreal x \ S" - show "eventually (\x. extreal (f x) \ S) net" - using `?r`[THEN topological_tendstoD, OF open_extreal_vimage, OF `open S`] - using `extreal x \ S` by auto -qed - -lemma lim_real_of_extreal[simp]: - assumes lim: "(f ---> extreal x) net" - shows "((\x. real (f x)) ---> x) net" -proof (intro topological_tendstoI) - fix S assume "open S" "x \ S" - then have S: "open S" "extreal x \ extreal ` S" - by (simp_all add: inj_image_mem_iff) - have "\x. f x \ extreal ` S \ real (f x) \ S" by auto - from this lim[THEN topological_tendstoD, OF open_extreal, OF S] - show "eventually (\x. real (f x) \ S) net" - by (rule eventually_mono) -qed - -lemma Lim_PInfty: "f ----> \ <-> (ALL B. EX N. ALL n>=N. f n >= extreal B)" (is "?l = ?r") -proof assume ?r show ?l apply(rule topological_tendstoI) - unfolding eventually_sequentially - proof- fix S assume "open S" "\ : S" - from open_PInfty[OF this] guess B .. note B=this - from `?r`[rule_format,of "B+1"] guess N .. note N=this - show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI) - proof safe case goal1 - have "extreal B < extreal (B + 1)" by auto - also have "... <= f n" using goal1 N by auto - finally show ?case using B by fastsimp - qed - qed -next assume ?l show ?r - proof fix B::real have "open {extreal B<..}" "\ : {extreal B<..}" by auto - from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially] - guess N .. note N=this - show "EX N. ALL n>=N. extreal B <= f n" apply(rule_tac x=N in exI) using N by auto - qed -qed - - -lemma Lim_MInfty: "f ----> (-\) <-> (ALL B. EX N. ALL n>=N. f n <= extreal B)" (is "?l = ?r") -proof assume ?r show ?l apply(rule topological_tendstoI) - unfolding eventually_sequentially - proof- fix S assume "open S" "(-\) : S" - from open_MInfty[OF this] guess B .. note B=this - from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this - show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI) - proof safe case goal1 - have "extreal (B - 1) >= f n" using goal1 N by auto - also have "... < extreal B" by auto - finally show ?case using B by fastsimp - qed - qed -next assume ?l show ?r - proof fix B::real have "open {..) : {..=N. extreal B >= f n" apply(rule_tac x=N in exI) using N by auto - qed -qed - - -lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= extreal B" shows "l ~= \" -proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\" - from lim[unfolded this Lim_PInfty,rule_format,of "?B"] - guess N .. note N=this[rule_format,OF le_refl] - hence "extreal ?B <= extreal B" using assms(2)[of N] by(rule order_trans) - hence "extreal ?B < extreal ?B" apply (rule le_less_trans) by auto - thus False by auto -qed - - -lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= extreal B" shows "l ~= (-\)" -proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\)" - from lim[unfolded this Lim_MInfty,rule_format,of "?B"] - guess N .. note N=this[rule_format,OF le_refl] - hence "extreal B <= extreal ?B" using assms(2)[of N] order_trans[of "extreal B" "f N" "extreal(B - 1)"] by blast - thus False by auto -qed - - -lemma tendsto_explicit: - "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))" - unfolding tendsto_def eventually_sequentially by auto - - -lemma tendsto_obtains_N: - assumes "f ----> f0" - assumes "open S" "f0 : S" - obtains N where "ALL n>=N. f n : S" - using tendsto_explicit[of f f0] assms by auto - - -lemma tail_same_limit: - fixes X Y N - assumes "X ----> L" "ALL n>=N. X n = Y n" - shows "Y ----> L" -proof- -{ fix S assume "open S" and "L:S" - from this obtain N1 where "ALL n>=N1. X n : S" - using assms unfolding tendsto_def eventually_sequentially by auto - hence "ALL n>=max N N1. Y n : S" using assms by auto - hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto -} -thus ?thesis using tendsto_explicit by auto -qed - - -lemma Lim_bounded_PInfty2: -assumes lim:"f ----> l" and "ALL n>=N. f n <= extreal B" -shows "l ~= \" -proof- - def g == "(%n. if n>=N then f n else extreal B)" - hence "g ----> l" using tail_same_limit[of f l N g] lim by auto - moreover have "!!n. g n <= extreal B" using g_def assms by auto - ultimately show ?thesis using Lim_bounded_PInfty by auto -qed - -lemma Lim_bounded_extreal: - assumes lim:"f ----> (l :: extreal)" - and "ALL n>=M. f n <= C" - shows "l<=C" -proof- -{ assume "l=(-\)" hence ?thesis by auto } -moreover -{ assume "~(l=(-\))" - { assume "C=\" hence ?thesis by auto } - moreover - { assume "C=(-\)" hence "ALL n>=M. f n = (-\)" using assms by auto - hence "l=(-\)" using assms - tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\n. -\" "-\" M f, OF tendsto_const] by auto - hence ?thesis by auto } - moreover - { assume "EX B. C = extreal B" - from this obtain B where B_def: "C=extreal B" by auto - hence "~(l=\)" using Lim_bounded_PInfty2 assms by auto - from this obtain m where m_def: "extreal m=l" using `~(l=(-\))` by (cases l) auto - from this obtain N where N_def: "ALL n>=N. f n : {extreal(m - 1) <..< extreal(m+1)}" - apply (subst tendsto_obtains_N[of f l "{extreal(m - 1) <..< extreal(m+1)}"]) using assms by auto - { fix n assume "n>=N" - hence "EX r. extreal r = f n" using N_def by (cases "f n") auto - } from this obtain g where g_def: "ALL n>=N. extreal (g n) = f n" by metis - hence "(%n. extreal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto - hence *: "(%n. g n) ----> m" using m_def by auto - { fix n assume "n>=max N M" - hence "extreal (g n) <= extreal B" using assms g_def B_def by auto - hence "g n <= B" by auto - } hence "EX N. ALL n>=N. g n <= B" by blast - hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto - hence ?thesis using m_def B_def by auto - } ultimately have ?thesis by (cases C) auto -} ultimately show ?thesis by blast -qed - -lemma real_of_extreal_mult[simp]: - fixes a b :: extreal shows "real (a * b) = real a * real b" - by (cases rule: extreal2_cases[of a b]) auto - -lemma real_of_extreal_eq_0: - "real x = 0 \ x = \ \ x = -\ \ x = 0" - by (cases x) auto - -lemma tendsto_extreal_realD: - fixes f :: "'a \ extreal" - assumes "x \ 0" and tendsto: "((\x. extreal (real (f x))) ---> x) net" - shows "(f ---> x) net" -proof (intro topological_tendstoI) - fix S assume S: "open S" "x \ S" - with `x \ 0` have "open (S - {0})" "x \ S - {0}" by auto - from tendsto[THEN topological_tendstoD, OF this] - show "eventually (\x. f x \ S) net" - by (rule eventually_rev_mp) (auto simp: extreal_real real_of_extreal_0) -qed - -lemma tendsto_extreal_realI: - fixes f :: "'a \ extreal" - assumes x: "\x\ \ \" and tendsto: "(f ---> x) net" - shows "((\x. extreal (real (f x))) ---> x) net" -proof (intro topological_tendstoI) - fix S assume "open S" "x \ S" - with x have "open (S - {\, -\})" "x \ S - {\, -\}" by auto - from tendsto[THEN topological_tendstoD, OF this] - show "eventually (\x. extreal (real (f x)) \ S) net" - by (elim eventually_elim1) (auto simp: extreal_real) -qed - -lemma extreal_mult_cancel_left: - fixes a b c :: extreal shows "a * b = a * c \ - ((\a\ = \ \ 0 < b * c) \ a = 0 \ b = c)" - by (cases rule: extreal3_cases[of a b c]) - (simp_all add: zero_less_mult_iff) - -lemma extreal_inj_affinity: - assumes "\m\ \ \" "m \ 0" "\t\ \ \" - shows "inj_on (\x. m * x + t) A" - using assms - by (cases rule: extreal2_cases[of m t]) - (auto intro!: inj_onI simp: extreal_add_cancel_right extreal_mult_cancel_left) - -lemma extreal_PInfty_eq_plus[simp]: - shows "\ = a + b \ a = \ \ b = \" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_MInfty_eq_plus[simp]: - shows "-\ = a + b \ (a = -\ \ b \ \) \ (b = -\ \ a \ \)" - by (cases rule: extreal2_cases[of a b]) auto - -lemma extreal_less_divide_pos: - "x > 0 \ x \ \ \ y < z / x \ x * y < z" - by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps) - -lemma extreal_divide_less_pos: - "x > 0 \ x \ \ \ y / x < z \ y < x * z" - by (cases rule: extreal3_cases[of x y z]) (auto simp: field_simps) - -lemma extreal_divide_eq: - "b \ 0 \ \b\ \ \ \ a / b = c \ a = b * c" - by (cases rule: extreal3_cases[of a b c]) - (simp_all add: field_simps) - -lemma extreal_inverse_not_MInfty[simp]: "inverse a \ -\" - by (cases a) auto - -lemma extreal_mult_m1[simp]: "x * extreal (-1) = -x" - by (cases x) auto - -lemma extreal_LimI_finite: - assumes "\x\ \ \" - assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r" - shows "u ----> x" -proof (rule topological_tendstoI, unfold eventually_sequentially) - obtain rx where rx_def: "x=extreal rx" using assms by (cases x) auto - fix S assume "open S" "x : S" - then have "open (extreal -` S)" unfolding open_extreal_def by auto - with `x \ S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> extreal y \ S" - unfolding open_real_def rx_def by auto - then obtain n where - upper: "!!N. n <= N ==> u N < x + extreal r" and - lower: "!!N. n <= N ==> x < u N + extreal r" using assms(2)[of "extreal r"] by auto - show "EX N. ALL n>=N. u n : S" - proof (safe intro!: exI[of _ n]) - fix N assume "n <= N" - from upper[OF this] lower[OF this] assms `0 < r` - have "u N ~: {\,(-\)}" by auto - from this obtain ra where ra_def: "(u N) = extreal ra" by (cases "u N") auto - hence "rx < ra + r" and "ra < rx + r" - using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto - hence "dist (real (u N)) rx < r" - using rx_def ra_def - by (auto simp: dist_real_def abs_diff_less_iff field_simps) - from dist[OF this] show "u N : S" using `u N ~: {\, -\}` - by (auto simp: extreal_real split: split_if_asm) - qed -qed - -lemma extreal_LimI_finite_iff: - assumes "\x\ \ \" - shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))" - (is "?lhs <-> ?rhs") -proof - assume lim: "u ----> x" - { fix r assume "(r::extreal)>0" - from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}" - apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"]) - using lim extreal_between[of x r] assms `r>0` by auto - hence "EX N. ALL n>=N. u n < x + r & x < u n + r" - using extreal_minus_less[of r x] by (cases r) auto - } then show "?rhs" by auto -next - assume ?rhs then show "u ----> x" - using extreal_LimI_finite[of x] assms by auto -qed - - -subsubsection {* @{text Liminf} and @{text Limsup} *} - -definition - "Liminf net f = (GREATEST l. \yx. y < f x) net)" - -definition - "Limsup net f = (LEAST l. \y>l. eventually (\x. f x < y) net)" - -lemma Liminf_Sup: - fixes f :: "'a => 'b::{complete_lattice, linorder}" - shows "Liminf net f = Sup {l. \yx. y < f x) net}" - by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def) - -lemma Limsup_Inf: - fixes f :: "'a => 'b::{complete_lattice, linorder}" - shows "Limsup net f = Inf {l. \y>l. eventually (\x. f x < y) net}" - by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def) - -lemma extreal_SupI: - fixes x :: extreal - assumes "\y. y \ A \ y \ x" - assumes "\y. (\z. z \ A \ z \ y) \ x \ y" - shows "Sup A = x" - unfolding Sup_extreal_def - using assms by (auto intro!: Least_equality) - -lemma extreal_InfI: - fixes x :: extreal - assumes "\i. i \ A \ x \ i" - assumes "\y. (\i. i \ A \ y \ i) \ y \ x" - shows "Inf A = x" - unfolding Inf_extreal_def - using assms by (auto intro!: Greatest_equality) - -lemma Limsup_const: - fixes c :: "'a::{complete_lattice, linorder}" - assumes ntriv: "\ trivial_limit net" - shows "Limsup net (\x. c) = c" - unfolding Limsup_Inf -proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower) - fix x assume *: "\y>x. eventually (\_. c < y) net" - show "c \ x" - proof (rule ccontr) - assume "\ c \ x" then have "x < c" by auto - then show False using ntriv * by (auto simp: trivial_limit_def) - qed -qed auto - -lemma Liminf_const: - fixes c :: "'a::{complete_lattice, linorder}" - assumes ntriv: "\ trivial_limit net" - shows "Liminf net (\x. c) = c" - unfolding Liminf_Sup -proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper) - fix x assume *: "\y_. y < c) net" - show "x \ c" - proof (rule ccontr) - assume "\ x \ c" then have "c < x" by auto - then show False using ntriv * by (auto simp: trivial_limit_def) - qed -qed auto - -lemma mono_set: - fixes S :: "('a::order) set" - shows "mono S \ (\x y. x \ y \ x \ S \ y \ S)" - by (auto simp: mono_def mem_def) - -lemma mono_greaterThan[intro, simp]: "mono {B<..}" unfolding mono_set by auto -lemma mono_atLeast[intro, simp]: "mono {B..}" unfolding mono_set by auto -lemma mono_UNIV[intro, simp]: "mono UNIV" unfolding mono_set by auto -lemma mono_empty[intro, simp]: "mono {}" unfolding mono_set by auto - -lemma mono_set_iff: - fixes S :: "'a::{linorder,complete_lattice} set" - defines "a \ Inf S" - shows "mono S \ (S = {a <..} \ S = {a..})" (is "_ = ?c") -proof - assume "mono S" - then have mono: "\x y. x \ y \ x \ S \ y \ S" by (auto simp: mono_set) - show ?c - proof cases - assume "a \ S" - show ?c - using mono[OF _ `a \ S`] - by (auto intro: complete_lattice_class.Inf_lower simp: a_def) - next - assume "a \ S" - have "S = {a <..}" - proof safe - fix x assume "x \ S" - then have "a \ x" unfolding a_def by (rule complete_lattice_class.Inf_lower) - then show "a < x" using `x \ S` `a \ S` by (cases "a = x") auto - next - fix x assume "a < x" - then obtain y where "y < x" "y \ S" unfolding a_def Inf_less_iff .. - with mono[of y x] show "x \ S" by auto - qed - then show ?c .. - qed -qed auto - -lemma lim_imp_Liminf: - fixes f :: "'a \ extreal" - assumes ntriv: "\ trivial_limit net" - assumes lim: "(f ---> f0) net" - shows "Liminf net f = f0" - unfolding Liminf_Sup -proof (safe intro!: extreal_SupI) - fix y assume *: "\y'x. y' < f x) net" - show "y \ f0" - proof (rule extreal_le_extreal) - fix B assume "B < y" - { assume "f0 < B" - then have "eventually (\x. f x < B \ B < f x) net" - using topological_tendstoD[OF lim, of "{..x. f x < B \ B < f x) = (\x. False)" by (auto simp: fun_eq_iff) - finally have False using ntriv[unfolded trivial_limit_def] by auto - } then show "B \ f0" by (metis linorder_le_less_linear) - qed -next - fix y assume *: "\z. z \ {l. \yx. y < f x) net} \ z \ y" - show "f0 \ y" - proof (safe intro!: *[rule_format]) - fix y assume "y < f0" then show "eventually (\x. y < f x) net" - using lim[THEN topological_tendstoD, of "{y <..}"] by auto - qed -qed - -lemma extreal_Liminf_le_Limsup: - fixes f :: "'a \ extreal" - assumes ntriv: "\ trivial_limit net" - shows "Liminf net f \ Limsup net f" - unfolding Limsup_Inf Liminf_Sup -proof (safe intro!: complete_lattice_class.Inf_greatest complete_lattice_class.Sup_least) - fix u v assume *: "\yx. y < f x) net" "\y>v. eventually (\x. f x < y) net" - show "u \ v" - proof (rule ccontr) - assume "\ u \ v" - then obtain t where "t < u" "v < t" - using extreal_dense[of v u] by (auto simp: not_le) - then have "eventually (\x. t < f x \ f x < t) net" - using * by (auto intro: eventually_conj) - also have "(\x. t < f x \ f x < t) = (\x. False)" by (auto simp: fun_eq_iff) - finally show False using ntriv by (auto simp: trivial_limit_def) - qed -qed - -lemma Liminf_mono: - fixes f g :: "'a => extreal" - assumes ev: "eventually (\x. f x \ g x) net" - shows "Liminf net f \ Liminf net g" - unfolding Liminf_Sup -proof (safe intro!: Sup_mono bexI) - fix a y assume "\yx. y < f x) net" and "y < a" - then have "eventually (\x. y < f x) net" by auto - then show "eventually (\x. y < g x) net" - by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto) -qed simp - -lemma Liminf_eq: - fixes f g :: "'a \ extreal" - assumes "eventually (\x. f x = g x) net" - shows "Liminf net f = Liminf net g" - by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto - -lemma Liminf_mono_all: - fixes f g :: "'a \ extreal" - assumes "\x. f x \ g x" - shows "Liminf net f \ Liminf net g" - using assms by (intro Liminf_mono always_eventually) auto - -lemma Limsup_mono: - fixes f g :: "'a \ extreal" - assumes ev: "eventually (\x. f x \ g x) net" - shows "Limsup net f \ Limsup net g" - unfolding Limsup_Inf -proof (safe intro!: Inf_mono bexI) - fix a y assume "\y>a. eventually (\x. g x < y) net" and "a < y" - then have "eventually (\x. g x < y) net" by auto - then show "eventually (\x. f x < y) net" - by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto) -qed simp - -lemma Limsup_mono_all: - fixes f g :: "'a \ extreal" - assumes "\x. f x \ g x" - shows "Limsup net f \ Limsup net g" - using assms by (intro Limsup_mono always_eventually) auto - -lemma Limsup_eq: - fixes f g :: "'a \ extreal" - assumes "eventually (\x. f x = g x) net" - shows "Limsup net f = Limsup net g" - by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto - -abbreviation "liminf \ Liminf sequentially" - -abbreviation "limsup \ Limsup sequentially" - -lemma (in complete_lattice) less_INFD: - assumes "y < INFI A f"" i \ A" shows "y < f i" -proof - - note `y < INFI A f` - also have "INFI A f \ f i" using `i \ A` by (rule INF_leI) - finally show "y < f i" . -qed - -lemma liminf_SUPR_INFI: - fixes f :: "nat \ extreal" - shows "liminf f = (SUP n. INF m:{n..}. f m)" - unfolding Liminf_Sup eventually_sequentially -proof (safe intro!: antisym complete_lattice_class.Sup_least) - fix x assume *: "\yN. \n\N. y < f n" show "x \ (SUP n. INF m:{n..}. f m)" - proof (rule extreal_le_extreal) - fix y assume "y < x" - with * obtain N where "\n. N \ n \ y < f n" by auto - then have "y \ (INF m:{N..}. f m)" by (force simp: le_INF_iff) - also have "\ \ (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto - finally show "y \ (SUP n. INF m:{n..}. f m)" . - qed -next - show "(SUP n. INF m:{n..}. f m) \ Sup {l. \yN. \n\N. y < f n}" - proof (unfold SUPR_def, safe intro!: Sup_mono bexI) - fix y n assume "y < INFI {n..} f" - from less_INFD[OF this] show "\N. \n\N. y < f n" by (intro exI[of _ n]) auto - qed (rule order_refl) -qed - -lemma tail_same_limsup: - fixes X Y :: "nat => extreal" - assumes "\n. N \ n \ X n = Y n" - shows "limsup X = limsup Y" - using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto - -lemma tail_same_liminf: - fixes X Y :: "nat => extreal" - assumes "\n. N \ n \ X n = Y n" - shows "liminf X = liminf Y" - using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto - -lemma liminf_mono: - fixes X Y :: "nat \ extreal" - assumes "\n. N \ n \ X n <= Y n" - shows "liminf X \ liminf Y" - using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto - -lemma limsup_mono: - fixes X Y :: "nat => extreal" - assumes "\n. N \ n \ X n <= Y n" - shows "limsup X \ limsup Y" - using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto - -declare trivial_limit_sequentially[simp] - -lemma - fixes X :: "nat \ extreal" - shows extreal_incseq_uminus[simp]: "incseq (\i. - X i) = decseq X" - and extreal_decseq_uminus[simp]: "decseq (\i. - X i) = incseq X" - unfolding incseq_def decseq_def by auto - -lemma liminf_bounded: - fixes X Y :: "nat \ extreal" - assumes "\n. N \ n \ C \ X n" - shows "C \ liminf X" - using liminf_mono[of N "\n. C" X] assms Liminf_const[of sequentially C] by simp - -lemma limsup_bounded: - fixes X Y :: "nat => extreal" - assumes "\n. N \ n \ X n <= C" - shows "limsup X \ C" - using limsup_mono[of N X "\n. C"] assms Limsup_const[of sequentially C] by simp - -lemma liminf_bounded_iff: - fixes x :: "nat \ extreal" - shows "C \ liminf x \ (\BN. \n\N. B < x n)" (is "?lhs <-> ?rhs") -proof safe - fix B assume "B < C" "C \ liminf x" - then have "B < liminf x" by auto - then obtain N where "B < (INF m:{N..}. x m)" - unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto - from less_INFD[OF this] show "\N. \n\N. B < x n" by auto -next - assume *: "\BN. \n\N. B < x n" - { fix B assume "Bn\N. B < x n" using `?rhs` by auto - hence "B \ (INF m:{N..}. x m)" by (intro le_INFI) auto - also have "... \ liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp - finally have "B \ liminf x" . - } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear) -qed - -lemma liminf_subseq_mono: - fixes X :: "nat \ extreal" - assumes "subseq r" - shows "liminf X \ liminf (X \ r) " -proof- - have "\n. (INF m:{n..}. X m) \ (INF m:{n..}. (X \ r) m)" - proof (safe intro!: INF_mono) - fix n m :: nat assume "n \ m" then show "\ma\{n..}. X ma \ (X \ r) m" - using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto - qed - then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def) -qed - -lemma extreal_real': assumes "\x\ \ \" shows "extreal (real x) = x" - using assms by auto - -lemma extreal_le_extreal_bounded: - fixes x y z :: extreal - assumes "z \ y" - assumes *: "\B. z < B \ B < x \ B \ y" - shows "x \ y" -proof (rule extreal_le_extreal) - fix B assume "B < x" - show "B \ y" - proof cases - assume "z < B" from *[OF this `B < x`] show "B \ y" . - next - assume "\ z < B" with `z \ y` show "B \ y" by auto - qed -qed - -lemma fixes x y :: extreal - shows Sup_atMost[simp]: "Sup {.. y} = y" - and Sup_lessThan[simp]: "Sup {..< y} = y" - and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" - and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" - and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" - by (auto simp: Sup_extreal_def intro!: Least_equality - intro: extreal_le_extreal extreal_le_extreal_bounded[of x]) - -lemma Sup_greaterThanlessThan[simp]: - fixes x y :: extreal assumes "x < y" shows "Sup { x <..< y} = y" - unfolding Sup_extreal_def -proof (intro Least_equality extreal_le_extreal_bounded[of _ _ y]) - fix z assume z: "\u\{x<.. z" - from extreal_dense[OF `x < y`] guess w .. note w = this - with z[THEN bspec, of w] show "x \ z" by auto -qed auto - -lemma real_extreal_id: "real o extreal = id" -proof- -{ fix x have "(real o extreal) x = id x" by auto } -from this show ?thesis using ext by blast -qed - -lemma open_image_extreal: "open(UNIV-{\,(-\)})" -by (metis range_extreal open_extreal open_UNIV) - -lemma extreal_le_distrib: - fixes a b c :: extreal shows "c * (a + b) \ c * a + c * b" - by (cases rule: extreal3_cases[of a b c]) - (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) - -lemma extreal_pos_distrib: - fixes a b c :: extreal assumes "0 \ c" "c \ \" shows "c * (a + b) = c * a + c * b" - using assms by (cases rule: extreal3_cases[of a b c]) - (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) - -lemma extreal_pos_le_distrib: -fixes a b c :: extreal -assumes "c>=0" -shows "c * (a + b) <= c * a + c * b" - using assms by (cases rule: extreal3_cases[of a b c]) - (auto simp add: field_simps) - -lemma extreal_max_mono: - "[| (a::extreal) <= b; c <= d |] ==> max a c <= max b d" - by (metis sup_extreal_def sup_mono) - - -lemma extreal_max_least: - "[| (a::extreal) <= x; c <= x |] ==> max a c <= x" - by (metis sup_extreal_def sup_least) - -end