hoelzl@44791: (* Title: HOL/Library/Extended_Real.thy wenzelm@42854: Author: Johannes Hölzl, TU München wenzelm@42854: Author: Robert Himmelmann, TU München wenzelm@42854: Author: Armin Heller, TU München wenzelm@42854: Author: Bogdan Grechuk, University of Edinburgh wenzelm@42854: *) hoelzl@42844: hoelzl@42844: header {* Extended real number line *} hoelzl@42844: hoelzl@44791: theory Extended_Real haftmann@44812: imports Complex_Main Extended_Nat hoelzl@42844: begin hoelzl@42844: hoelzl@42851: text {* hoelzl@42851: hoelzl@42851: For more lemmas about the extended real numbers go to bulwahn@43471: @{text "src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy"} hoelzl@42851: hoelzl@42851: *} hoelzl@42851: hoelzl@42850: lemma (in complete_lattice) atLeast_eq_UNIV_iff: "{x..} = UNIV \ x = bot" hoelzl@42850: proof hoelzl@42850: assume "{x..} = UNIV" hoelzl@42850: show "x = bot" hoelzl@42850: proof (rule ccontr) hoelzl@42850: assume "x \ bot" then have "bot \ {x..}" by (simp add: le_less) hoelzl@42850: then show False using `{x..} = UNIV` by simp hoelzl@42850: qed hoelzl@42850: qed auto hoelzl@42850: hoelzl@42850: lemma SUPR_pair: hoelzl@42850: "(SUP i : A. SUP j : B. f i j) = (SUP p : A \ B. f (fst p) (snd p))" hoelzl@42851: by (rule antisym) (auto intro!: SUP_leI le_SUPI2) hoelzl@42850: hoelzl@42850: lemma INFI_pair: hoelzl@42850: "(INF i : A. INF j : B. f i j) = (INF p : A \ B. f (fst p) (snd p))" hoelzl@42851: by (rule antisym) (auto intro!: le_INFI INF_leI2) hoelzl@42850: hoelzl@42844: subsection {* Definition and basic properties *} hoelzl@42844: hoelzl@44791: datatype ereal = ereal real | PInfty | MInfty hoelzl@42844: hoelzl@44791: instantiation ereal :: uminus hoelzl@42844: begin hoelzl@44791: fun uminus_ereal where hoelzl@44791: "- (ereal r) = ereal (- r)" hoelzl@44794: | "- PInfty = MInfty" hoelzl@44794: | "- MInfty = PInfty" hoelzl@42844: instance .. hoelzl@42844: end hoelzl@42844: hoelzl@44794: instantiation ereal :: infinity hoelzl@44794: begin hoelzl@44794: definition "(\::ereal) = PInfty" hoelzl@44794: instance .. hoelzl@44794: end hoelzl@42847: hoelzl@44795: definition "ereal_of_enat n = (case n of enat n \ ereal (real n) | \ \ \)" hoelzl@42844: hoelzl@44794: declare [[coercion "ereal :: real \ ereal"]] hoelzl@44794: declare [[coercion "ereal_of_enat :: enat \ ereal"]] hoelzl@44794: declare [[coercion "(\n. ereal (of_nat n)) :: nat \ ereal"]] hoelzl@42844: hoelzl@44791: lemma ereal_uminus_uminus[simp]: hoelzl@44791: fixes a :: ereal shows "- (- a) = a" hoelzl@42844: by (cases a) simp_all hoelzl@42844: hoelzl@44794: lemma hoelzl@44794: shows PInfty_eq_infinity[simp]: "PInfty = \" hoelzl@44794: and MInfty_eq_minfinity[simp]: "MInfty = - \" hoelzl@44794: and MInfty_neq_PInfty[simp]: "\ \ - (\::ereal)" "- \ \ (\::ereal)" hoelzl@44794: and MInfty_neq_ereal[simp]: "ereal r \ - \" "- \ \ ereal r" hoelzl@44794: and PInfty_neq_ereal[simp]: "ereal r \ \" "\ \ ereal r" hoelzl@44794: and PInfty_cases[simp]: "(case \ of ereal r \ f r | PInfty \ y | MInfty \ z) = y" hoelzl@44794: and MInfty_cases[simp]: "(case - \ of ereal r \ f r | PInfty \ y | MInfty \ z) = z" hoelzl@44794: by (simp_all add: infinity_ereal_def) hoelzl@42844: hoelzl@44804: declare hoelzl@44804: PInfty_eq_infinity[code_post] hoelzl@44804: MInfty_eq_minfinity[code_post] hoelzl@44804: hoelzl@44804: lemma [code_unfold]: hoelzl@44804: "\ = PInfty" hoelzl@44804: "-PInfty = MInfty" hoelzl@44804: by simp_all hoelzl@44804: hoelzl@44794: lemma inj_ereal[simp]: "inj_on ereal A" hoelzl@44794: unfolding inj_on_def by auto hoelzl@42844: hoelzl@44791: lemma ereal_cases[case_names real PInf MInf, cases type: ereal]: hoelzl@44791: assumes "\r. x = ereal r \ P" hoelzl@42844: assumes "x = \ \ P" hoelzl@42844: assumes "x = -\ \ P" hoelzl@42844: shows P hoelzl@42844: using assms by (cases x) auto hoelzl@42844: hoelzl@44791: lemmas ereal2_cases = ereal_cases[case_product ereal_cases] hoelzl@44791: lemmas ereal3_cases = ereal2_cases[case_product ereal_cases] hoelzl@42844: hoelzl@44791: lemma ereal_uminus_eq_iff[simp]: hoelzl@44791: fixes a b :: ereal shows "-a = -b \ a = b" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) simp_all hoelzl@42844: hoelzl@44791: function of_ereal :: "ereal \ real" where hoelzl@44791: "of_ereal (ereal r) = r" | hoelzl@44791: "of_ereal \ = 0" | hoelzl@44791: "of_ereal (-\) = 0" hoelzl@44791: by (auto intro: ereal_cases) hoelzl@42844: termination proof qed (rule wf_empty) hoelzl@42844: hoelzl@42844: defs (overloaded) hoelzl@44791: real_of_ereal_def [code_unfold]: "real \ of_ereal" hoelzl@42844: hoelzl@44791: lemma real_of_ereal[simp]: hoelzl@44791: "real (- x :: ereal) = - (real x)" hoelzl@44791: "real (ereal r) = r" hoelzl@44794: "real (\::ereal) = 0" hoelzl@44791: by (cases x) (simp_all add: real_of_ereal_def) hoelzl@42844: hoelzl@44791: lemma range_ereal[simp]: "range ereal = UNIV - {\, -\}" hoelzl@42844: proof safe hoelzl@44791: fix x assume "x \ range ereal" "x \ \" hoelzl@42844: then show "x = -\" by (cases x) auto hoelzl@42844: qed auto hoelzl@42844: hoelzl@44791: lemma ereal_range_uminus[simp]: "range uminus = (UNIV::ereal set)" hoelzl@42850: proof safe hoelzl@44791: fix x :: ereal show "x \ range uminus" by (intro image_eqI[of _ _ "-x"]) auto hoelzl@42850: qed auto hoelzl@42850: hoelzl@44791: instantiation ereal :: number hoelzl@42844: begin hoelzl@44791: definition [simp]: "number_of x = ereal (number_of x)" hoelzl@42844: instance proof qed hoelzl@42844: end hoelzl@42844: hoelzl@44791: instantiation ereal :: abs hoelzl@42847: begin hoelzl@44791: function abs_ereal where hoelzl@44791: "\ereal r\ = ereal \r\" hoelzl@44794: | "\-\\ = (\::ereal)" hoelzl@44794: | "\\\ = (\::ereal)" hoelzl@44791: by (auto intro: ereal_cases) hoelzl@42847: termination proof qed (rule wf_empty) hoelzl@42847: instance .. hoelzl@42847: end hoelzl@42847: hoelzl@44794: lemma abs_eq_infinity_cases[elim!]: "\ \x :: ereal\ = \ ; x = \ \ P ; x = -\ \ P \ \ P" hoelzl@42847: by (cases x) auto hoelzl@42847: hoelzl@44794: lemma abs_neq_infinity_cases[elim!]: "\ \x :: ereal\ \ \ ; \r. x = ereal r \ P \ \ P" hoelzl@42847: by (cases x) auto hoelzl@42847: hoelzl@44791: lemma abs_ereal_uminus[simp]: "\- x\ = \x::ereal\" hoelzl@42847: by (cases x) auto hoelzl@42847: hoelzl@42844: subsubsection "Addition" hoelzl@42844: hoelzl@44791: instantiation ereal :: comm_monoid_add hoelzl@42844: begin hoelzl@42844: hoelzl@44791: definition "0 = ereal 0" hoelzl@42844: hoelzl@44791: function plus_ereal where hoelzl@44791: "ereal r + ereal p = ereal (r + p)" | hoelzl@44794: "\ + a = (\::ereal)" | hoelzl@44794: "a + \ = (\::ereal)" | hoelzl@44791: "ereal r + -\ = - \" | hoelzl@44794: "-\ + ereal p = -(\::ereal)" | hoelzl@44794: "-\ + -\ = -(\::ereal)" hoelzl@42844: proof - hoelzl@42844: case (goal1 P x) hoelzl@42844: moreover then obtain a b where "x = (a, b)" by (cases x) auto hoelzl@42844: ultimately show P hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: qed auto hoelzl@42844: termination proof qed (rule wf_empty) hoelzl@42844: hoelzl@42844: lemma Infty_neq_0[simp]: hoelzl@44794: "(\::ereal) \ 0" "0 \ (\::ereal)" hoelzl@44794: "-(\::ereal) \ 0" "0 \ -(\::ereal)" hoelzl@44791: by (simp_all add: zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_eq_0[simp]: hoelzl@44791: "ereal r = 0 \ r = 0" hoelzl@44791: "0 = ereal r \ r = 0" hoelzl@44791: unfolding zero_ereal_def by simp_all hoelzl@42844: hoelzl@42844: instance hoelzl@42844: proof hoelzl@44791: fix a :: ereal show "0 + a = a" hoelzl@44791: by (cases a) (simp_all add: zero_ereal_def) hoelzl@44791: fix b :: ereal show "a + b = b + a" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) simp_all hoelzl@44791: fix c :: ereal show "a + b + c = a + (b + c)" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) simp_all hoelzl@42844: qed hoelzl@42844: end hoelzl@42844: hoelzl@44791: lemma real_of_ereal_0[simp]: "real (0::ereal) = 0" hoelzl@44791: unfolding real_of_ereal_def zero_ereal_def by simp hoelzl@43791: hoelzl@44791: lemma abs_ereal_zero[simp]: "\0\ = (0::ereal)" hoelzl@44791: unfolding zero_ereal_def abs_ereal.simps by simp hoelzl@42847: hoelzl@44791: lemma ereal_uminus_zero[simp]: hoelzl@44791: "- 0 = (0::ereal)" hoelzl@44791: by (simp add: zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_uminus_zero_iff[simp]: hoelzl@44791: fixes a :: ereal shows "-a = 0 \ a = 0" hoelzl@42844: by (cases a) simp_all hoelzl@42844: hoelzl@44791: lemma ereal_plus_eq_PInfty[simp]: hoelzl@44794: fixes a b :: ereal shows "a + b = \ \ a = \ \ b = \" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_plus_eq_MInfty[simp]: hoelzl@44794: fixes a b :: ereal shows "a + b = -\ \ hoelzl@42844: (a = -\ \ b = -\) \ a \ \ \ b \ \" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_add_cancel_left: hoelzl@44794: fixes a b :: ereal assumes "a \ -\" hoelzl@42844: shows "a + b = a + c \ (a = \ \ b = c)" hoelzl@44791: using assms by (cases rule: ereal3_cases[of a b c]) auto hoelzl@42844: hoelzl@44791: lemma ereal_add_cancel_right: hoelzl@44794: fixes a b :: ereal assumes "a \ -\" hoelzl@42844: shows "b + a = c + a \ (a = \ \ b = c)" hoelzl@44791: using assms by (cases rule: ereal3_cases[of a b c]) auto hoelzl@42844: hoelzl@44791: lemma ereal_real: hoelzl@44791: "ereal (real x) = (if \x\ = \ then 0 else x)" hoelzl@42844: by (cases x) simp_all hoelzl@42844: hoelzl@44791: lemma real_of_ereal_add: hoelzl@44791: fixes a b :: ereal hoelzl@42850: shows "real (a + b) = (if (\a\ = \) \ (\b\ = \) \ (\a\ \ \) \ (\b\ \ \) then real a + real b else 0)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42850: hoelzl@44791: subsubsection "Linear order on @{typ ereal}" hoelzl@42844: hoelzl@44791: instantiation ereal :: linorder hoelzl@42844: begin hoelzl@42844: hoelzl@44791: function less_ereal where hoelzl@44794: " ereal x < ereal y \ x < y" | hoelzl@44794: "(\::ereal) < a \ False" | hoelzl@44794: " a < -(\::ereal) \ False" | hoelzl@44794: "ereal x < \ \ True" | hoelzl@44794: " -\ < ereal r \ True" | hoelzl@44794: " -\ < (\::ereal) \ True" hoelzl@42844: proof - hoelzl@42844: case (goal1 P x) hoelzl@42844: moreover then obtain a b where "x = (a,b)" by (cases x) auto hoelzl@44791: ultimately show P by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: qed simp_all hoelzl@42844: termination by (relation "{}") simp hoelzl@42844: hoelzl@44791: definition "x \ (y::ereal) \ x < y \ x = y" hoelzl@42844: hoelzl@44791: lemma ereal_infty_less[simp]: hoelzl@44794: fixes x :: ereal hoelzl@44794: shows "x < \ \ (x \ \)" hoelzl@44794: "-\ < x \ (x \ -\)" hoelzl@42844: by (cases x, simp_all) (cases x, simp_all) hoelzl@42844: hoelzl@44791: lemma ereal_infty_less_eq[simp]: hoelzl@44794: fixes x :: ereal hoelzl@44794: shows "\ \ x \ x = \" hoelzl@42844: "x \ -\ \ x = -\" hoelzl@44791: by (auto simp add: less_eq_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_less[simp]: hoelzl@44791: "ereal r < 0 \ (r < 0)" hoelzl@44791: "0 < ereal r \ (0 < r)" hoelzl@44794: "0 < (\::ereal)" hoelzl@44794: "-(\::ereal) < 0" hoelzl@44791: by (simp_all add: zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_less_eq[simp]: hoelzl@44794: "x \ (\::ereal)" hoelzl@44794: "-(\::ereal) \ x" hoelzl@44791: "ereal r \ ereal p \ r \ p" hoelzl@44791: "ereal r \ 0 \ r \ 0" hoelzl@44791: "0 \ ereal r \ 0 \ r" hoelzl@44791: by (auto simp add: less_eq_ereal_def zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_infty_less_eq2: hoelzl@44794: "a \ b \ a = \ \ b = (\::ereal)" hoelzl@44794: "a \ b \ b = -\ \ a = -(\::ereal)" hoelzl@42844: by simp_all hoelzl@42844: hoelzl@42844: instance hoelzl@42844: proof hoelzl@44791: fix x :: ereal show "x \ x" hoelzl@42844: by (cases x) simp_all hoelzl@44791: fix y :: ereal show "x < y \ x \ y \ \ y \ x" hoelzl@44791: by (cases rule: ereal2_cases[of x y]) auto hoelzl@42844: show "x \ y \ y \ x " hoelzl@44791: by (cases rule: ereal2_cases[of x y]) auto hoelzl@42844: { assume "x \ y" "y \ x" then show "x = y" hoelzl@44791: by (cases rule: ereal2_cases[of x y]) auto } hoelzl@42844: { fix z assume "x \ y" "y \ z" then show "x \ z" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) auto } hoelzl@42844: qed hoelzl@42844: end hoelzl@42844: hoelzl@44791: instance ereal :: ordered_ab_semigroup_add hoelzl@42849: proof hoelzl@44791: fix a b c :: ereal assume "a \ b" then show "c + a \ c + b" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) auto hoelzl@42849: qed hoelzl@42849: hoelzl@44791: lemma real_of_ereal_positive_mono: hoelzl@44794: fixes x y :: ereal shows "\0 \ x; x \ y; y \ \\ \ real x \ real y" hoelzl@44791: by (cases rule: ereal2_cases[of x y]) auto hoelzl@43791: hoelzl@44791: lemma ereal_MInfty_lessI[intro, simp]: hoelzl@44794: fixes a :: ereal shows "a \ -\ \ -\ < a" hoelzl@42844: by (cases a) auto hoelzl@42844: hoelzl@44791: lemma ereal_less_PInfty[intro, simp]: hoelzl@44794: fixes a :: ereal shows "a \ \ \ a < \" hoelzl@42844: by (cases a) auto hoelzl@42844: hoelzl@44791: lemma ereal_less_ereal_Ex: hoelzl@44791: fixes a b :: ereal hoelzl@44791: shows "x < ereal r \ x = -\ \ (\p. p < r \ x = ereal p)" hoelzl@42844: by (cases x) auto hoelzl@42844: hoelzl@44791: lemma less_PInf_Ex_of_nat: "x \ \ \ (\n::nat. x < ereal (real n))" hoelzl@42850: proof (cases x) hoelzl@42850: case (real r) then show ?thesis hoelzl@42851: using reals_Archimedean2[of r] by simp hoelzl@42850: qed simp_all hoelzl@42850: hoelzl@44791: lemma ereal_add_mono: hoelzl@44791: fixes a b c d :: ereal assumes "a \ b" "c \ d" shows "a + c \ b + d" hoelzl@42844: using assms hoelzl@42844: apply (cases a) hoelzl@44791: apply (cases rule: ereal3_cases[of b c d], auto) hoelzl@44791: apply (cases rule: ereal3_cases[of b c d], auto) hoelzl@42844: done hoelzl@42844: hoelzl@44791: lemma ereal_minus_le_minus[simp]: hoelzl@44791: fixes a b :: ereal shows "- a \ - b \ b \ a" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_minus_less_minus[simp]: hoelzl@44791: fixes a b :: ereal shows "- a < - b \ b < a" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_le_real_iff: hoelzl@44791: "x \ real y \ ((\y\ \ \ \ ereal x \ y) \ (\y\ = \ \ x \ 0))" hoelzl@42844: by (cases y) auto hoelzl@42844: hoelzl@44791: lemma real_le_ereal_iff: hoelzl@44791: "real y \ x \ ((\y\ \ \ \ y \ ereal x) \ (\y\ = \ \ 0 \ x))" hoelzl@42844: by (cases y) auto hoelzl@42844: hoelzl@44791: lemma ereal_less_real_iff: hoelzl@44791: "x < real y \ ((\y\ \ \ \ ereal x < y) \ (\y\ = \ \ x < 0))" hoelzl@42844: by (cases y) auto hoelzl@42844: hoelzl@44791: lemma real_less_ereal_iff: hoelzl@44791: "real y < x \ ((\y\ \ \ \ y < ereal x) \ (\y\ = \ \ 0 < x))" hoelzl@42844: by (cases y) auto hoelzl@42844: hoelzl@44791: lemma real_of_ereal_pos: hoelzl@44791: fixes x :: ereal shows "0 \ x \ 0 \ real x" by (cases x) auto hoelzl@42850: hoelzl@44791: lemmas real_of_ereal_ord_simps = hoelzl@44791: ereal_le_real_iff real_le_ereal_iff ereal_less_real_iff real_less_ereal_iff hoelzl@42844: hoelzl@44791: lemma abs_ereal_ge0[simp]: "0 \ x \ \x :: ereal\ = x" hoelzl@43791: by (cases x) auto hoelzl@43791: hoelzl@44791: lemma abs_ereal_less0[simp]: "x < 0 \ \x :: ereal\ = -x" hoelzl@43791: by (cases x) auto hoelzl@43791: hoelzl@44791: lemma abs_ereal_pos[simp]: "0 \ \x :: ereal\" hoelzl@43791: by (cases x) auto hoelzl@43791: hoelzl@44794: lemma real_of_ereal_le_0[simp]: "real (x :: ereal) \ 0 \ (x \ 0 \ x = \)" hoelzl@44794: by (cases x) auto hoelzl@43791: hoelzl@44794: lemma abs_real_of_ereal[simp]: "\real (x :: ereal)\ = real \x\" hoelzl@44794: by (cases x) auto hoelzl@43791: hoelzl@44794: lemma zero_less_real_of_ereal: hoelzl@44794: fixes x :: ereal shows "0 < real x \ (0 < x \ x \ \)" hoelzl@44794: by (cases x) auto hoelzl@43791: hoelzl@44791: lemma ereal_0_le_uminus_iff[simp]: hoelzl@44791: fixes a :: ereal shows "0 \ -a \ a \ 0" hoelzl@44791: by (cases rule: ereal2_cases[of a]) auto hoelzl@43791: hoelzl@44791: lemma ereal_uminus_le_0_iff[simp]: hoelzl@44791: fixes a :: ereal shows "-a \ 0 \ 0 \ a" hoelzl@44791: by (cases rule: ereal2_cases[of a]) auto hoelzl@43791: hoelzl@44794: lemma ereal_dense2: "x < y \ \z. x < ereal z \ ereal z < y" hoelzl@44794: using lt_ex gt_ex dense by (cases x y rule: ereal2_cases) auto hoelzl@44794: hoelzl@44791: lemma ereal_dense: hoelzl@44791: fixes x y :: ereal assumes "x < y" hoelzl@44794: shows "\z. x < z \ z < y" hoelzl@44794: using ereal_dense2[OF `x < y`] by blast hoelzl@42844: hoelzl@44791: lemma ereal_add_strict_mono: hoelzl@44791: fixes a b c d :: ereal hoelzl@42850: assumes "a = b" "0 \ a" "a \ \" "c < d" hoelzl@42850: shows "a + c < b + d" hoelzl@44791: using assms by (cases rule: ereal3_cases[case_product ereal_cases, of a b c d]) auto hoelzl@42850: hoelzl@44794: lemma ereal_less_add: hoelzl@44794: fixes a b c :: ereal shows "\a\ \ \ \ c < b \ a + c < a + b" hoelzl@44791: by (cases rule: ereal2_cases[of b c]) auto hoelzl@42850: hoelzl@44791: lemma ereal_uminus_eq_reorder: "- a = b \ a = (-b::ereal)" by auto hoelzl@42850: hoelzl@44791: lemma ereal_uminus_less_reorder: "- a < b \ -b < (a::ereal)" hoelzl@44791: by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_less_minus) hoelzl@42850: hoelzl@44791: lemma ereal_uminus_le_reorder: "- a \ b \ -b \ (a::ereal)" hoelzl@44791: by (subst (3) ereal_uminus_uminus[symmetric]) (simp only: ereal_minus_le_minus) hoelzl@42850: hoelzl@44791: lemmas ereal_uminus_reorder = hoelzl@44791: ereal_uminus_eq_reorder ereal_uminus_less_reorder ereal_uminus_le_reorder hoelzl@42850: hoelzl@44791: lemma ereal_bot: hoelzl@44791: fixes x :: ereal assumes "\B. x \ ereal B" shows "x = - \" hoelzl@42850: proof (cases x) hoelzl@42850: case (real r) with assms[of "r - 1"] show ?thesis by auto hoelzl@42850: next case PInf with assms[of 0] show ?thesis by auto hoelzl@42850: next case MInf then show ?thesis by simp hoelzl@42850: qed hoelzl@42850: hoelzl@44791: lemma ereal_top: hoelzl@44791: fixes x :: ereal assumes "\B. x \ ereal B" shows "x = \" hoelzl@42850: proof (cases x) hoelzl@42850: case (real r) with assms[of "r + 1"] show ?thesis by auto hoelzl@42850: next case MInf with assms[of 0] show ?thesis by auto hoelzl@42850: next case PInf then show ?thesis by simp hoelzl@42850: qed hoelzl@42850: hoelzl@42850: lemma hoelzl@44791: shows ereal_max[simp]: "ereal (max x y) = max (ereal x) (ereal y)" hoelzl@44791: and ereal_min[simp]: "ereal (min x y) = min (ereal x) (ereal y)" hoelzl@42850: by (simp_all add: min_def max_def) hoelzl@42850: hoelzl@44791: lemma ereal_max_0: "max 0 (ereal r) = ereal (max 0 r)" hoelzl@44791: by (auto simp: zero_ereal_def) hoelzl@42850: hoelzl@42849: lemma hoelzl@44791: fixes f :: "nat \ ereal" hoelzl@42849: shows incseq_uminus[simp]: "incseq (\x. - f x) \ decseq f" hoelzl@42849: and decseq_uminus[simp]: "decseq (\x. - f x) \ incseq f" hoelzl@42849: unfolding decseq_def incseq_def by auto hoelzl@42849: hoelzl@44791: lemma incseq_ereal: "incseq f \ incseq (\x. ereal (f x))" hoelzl@43791: unfolding incseq_def by auto hoelzl@43791: hoelzl@44791: lemma ereal_add_nonneg_nonneg: hoelzl@44791: fixes a b :: ereal shows "0 \ a \ 0 \ b \ 0 \ a + b" hoelzl@42849: using add_mono[of 0 a 0 b] by simp hoelzl@42849: hoelzl@42849: lemma image_eqD: "f ` A = B \ (\x\A. f x \ B)" hoelzl@42849: by auto hoelzl@42849: hoelzl@42849: lemma incseq_setsumI: hoelzl@42850: fixes f :: "nat \ 'a::{comm_monoid_add, ordered_ab_semigroup_add}" hoelzl@42849: assumes "\i. 0 \ f i" hoelzl@42849: shows "incseq (\i. setsum f {..< i})" hoelzl@42849: proof (intro incseq_SucI) hoelzl@42849: fix n have "setsum f {..< n} + 0 \ setsum f {.. setsum f {..< Suc n}" hoelzl@42849: by auto hoelzl@42849: qed hoelzl@42849: hoelzl@42850: lemma incseq_setsumI2: hoelzl@42850: fixes f :: "'i \ nat \ 'a::{comm_monoid_add, ordered_ab_semigroup_add}" hoelzl@42850: assumes "\n. n \ A \ incseq (f n)" hoelzl@42850: shows "incseq (\i. \n\A. f n i)" hoelzl@42850: using assms unfolding incseq_def by (auto intro: setsum_mono) hoelzl@42850: hoelzl@42844: subsubsection "Multiplication" hoelzl@42844: hoelzl@44791: instantiation ereal :: "{comm_monoid_mult, sgn}" hoelzl@42844: begin hoelzl@42844: hoelzl@44791: definition "1 = ereal 1" hoelzl@42844: hoelzl@44791: function sgn_ereal where hoelzl@44791: "sgn (ereal r) = ereal (sgn r)" hoelzl@44794: | "sgn (\::ereal) = 1" hoelzl@44794: | "sgn (-\::ereal) = -1" hoelzl@44791: by (auto intro: ereal_cases) hoelzl@42847: termination proof qed (rule wf_empty) hoelzl@42847: hoelzl@44791: function times_ereal where hoelzl@44791: "ereal r * ereal p = ereal (r * p)" | hoelzl@44791: "ereal r * \ = (if r = 0 then 0 else if r > 0 then \ else -\)" | hoelzl@44791: "\ * ereal r = (if r = 0 then 0 else if r > 0 then \ else -\)" | hoelzl@44791: "ereal r * -\ = (if r = 0 then 0 else if r > 0 then -\ else \)" | hoelzl@44791: "-\ * ereal r = (if r = 0 then 0 else if r > 0 then -\ else \)" | hoelzl@44794: "(\::ereal) * \ = \" | hoelzl@44794: "-(\::ereal) * \ = -\" | hoelzl@44794: "(\::ereal) * -\ = -\" | hoelzl@44794: "-(\::ereal) * -\ = \" hoelzl@42844: proof - hoelzl@42844: case (goal1 P x) hoelzl@42844: moreover then obtain a b where "x = (a, b)" by (cases x) auto hoelzl@44791: ultimately show P by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: qed simp_all hoelzl@42844: termination by (relation "{}") simp hoelzl@42844: hoelzl@42844: instance hoelzl@42844: proof hoelzl@44791: fix a :: ereal show "1 * a = a" hoelzl@44791: by (cases a) (simp_all add: one_ereal_def) hoelzl@44791: fix b :: ereal show "a * b = b * a" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) simp_all hoelzl@44791: fix c :: ereal show "a * b * c = a * (b * c)" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@44791: (simp_all add: zero_ereal_def zero_less_mult_iff) hoelzl@42844: qed hoelzl@42844: end hoelzl@42844: hoelzl@44791: lemma real_of_ereal_le_1: hoelzl@44791: fixes a :: ereal shows "a \ 1 \ real a \ 1" hoelzl@44791: by (cases a) (auto simp: one_ereal_def) hoelzl@43791: hoelzl@44791: lemma abs_ereal_one[simp]: "\1\ = (1::ereal)" hoelzl@44791: unfolding one_ereal_def by simp hoelzl@42847: hoelzl@44791: lemma ereal_mult_zero[simp]: hoelzl@44791: fixes a :: ereal shows "a * 0 = 0" hoelzl@44791: by (cases a) (simp_all add: zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_zero_mult[simp]: hoelzl@44791: fixes a :: ereal shows "0 * a = 0" hoelzl@44791: by (cases a) (simp_all add: zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_m1_less_0[simp]: hoelzl@44791: "-(1::ereal) < 0" hoelzl@44791: by (simp add: zero_ereal_def one_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_zero_m1[simp]: hoelzl@44791: "1 \ (0::ereal)" hoelzl@44791: by (simp add: zero_ereal_def one_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_times_0[simp]: hoelzl@44791: fixes x :: ereal shows "0 * x = 0" hoelzl@44791: by (cases x) (auto simp: zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_times[simp]: hoelzl@44794: "1 \ (\::ereal)" "(\::ereal) \ 1" hoelzl@44794: "1 \ -(\::ereal)" "-(\::ereal) \ 1" hoelzl@44791: by (auto simp add: times_ereal_def one_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_plus_1[simp]: hoelzl@44791: "1 + ereal r = ereal (r + 1)" "ereal r + 1 = ereal (r + 1)" hoelzl@44794: "1 + -(\::ereal) = -\" "-(\::ereal) + 1 = -\" hoelzl@44791: unfolding one_ereal_def by auto hoelzl@42844: hoelzl@44791: lemma ereal_zero_times[simp]: hoelzl@44791: fixes a b :: ereal shows "a * b = 0 \ a = 0 \ b = 0" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_mult_eq_PInfty[simp]: hoelzl@44794: shows "a * b = (\::ereal) \ hoelzl@42844: (a = \ \ b > 0) \ (a > 0 \ b = \) \ (a = -\ \ b < 0) \ (a < 0 \ b = -\)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_mult_eq_MInfty[simp]: hoelzl@44794: shows "a * b = -(\::ereal) \ hoelzl@42844: (a = \ \ b < 0) \ (a < 0 \ b = \) \ (a = -\ \ b > 0) \ (a > 0 \ b = -\)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_0_less_1[simp]: "0 < (1::ereal)" hoelzl@44791: by (simp_all add: zero_ereal_def one_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_zero_one[simp]: "0 \ (1::ereal)" hoelzl@44791: by (simp_all add: zero_ereal_def one_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_mult_minus_left[simp]: hoelzl@44791: fixes a b :: ereal shows "-a * b = - (a * b)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_mult_minus_right[simp]: hoelzl@44791: fixes a b :: ereal shows "a * -b = - (a * b)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_mult_infty[simp]: hoelzl@44794: "a * (\::ereal) = (if a = 0 then 0 else if 0 < a then \ else - \)" hoelzl@42844: by (cases a) auto hoelzl@42844: hoelzl@44791: lemma ereal_infty_mult[simp]: hoelzl@44794: "(\::ereal) * a = (if a = 0 then 0 else if 0 < a then \ else - \)" hoelzl@42844: by (cases a) auto hoelzl@42844: hoelzl@44791: lemma ereal_mult_strict_right_mono: hoelzl@44794: assumes "a < b" and "0 < c" "c < (\::ereal)" hoelzl@42844: shows "a * c < b * c" hoelzl@42844: using assms hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@44791: (auto simp: zero_le_mult_iff ereal_less_PInfty) hoelzl@42844: hoelzl@44791: lemma ereal_mult_strict_left_mono: hoelzl@44794: "\ a < b ; 0 < c ; c < (\::ereal)\ \ c * a < c * b" hoelzl@44791: using ereal_mult_strict_right_mono by (simp add: mult_commute[of c]) hoelzl@42844: hoelzl@44791: lemma ereal_mult_right_mono: hoelzl@44791: fixes a b c :: ereal shows "\a \ b; 0 \ c\ \ a*c \ b*c" hoelzl@42844: using assms hoelzl@42844: apply (cases "c = 0") apply simp hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@44791: (auto simp: zero_le_mult_iff ereal_less_PInfty) hoelzl@42844: hoelzl@44791: lemma ereal_mult_left_mono: hoelzl@44791: fixes a b c :: ereal shows "\a \ b; 0 \ c\ \ c * a \ c * b" hoelzl@44791: using ereal_mult_right_mono by (simp add: mult_commute[of c]) hoelzl@42844: hoelzl@44791: lemma zero_less_one_ereal[simp]: "0 \ (1::ereal)" hoelzl@44791: by (simp add: one_ereal_def zero_ereal_def) hoelzl@42849: hoelzl@44791: lemma ereal_0_le_mult[simp]: "0 \ a \ 0 \ b \ 0 \ a * (b :: ereal)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) (auto simp: mult_nonneg_nonneg) hoelzl@42850: hoelzl@44791: lemma ereal_right_distrib: hoelzl@44791: fixes r a b :: ereal shows "0 \ a \ 0 \ b \ r * (a + b) = r * a + r * b" hoelzl@44791: by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps) hoelzl@42850: hoelzl@44791: lemma ereal_left_distrib: hoelzl@44791: fixes r a b :: ereal shows "0 \ a \ 0 \ b \ (a + b) * r = a * r + b * r" hoelzl@44791: by (cases rule: ereal3_cases[of r a b]) (simp_all add: field_simps) hoelzl@42850: hoelzl@44791: lemma ereal_mult_le_0_iff: hoelzl@44791: fixes a b :: ereal hoelzl@42850: shows "a * b \ 0 \ (0 \ a \ b \ 0) \ (a \ 0 \ 0 \ b)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_le_0_iff) hoelzl@42850: hoelzl@44791: lemma ereal_zero_le_0_iff: hoelzl@44791: fixes a b :: ereal hoelzl@42850: shows "0 \ a * b \ (0 \ a \ 0 \ b) \ (a \ 0 \ b \ 0)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_le_mult_iff) hoelzl@42850: hoelzl@44791: lemma ereal_mult_less_0_iff: hoelzl@44791: fixes a b :: ereal hoelzl@42850: shows "a * b < 0 \ (0 < a \ b < 0) \ (a < 0 \ 0 < b)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) (simp_all add: mult_less_0_iff) hoelzl@42850: hoelzl@44791: lemma ereal_zero_less_0_iff: hoelzl@44791: fixes a b :: ereal hoelzl@42850: shows "0 < a * b \ (0 < a \ 0 < b) \ (a < 0 \ b < 0)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) (simp_all add: zero_less_mult_iff) hoelzl@42850: hoelzl@44791: lemma ereal_distrib: hoelzl@44791: fixes a b c :: ereal hoelzl@42850: assumes "a \ \ \ b \ -\" "a \ -\ \ b \ \" "\c\ \ \" hoelzl@42850: shows "(a + b) * c = a * c + b * c" hoelzl@42850: using assms hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) hoelzl@42850: hoelzl@44791: lemma ereal_le_epsilon: hoelzl@44791: fixes x y :: ereal hoelzl@42850: assumes "ALL e. 0 < e --> x <= y + e" hoelzl@42850: shows "x <= y" hoelzl@42850: proof- hoelzl@44791: { assume a: "EX r. y = ereal r" hoelzl@44791: from this obtain r where r_def: "y = ereal r" by auto hoelzl@42850: { assume "x=(-\)" hence ?thesis by auto } hoelzl@42850: moreover hoelzl@42850: { assume "~(x=(-\))" hoelzl@44791: from this obtain p where p_def: "x = ereal p" hoelzl@42850: using a assms[rule_format, of 1] by (cases x) auto hoelzl@42850: { fix e have "0 < e --> p <= r + e" hoelzl@44791: using assms[rule_format, of "ereal e"] p_def r_def by auto } hoelzl@42850: hence "p <= r" apply (subst field_le_epsilon) by auto hoelzl@42850: hence ?thesis using r_def p_def by auto hoelzl@42850: } ultimately have ?thesis by blast hoelzl@42850: } hoelzl@42850: moreover hoelzl@42850: { assume "y=(-\) | y=\" hence ?thesis hoelzl@42850: using assms[rule_format, of 1] by (cases x) auto hoelzl@42850: } ultimately show ?thesis by (cases y) auto hoelzl@42850: qed hoelzl@42850: hoelzl@42850: hoelzl@44791: lemma ereal_le_epsilon2: hoelzl@44791: fixes x y :: ereal hoelzl@44791: assumes "ALL e. 0 < e --> x <= y + ereal e" hoelzl@42850: shows "x <= y" hoelzl@42850: proof- hoelzl@44791: { fix e :: ereal assume "e>0" hoelzl@42850: { assume "e=\" hence "x<=y+e" by auto } hoelzl@42850: moreover hoelzl@42850: { assume "e~=\" hoelzl@44791: from this obtain r where "e = ereal r" using `e>0` apply (cases e) by auto hoelzl@42850: hence "x<=y+e" using assms[rule_format, of r] `e>0` by auto hoelzl@42850: } ultimately have "x<=y+e" by blast hoelzl@44791: } from this show ?thesis using ereal_le_epsilon by auto hoelzl@42850: qed hoelzl@42850: hoelzl@44791: lemma ereal_le_real: hoelzl@44791: fixes x y :: ereal hoelzl@44791: assumes "ALL z. x <= ereal z --> y <= ereal z" hoelzl@42850: shows "y <= x" hoelzl@44794: by (metis assms ereal_bot ereal_cases ereal_infty_less_eq ereal_less_eq linorder_le_cases) hoelzl@42850: hoelzl@44791: lemma ereal_le_ereal: hoelzl@44791: fixes x y :: ereal hoelzl@42850: assumes "\B. B < x \ B <= y" hoelzl@42850: shows "x <= y" hoelzl@44791: by (metis assms ereal_dense leD linorder_le_less_linear) hoelzl@42850: hoelzl@44791: lemma ereal_ge_ereal: hoelzl@44791: fixes x y :: ereal hoelzl@42850: assumes "ALL B. B>x --> B >= y" hoelzl@42850: shows "x >= y" hoelzl@44791: by (metis assms ereal_dense leD linorder_le_less_linear) hoelzl@42849: hoelzl@44791: lemma setprod_ereal_0: hoelzl@44791: fixes f :: "'a \ ereal" hoelzl@43791: shows "(\i\A. f i) = 0 \ (finite A \ (\i\A. f i = 0))" hoelzl@43791: proof cases hoelzl@43791: assume "finite A" hoelzl@43791: then show ?thesis by (induct A) auto hoelzl@43791: qed auto hoelzl@43791: hoelzl@44791: lemma setprod_ereal_pos: hoelzl@44791: fixes f :: "'a \ ereal" assumes pos: "\i. i \ I \ 0 \ f i" shows "0 \ (\i\I. f i)" hoelzl@43791: proof cases hoelzl@43791: assume "finite I" from this pos show ?thesis by induct auto hoelzl@43791: qed simp hoelzl@43791: hoelzl@43791: lemma setprod_PInf: hoelzl@44794: fixes f :: "'a \ ereal" hoelzl@43791: assumes "\i. i \ I \ 0 \ f i" hoelzl@43791: shows "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" hoelzl@43791: proof cases hoelzl@43791: assume "finite I" from this assms show ?thesis hoelzl@43791: proof (induct I) hoelzl@43791: case (insert i I) hoelzl@44791: then have pos: "0 \ f i" "0 \ setprod f I" by (auto intro!: setprod_ereal_pos) hoelzl@43791: from insert have "(\j\insert i I. f j) = \ \ setprod f I * f i = \" by auto hoelzl@43791: also have "\ \ (setprod f I = \ \ f i = \) \ f i \ 0 \ setprod f I \ 0" hoelzl@44791: using setprod_ereal_pos[of I f] pos hoelzl@44791: by (cases rule: ereal2_cases[of "f i" "setprod f I"]) auto hoelzl@43791: also have "\ \ finite (insert i I) \ (\j\insert i I. f j = \) \ (\j\insert i I. f j \ 0)" hoelzl@44791: using insert by (auto simp: setprod_ereal_0) hoelzl@43791: finally show ?case . hoelzl@43791: qed simp hoelzl@43791: qed simp hoelzl@43791: hoelzl@44791: lemma setprod_ereal: "(\i\A. ereal (f i)) = ereal (setprod f A)" hoelzl@43791: proof cases hoelzl@43791: assume "finite A" then show ?thesis hoelzl@44791: by induct (auto simp: one_ereal_def) hoelzl@44791: qed (simp add: one_ereal_def) hoelzl@43791: hoelzl@42849: subsubsection {* Power *} hoelzl@42849: hoelzl@44791: instantiation ereal :: power hoelzl@42849: begin hoelzl@44791: primrec power_ereal where hoelzl@44791: "power_ereal x 0 = 1" | hoelzl@44791: "power_ereal x (Suc n) = x * x ^ n" hoelzl@42849: instance .. hoelzl@42849: end hoelzl@42849: hoelzl@44791: lemma ereal_power[simp]: "(ereal x) ^ n = ereal (x^n)" hoelzl@44791: by (induct n) (auto simp: one_ereal_def) hoelzl@42849: hoelzl@44794: lemma ereal_power_PInf[simp]: "(\::ereal) ^ n = (if n = 0 then 1 else \)" hoelzl@44791: by (induct n) (auto simp: one_ereal_def) hoelzl@42849: hoelzl@44791: lemma ereal_power_uminus[simp]: hoelzl@44791: fixes x :: ereal hoelzl@42849: shows "(- x) ^ n = (if even n then x ^ n else - (x^n))" hoelzl@44791: by (induct n) (auto simp: one_ereal_def) hoelzl@42849: hoelzl@44791: lemma ereal_power_number_of[simp]: hoelzl@44791: "(number_of num :: ereal) ^ n = ereal (number_of num ^ n)" hoelzl@44791: by (induct n) (auto simp: one_ereal_def) hoelzl@42850: hoelzl@44791: lemma zero_le_power_ereal[simp]: hoelzl@44791: fixes a :: ereal assumes "0 \ a" hoelzl@42850: shows "0 \ a ^ n" hoelzl@44791: using assms by (induct n) (auto simp: ereal_zero_le_0_iff) hoelzl@42850: hoelzl@42844: subsubsection {* Subtraction *} hoelzl@42844: hoelzl@44791: lemma ereal_minus_minus_image[simp]: hoelzl@44791: fixes S :: "ereal set" hoelzl@42844: shows "uminus ` uminus ` S = S" hoelzl@42844: by (auto simp: image_iff) hoelzl@42844: hoelzl@44791: lemma ereal_uminus_lessThan[simp]: hoelzl@44791: fixes a :: ereal shows "uminus ` {.. - ereal r = -\" hoelzl@44791: "ereal r - \ = -\" hoelzl@44794: "(\::ereal) - x = \" hoelzl@44794: "-(\::ereal) - \ = -\" hoelzl@42844: "x - -y = x + y" hoelzl@42844: "x - 0 = x" hoelzl@42844: "0 - x = -x" hoelzl@44791: by (simp_all add: minus_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_x_minus_x[simp]: hoelzl@44794: "x - x = (if \x\ = \ then \ else 0::ereal)" hoelzl@42844: by (cases x) simp_all hoelzl@42844: hoelzl@44791: lemma ereal_eq_minus_iff: hoelzl@44791: fixes x y z :: ereal hoelzl@42844: shows "x = z - y \ hoelzl@42847: (\y\ \ \ \ x + y = z) \ hoelzl@42844: (y = -\ \ x = \) \ hoelzl@42844: (y = \ \ z = \ \ x = \) \ hoelzl@42844: (y = \ \ z \ \ \ x = -\)" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) auto hoelzl@42844: hoelzl@44791: lemma ereal_eq_minus: hoelzl@44791: fixes x y z :: ereal hoelzl@42847: shows "\y\ \ \ \ x = z - y \ x + y = z" hoelzl@44791: by (auto simp: ereal_eq_minus_iff) hoelzl@42844: hoelzl@44791: lemma ereal_less_minus_iff: hoelzl@44791: fixes x y z :: ereal hoelzl@42844: shows "x < z - y \ hoelzl@42844: (y = \ \ z = \ \ x \ \) \ hoelzl@42844: (y = -\ \ x \ \) \ hoelzl@42847: (\y\ \ \\ x + y < z)" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) auto hoelzl@42844: hoelzl@44791: lemma ereal_less_minus: hoelzl@44791: fixes x y z :: ereal hoelzl@42847: shows "\y\ \ \ \ x < z - y \ x + y < z" hoelzl@44791: by (auto simp: ereal_less_minus_iff) hoelzl@42844: hoelzl@44791: lemma ereal_le_minus_iff: hoelzl@44791: fixes x y z :: ereal hoelzl@42844: shows "x \ z - y \ hoelzl@42844: (y = \ \ z \ \ \ x = -\) \ hoelzl@42847: (\y\ \ \ \ x + y \ z)" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) auto hoelzl@42844: hoelzl@44791: lemma ereal_le_minus: hoelzl@44791: fixes x y z :: ereal hoelzl@42847: shows "\y\ \ \ \ x \ z - y \ x + y \ z" hoelzl@44791: by (auto simp: ereal_le_minus_iff) hoelzl@42844: hoelzl@44791: lemma ereal_minus_less_iff: hoelzl@44791: fixes x y z :: ereal hoelzl@42844: shows "x - y < z \ hoelzl@42844: y \ -\ \ (y = \ \ x \ \ \ z \ -\) \ hoelzl@42844: (y \ \ \ x < z + y)" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) auto hoelzl@42844: hoelzl@44791: lemma ereal_minus_less: hoelzl@44791: fixes x y z :: ereal hoelzl@42847: shows "\y\ \ \ \ x - y < z \ x < z + y" hoelzl@44791: by (auto simp: ereal_minus_less_iff) hoelzl@42844: hoelzl@44791: lemma ereal_minus_le_iff: hoelzl@44791: fixes x y z :: ereal hoelzl@42844: shows "x - y \ z \ hoelzl@42844: (y = -\ \ z = \) \ hoelzl@42844: (y = \ \ x = \ \ z = \) \ hoelzl@42847: (\y\ \ \ \ x \ z + y)" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) auto hoelzl@42844: hoelzl@44791: lemma ereal_minus_le: hoelzl@44791: fixes x y z :: ereal hoelzl@42847: shows "\y\ \ \ \ x - y \ z \ x \ z + y" hoelzl@44791: by (auto simp: ereal_minus_le_iff) hoelzl@42844: hoelzl@44791: lemma ereal_minus_eq_minus_iff: hoelzl@44791: fixes a b c :: ereal hoelzl@42844: shows "a - b = a - c \ hoelzl@42844: b = c \ a = \ \ (a = -\ \ b \ -\ \ c \ -\)" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) auto hoelzl@42844: hoelzl@44791: lemma ereal_add_le_add_iff: hoelzl@44794: fixes a b c :: ereal hoelzl@44794: shows "c + a \ c + b \ hoelzl@42844: a \ b \ c = \ \ (c = -\ \ a \ \ \ b \ \)" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) (simp_all add: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_mult_le_mult_iff: hoelzl@44794: fixes a b c :: ereal hoelzl@44794: shows "\c\ \ \ \ c * a \ c * b \ (0 < c \ a \ b) \ (c < 0 \ b \ a)" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) (simp_all add: mult_le_cancel_left) hoelzl@42844: hoelzl@44791: lemma ereal_minus_mono: hoelzl@44791: fixes A B C D :: ereal assumes "A \ B" "D \ C" hoelzl@42850: shows "A - C \ B - D" hoelzl@42850: using assms hoelzl@44791: by (cases rule: ereal3_cases[case_product ereal_cases, of A B C D]) simp_all hoelzl@42850: hoelzl@44791: lemma real_of_ereal_minus: hoelzl@44794: fixes a b :: ereal hoelzl@44794: shows "real (a - b) = (if \a\ = \ \ \b\ = \ then 0 else real a - real b)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42850: hoelzl@44791: lemma ereal_diff_positive: hoelzl@44791: fixes a b :: ereal shows "a \ b \ 0 \ b - a" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42850: hoelzl@44791: lemma ereal_between: hoelzl@44791: fixes x e :: ereal hoelzl@42847: assumes "\x\ \ \" "0 < e" hoelzl@42844: shows "x - e < x" "x < x + e" hoelzl@42844: using assms apply (cases x, cases e) apply auto hoelzl@42844: using assms by (cases x, cases e) auto hoelzl@42844: hoelzl@42844: subsubsection {* Division *} hoelzl@42844: hoelzl@44791: instantiation ereal :: inverse hoelzl@42844: begin hoelzl@42844: hoelzl@44791: function inverse_ereal where hoelzl@44791: "inverse (ereal r) = (if r = 0 then \ else ereal (inverse r))" | hoelzl@44794: "inverse (\::ereal) = 0" | hoelzl@44794: "inverse (-\::ereal) = 0" hoelzl@44791: by (auto intro: ereal_cases) hoelzl@42844: termination by (relation "{}") simp hoelzl@42844: hoelzl@44791: definition "x / y = x * inverse (y :: ereal)" hoelzl@42844: hoelzl@42844: instance proof qed hoelzl@42844: end hoelzl@42844: hoelzl@44791: lemma real_of_ereal_inverse[simp]: hoelzl@44791: fixes a :: ereal hoelzl@43791: shows "real (inverse a) = 1 / real a" hoelzl@43791: by (cases a) (auto simp: inverse_eq_divide) hoelzl@43791: hoelzl@44791: lemma ereal_inverse[simp]: hoelzl@44794: "inverse (0::ereal) = \" hoelzl@44791: "inverse (1::ereal) = 1" hoelzl@44791: by (simp_all add: one_ereal_def zero_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_divide[simp]: hoelzl@44791: "ereal r / ereal p = (if p = 0 then ereal r * \ else ereal (r / p))" hoelzl@44791: unfolding divide_ereal_def by (auto simp: divide_real_def) hoelzl@42844: hoelzl@44791: lemma ereal_divide_same[simp]: hoelzl@44794: fixes x :: ereal shows "x / x = (if \x\ = \ \ x = 0 then 0 else 1)" hoelzl@42844: by (cases x) hoelzl@44791: (simp_all add: divide_real_def divide_ereal_def one_ereal_def) hoelzl@42844: hoelzl@44791: lemma ereal_inv_inv[simp]: hoelzl@44794: fixes x :: ereal shows "inverse (inverse x) = (if x \ -\ then x else \)" hoelzl@42844: by (cases x) auto hoelzl@42844: hoelzl@44791: lemma ereal_inverse_minus[simp]: hoelzl@44794: fixes x :: ereal shows "inverse (- x) = (if x = 0 then \ else -inverse x)" hoelzl@42844: by (cases x) simp_all hoelzl@42844: hoelzl@44791: lemma ereal_uminus_divide[simp]: hoelzl@44791: fixes x y :: ereal shows "- x / y = - (x / y)" hoelzl@44791: unfolding divide_ereal_def by simp hoelzl@42844: hoelzl@44791: lemma ereal_divide_Infty[simp]: hoelzl@44794: fixes x :: ereal shows "x / \ = 0" "x / -\ = 0" hoelzl@44791: unfolding divide_ereal_def by simp_all hoelzl@42844: hoelzl@44791: lemma ereal_divide_one[simp]: hoelzl@44791: "x / 1 = (x::ereal)" hoelzl@44791: unfolding divide_ereal_def by simp hoelzl@42844: hoelzl@44791: lemma ereal_divide_ereal[simp]: hoelzl@44791: "\ / ereal r = (if 0 \ r then \ else -\)" hoelzl@44791: unfolding divide_ereal_def by simp hoelzl@42844: hoelzl@44791: lemma zero_le_divide_ereal[simp]: hoelzl@44791: fixes a :: ereal assumes "0 \ a" "0 \ b" hoelzl@42849: shows "0 \ a / b" hoelzl@44791: using assms by (cases rule: ereal2_cases[of a b]) (auto simp: zero_le_divide_iff) hoelzl@42849: hoelzl@44791: lemma ereal_le_divide_pos: hoelzl@44794: fixes x y z :: ereal shows "x > 0 \ x \ \ \ y \ z / x \ x * y \ z" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_divide_le_pos: hoelzl@44794: fixes x y z :: ereal shows "x > 0 \ x \ \ \ z / x \ y \ z \ x * y" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_le_divide_neg: hoelzl@44794: fixes x y z :: ereal shows "x < 0 \ x \ -\ \ y \ z / x \ z \ x * y" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_divide_le_neg: hoelzl@44794: fixes x y z :: ereal shows "x < 0 \ x \ -\ \ z / x \ y \ x * y \ z" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_inverse_antimono_strict: hoelzl@44791: fixes x y :: ereal hoelzl@42844: shows "0 \ x \ x < y \ inverse y < inverse x" hoelzl@44791: by (cases rule: ereal2_cases[of x y]) auto hoelzl@42844: hoelzl@44791: lemma ereal_inverse_antimono: hoelzl@44791: fixes x y :: ereal hoelzl@42844: shows "0 \ x \ x <= y \ inverse y <= inverse x" hoelzl@44791: by (cases rule: ereal2_cases[of x y]) auto hoelzl@42844: hoelzl@42844: lemma inverse_inverse_Pinfty_iff[simp]: hoelzl@44794: fixes x :: ereal shows "inverse x = \ \ x = 0" hoelzl@42844: by (cases x) auto hoelzl@42844: hoelzl@44791: lemma ereal_inverse_eq_0: hoelzl@44794: fixes x :: ereal shows "inverse x = 0 \ x = \ \ x = -\" hoelzl@42844: by (cases x) auto hoelzl@42844: hoelzl@44791: lemma ereal_0_gt_inverse: hoelzl@44791: fixes x :: ereal shows "0 < inverse x \ x \ \ \ 0 \ x" hoelzl@42850: by (cases x) auto hoelzl@42850: hoelzl@44791: lemma ereal_mult_less_right: hoelzl@44794: fixes a b c :: ereal hoelzl@42844: assumes "b * a < c * a" "0 < a" "a < \" hoelzl@42844: shows "b < c" hoelzl@42844: using assms hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@42844: (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff) hoelzl@42844: hoelzl@44791: lemma ereal_power_divide: hoelzl@44794: fixes x y :: ereal shows "y \ 0 \ (x / y) ^ n = x^n / y^n" hoelzl@44791: by (cases rule: ereal2_cases[of x y]) hoelzl@44791: (auto simp: one_ereal_def zero_ereal_def power_divide not_le hoelzl@42850: power_less_zero_eq zero_le_power_iff) hoelzl@42850: hoelzl@44791: lemma ereal_le_mult_one_interval: hoelzl@44791: fixes x y :: ereal hoelzl@42850: assumes y: "y \ -\" hoelzl@42850: assumes z: "\z. \ 0 < z ; z < 1 \ \ z * x \ y" hoelzl@42850: shows "x \ y" hoelzl@42850: proof (cases x) hoelzl@44791: case PInf with z[of "1 / 2"] show "x \ y" by (simp add: one_ereal_def) hoelzl@42850: next hoelzl@42850: case (real r) note r = this hoelzl@42850: show "x \ y" hoelzl@42850: proof (cases y) hoelzl@42850: case (real p) note p = this hoelzl@42850: have "r \ p" hoelzl@42850: proof (rule field_le_mult_one_interval) hoelzl@42850: fix z :: real assume "0 < z" and "z < 1" hoelzl@44791: with z[of "ereal z"] hoelzl@44791: show "z * r \ p" using p r by (auto simp: zero_le_mult_iff one_ereal_def) hoelzl@42850: qed hoelzl@42850: then show "x \ y" using p r by simp hoelzl@42850: qed (insert y, simp_all) hoelzl@42850: qed simp hoelzl@42849: hoelzl@42844: subsection "Complete lattice" hoelzl@42844: hoelzl@44791: instantiation ereal :: lattice hoelzl@42844: begin hoelzl@44791: definition [simp]: "sup x y = (max x y :: ereal)" hoelzl@44791: definition [simp]: "inf x y = (min x y :: ereal)" hoelzl@42844: instance proof qed simp_all hoelzl@42844: end hoelzl@42844: hoelzl@44791: instantiation ereal :: complete_lattice hoelzl@42844: begin hoelzl@42844: hoelzl@44794: definition "bot = (-\::ereal)" hoelzl@44794: definition "top = (\::ereal)" hoelzl@42844: hoelzl@44794: definition "Sup S = (LEAST z. \x\S. x \ z :: ereal)" hoelzl@44794: definition "Inf S = (GREATEST z. \x\S. z \ x :: ereal)" hoelzl@42844: hoelzl@44791: lemma ereal_complete_Sup: hoelzl@44791: fixes S :: "ereal set" assumes "S \ {}" hoelzl@42844: shows "\x. (\y\S. y \ x) \ (\z. (\y\S. y \ z) \ x \ z)" hoelzl@42844: proof cases hoelzl@44791: assume "\x. \a\S. a \ ereal x" hoelzl@44791: then obtain y where y: "\a. a\S \ a \ ereal y" by auto hoelzl@42844: then have "\ \ S" by force hoelzl@42844: show ?thesis hoelzl@42844: proof cases hoelzl@42844: assume "S = {-\}" hoelzl@42844: then show ?thesis by (auto intro!: exI[of _ "-\"]) hoelzl@42844: next hoelzl@42844: assume "S \ {-\}" hoelzl@42844: with `S \ {}` `\ \ S` obtain x where "x \ S - {-\}" "x \ \" by auto hoelzl@42844: with y `\ \ S` have "\z\real ` (S - {-\}). z \ y" hoelzl@44791: by (auto simp: real_of_ereal_ord_simps) hoelzl@42844: with reals_complete2[of "real ` (S - {-\})"] `x \ S - {-\}` hoelzl@42844: obtain s where s: hoelzl@42844: "\y\S - {-\}. real y \ s" "\z. (\y\S - {-\}. real y \ z) \ s \ z" hoelzl@42844: by auto hoelzl@42844: show ?thesis hoelzl@44791: proof (safe intro!: exI[of _ "ereal s"]) hoelzl@44791: fix z assume "z \ S" with `\ \ S` show "z \ ereal s" hoelzl@42844: proof (cases z) hoelzl@42844: case (real r) hoelzl@42844: then show ?thesis hoelzl@44791: using s(1)[rule_format, of z] `z \ S` `z = ereal r` by auto hoelzl@42844: qed auto hoelzl@42844: next hoelzl@42844: fix z assume *: "\y\S. y \ z" hoelzl@44791: with `S \ {-\}` `S \ {}` show "ereal s \ z" hoelzl@42844: proof (cases z) hoelzl@42844: case (real u) hoelzl@42844: with * have "s \ u" hoelzl@44791: by (intro s(2)[of u]) (auto simp: real_of_ereal_ord_simps) hoelzl@42844: then show ?thesis using real by simp hoelzl@42844: qed auto hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: next hoelzl@44791: assume *: "\ (\x. \a\S. a \ ereal x)" hoelzl@42844: show ?thesis hoelzl@42844: proof (safe intro!: exI[of _ \]) hoelzl@42844: fix y assume **: "\z\S. z \ y" hoelzl@42844: with * show "\ \ y" hoelzl@42844: proof (cases y) hoelzl@42844: case MInf with * ** show ?thesis by (force simp: not_le) hoelzl@42844: qed auto hoelzl@42844: qed simp hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_complete_Inf: hoelzl@44791: fixes S :: "ereal set" assumes "S ~= {}" hoelzl@42844: shows "EX x. (ALL y:S. x <= y) & (ALL z. (ALL y:S. z <= y) --> z <= x)" hoelzl@42844: proof- hoelzl@42844: def S1 == "uminus ` S" hoelzl@42844: hence "S1 ~= {}" using assms by auto hoelzl@42844: from this obtain x where x_def: "(ALL y:S1. y <= x) & (ALL z. (ALL y:S1. y <= z) --> x <= z)" hoelzl@44791: using ereal_complete_Sup[of S1] by auto hoelzl@42844: { fix z assume "ALL y:S. z <= y" hoelzl@42844: hence "ALL y:S1. y <= -z" unfolding S1_def by auto hoelzl@42844: hence "x <= -z" using x_def by auto hoelzl@42844: hence "z <= -x" hoelzl@44791: apply (subst ereal_uminus_uminus[symmetric]) hoelzl@44791: unfolding ereal_minus_le_minus . } hoelzl@42844: moreover have "(ALL y:S. -x <= y)" hoelzl@42844: using x_def unfolding S1_def hoelzl@42844: apply simp hoelzl@44791: apply (subst (3) ereal_uminus_uminus[symmetric]) hoelzl@44791: unfolding ereal_minus_le_minus by simp hoelzl@42844: ultimately show ?thesis by auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_complete_uminus_eq: hoelzl@44791: fixes S :: "ereal set" hoelzl@42844: shows "(\y\uminus`S. y \ x) \ (\z. (\y\uminus`S. y \ z) \ x \ z) hoelzl@42844: \ (\y\S. -x \ y) \ (\z. (\y\S. z \ y) \ z \ -x)" hoelzl@44791: by simp (metis ereal_minus_le_minus ereal_uminus_uminus) hoelzl@42844: hoelzl@44791: lemma ereal_Sup_uminus_image_eq: hoelzl@44791: fixes S :: "ereal set" hoelzl@42844: shows "Sup (uminus ` S) = - Inf S" hoelzl@42844: proof cases hoelzl@42844: assume "S = {}" hoelzl@44791: moreover have "(THE x. All (op \ x)) = (-\::ereal)" hoelzl@44791: by (rule the_equality) (auto intro!: ereal_bot) hoelzl@44791: moreover have "(SOME x. \y. y \ x) = (\::ereal)" hoelzl@44791: by (rule some_equality) (auto intro!: ereal_top) hoelzl@44791: ultimately show ?thesis unfolding Inf_ereal_def Sup_ereal_def hoelzl@42844: Least_def Greatest_def GreatestM_def by simp hoelzl@42844: next hoelzl@42844: assume "S \ {}" hoelzl@44791: with ereal_complete_Sup[of "uminus`S"] hoelzl@42844: obtain x where x: "(\y\S. -x \ y) \ (\z. (\y\S. z \ y) \ z \ -x)" hoelzl@44791: unfolding ereal_complete_uminus_eq by auto hoelzl@42844: show "Sup (uminus ` S) = - Inf S" hoelzl@44791: unfolding Inf_ereal_def Greatest_def GreatestM_def hoelzl@42844: proof (intro someI2[of _ _ "\x. Sup (uminus`S) = - x"]) hoelzl@42844: show "(\y\S. -x \ y) \ (\y. (\z\S. y \ z) \ y \ -x)" hoelzl@42844: using x . hoelzl@42844: fix x' assume "(\y\S. x' \ y) \ (\y. (\z\S. y \ z) \ y \ x')" hoelzl@42844: then have "(\y\uminus`S. y \ - x') \ (\y. (\z\uminus`S. z \ y) \ - x' \ y)" hoelzl@44791: unfolding ereal_complete_uminus_eq by simp hoelzl@42844: then show "Sup (uminus ` S) = -x'" hoelzl@44791: unfolding Sup_ereal_def ereal_uminus_eq_iff hoelzl@42844: by (intro Least_equality) auto hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: hoelzl@42844: instance hoelzl@42844: proof hoelzl@44791: { fix x :: ereal and A hoelzl@44791: show "bot <= x" by (cases x) (simp_all add: bot_ereal_def) hoelzl@44791: show "x <= top" by (simp add: top_ereal_def) } hoelzl@42844: hoelzl@44791: { fix x :: ereal and A assume "x : A" hoelzl@44791: with ereal_complete_Sup[of A] hoelzl@42844: obtain s where s: "\y\A. y <= s" "\z. (\y\A. y <= z) \ s <= z" by auto hoelzl@42844: hence "x <= s" using `x : A` by auto hoelzl@44791: also have "... = Sup A" using s unfolding Sup_ereal_def hoelzl@42844: by (auto intro!: Least_equality[symmetric]) hoelzl@42844: finally show "x <= Sup A" . } hoelzl@42844: note le_Sup = this hoelzl@42844: hoelzl@44791: { fix x :: ereal and A assume *: "!!z. (z : A ==> z <= x)" hoelzl@42844: show "Sup A <= x" hoelzl@42844: proof (cases "A = {}") hoelzl@42844: case True hoelzl@44791: hence "Sup A = -\" unfolding Sup_ereal_def hoelzl@42844: by (auto intro!: Least_equality) hoelzl@42844: thus "Sup A <= x" by simp hoelzl@42844: next hoelzl@42844: case False hoelzl@44791: with ereal_complete_Sup[of A] hoelzl@42844: obtain s where s: "\y\A. y <= s" "\z. (\y\A. y <= z) \ s <= z" by auto hoelzl@42844: hence "Sup A = s" hoelzl@44791: unfolding Sup_ereal_def by (auto intro!: Least_equality) hoelzl@42844: also have "s <= x" using * s by auto hoelzl@42844: finally show "Sup A <= x" . hoelzl@42844: qed } hoelzl@42844: note Sup_le = this hoelzl@42844: hoelzl@44791: { fix x :: ereal and A assume "x \ A" hoelzl@42844: with le_Sup[of "-x" "uminus`A"] show "Inf A \ x" hoelzl@44791: unfolding ereal_Sup_uminus_image_eq by simp } hoelzl@42844: hoelzl@44791: { fix x :: ereal and A assume *: "!!z. (z : A ==> x <= z)" hoelzl@42844: with Sup_le[of "uminus`A" "-x"] show "x \ Inf A" hoelzl@44791: unfolding ereal_Sup_uminus_image_eq by force } hoelzl@42844: qed haftmann@44812: hoelzl@42844: end hoelzl@42844: haftmann@44812: instance ereal :: complete_linorder .. haftmann@44812: hoelzl@44791: lemma ereal_SUPR_uminus: hoelzl@44791: fixes f :: "'a => ereal" hoelzl@42844: shows "(SUP i : R. -(f i)) = -(INF i : R. f i)" hoelzl@42844: unfolding SUPR_def INFI_def hoelzl@44791: using ereal_Sup_uminus_image_eq[of "f`R"] hoelzl@42844: by (simp add: image_image) hoelzl@42844: hoelzl@44791: lemma ereal_INFI_uminus: hoelzl@44791: fixes f :: "'a => ereal" hoelzl@42844: shows "(INF i : R. -(f i)) = -(SUP i : R. f i)" hoelzl@44791: using ereal_SUPR_uminus[of _ "\x. - f x"] by simp hoelzl@42844: hoelzl@44791: lemma ereal_Inf_uminus_image_eq: "Inf (uminus ` S) = - Sup (S::ereal set)" hoelzl@44791: using ereal_Sup_uminus_image_eq[of "uminus ` S"] by (simp add: image_image) hoelzl@42850: hoelzl@44791: lemma ereal_inj_on_uminus[intro, simp]: "inj_on uminus (A :: ereal set)" hoelzl@42844: by (auto intro!: inj_onI) hoelzl@42844: hoelzl@44791: lemma ereal_image_uminus_shift: hoelzl@44791: fixes X Y :: "ereal set" shows "uminus ` X = Y \ X = uminus ` Y" hoelzl@42844: proof hoelzl@42844: assume "uminus ` X = Y" hoelzl@42844: then have "uminus ` uminus ` X = uminus ` Y" hoelzl@42844: by (simp add: inj_image_eq_iff) hoelzl@42844: then show "X = uminus ` Y" by (simp add: image_image) hoelzl@42844: qed (simp add: image_image) hoelzl@42844: hoelzl@44791: lemma Inf_ereal_iff: hoelzl@44791: fixes z :: ereal hoelzl@42844: shows "(!!x. x:X ==> z <= x) ==> (EX x:X. x Inf X < y" hoelzl@42844: by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear hoelzl@42844: order_less_le_trans) hoelzl@42844: hoelzl@42844: lemma Sup_eq_MInfty: hoelzl@44791: fixes S :: "ereal set" shows "Sup S = -\ \ S = {} \ S = {-\}" hoelzl@42844: proof hoelzl@42844: assume a: "Sup S = -\" hoelzl@42844: with complete_lattice_class.Sup_upper[of _ S] hoelzl@42844: show "S={} \ S={-\}" by auto hoelzl@42844: next hoelzl@42844: assume "S={} \ S={-\}" then show "Sup S = -\" hoelzl@44791: unfolding Sup_ereal_def by (auto intro!: Least_equality) hoelzl@42844: qed hoelzl@42844: hoelzl@42844: lemma Inf_eq_PInfty: hoelzl@44791: fixes S :: "ereal set" shows "Inf S = \ \ S = {} \ S = {\}" hoelzl@42844: using Sup_eq_MInfty[of "uminus`S"] hoelzl@44791: unfolding ereal_Sup_uminus_image_eq ereal_image_uminus_shift by simp hoelzl@42844: hoelzl@44794: lemma Inf_eq_MInfty: hoelzl@44794: fixes S :: "ereal set" shows "-\ \ S \ Inf S = -\" hoelzl@44791: unfolding Inf_ereal_def hoelzl@42844: by (auto intro!: Greatest_equality) hoelzl@42844: hoelzl@44794: lemma Sup_eq_PInfty: hoelzl@44794: fixes S :: "ereal set" shows "\ \ S \ Sup S = \" hoelzl@44791: unfolding Sup_ereal_def hoelzl@42844: by (auto intro!: Least_equality) hoelzl@42844: hoelzl@44791: lemma ereal_SUPI: hoelzl@44791: fixes x :: ereal hoelzl@42844: assumes "!!i. i : A ==> f i <= x" hoelzl@42844: assumes "!!y. (!!i. i : A ==> f i <= y) ==> x <= y" hoelzl@42844: shows "(SUP i:A. f i) = x" hoelzl@44791: unfolding SUPR_def Sup_ereal_def hoelzl@42844: using assms by (auto intro!: Least_equality) hoelzl@42844: hoelzl@44791: lemma ereal_INFI: hoelzl@44791: fixes x :: ereal hoelzl@42844: assumes "!!i. i : A ==> f i >= x" hoelzl@42844: assumes "!!y. (!!i. i : A ==> f i >= y) ==> x >= y" hoelzl@42844: shows "(INF i:A. f i) = x" hoelzl@44791: unfolding INFI_def Inf_ereal_def hoelzl@42844: using assms by (auto intro!: Greatest_equality) hoelzl@42844: hoelzl@44791: lemma Sup_ereal_close: hoelzl@44791: fixes e :: ereal hoelzl@42847: assumes "0 < e" and S: "\Sup S\ \ \" "S \ {}" hoelzl@42844: shows "\x\S. Sup S - e < x" hoelzl@42847: using assms by (cases e) (auto intro!: less_Sup_iff[THEN iffD1]) hoelzl@42844: hoelzl@44791: lemma Inf_ereal_close: hoelzl@44791: fixes e :: ereal assumes "\Inf X\ \ \" "0 < e" hoelzl@42844: shows "\x\X. x < Inf X + e" hoelzl@42844: proof (rule Inf_less_iff[THEN iffD1]) hoelzl@42844: show "Inf X < Inf X + e" using assms hoelzl@42847: by (cases e) auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma SUP_nat_Infty: "(SUP i::nat. ereal (real i)) = \" hoelzl@42844: proof - hoelzl@44794: { fix x ::ereal assume "x \ \" hoelzl@44791: then have "\k::nat. x < ereal (real k)" hoelzl@42844: proof (cases x) hoelzl@42844: case MInf then show ?thesis by (intro exI[of _ 0]) auto hoelzl@42844: next hoelzl@42844: case (real r) hoelzl@42844: moreover obtain k :: nat where "r < real k" hoelzl@42844: using ex_less_of_nat by (auto simp: real_eq_of_nat) hoelzl@42844: ultimately show ?thesis by auto hoelzl@42844: qed simp } hoelzl@42844: then show ?thesis hoelzl@44791: using SUP_eq_top_iff[of UNIV "\n::nat. ereal (real n)"] hoelzl@44791: by (auto simp: top_ereal_def) hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_le_Sup: hoelzl@44791: fixes x :: ereal hoelzl@42844: shows "(x <= (SUP i:A. f i)) <-> (ALL y. y < x --> (EX i. i : A & y <= f i))" hoelzl@42844: (is "?lhs <-> ?rhs") hoelzl@42844: proof- hoelzl@42844: { assume "?rhs" hoelzl@42844: { assume "~(x <= (SUP i:A. f i))" hence "(SUP i:A. f i) (ALL y. x < y --> (EX i. i : A & f i <= y))" hoelzl@42844: (is "?lhs <-> ?rhs") hoelzl@42844: proof- hoelzl@42844: { assume "?rhs" hoelzl@42844: { assume "~((INF i:A. f i) <= x)" hence "x < (INF i:A. f i)" by (simp add: not_le) hoelzl@44791: from this obtain y where y_def: "x x" by auto hoelzl@42844: hence "(INF i:A. f i) >= x" apply (subst le_INFI) by auto hoelzl@42844: thus False using assms by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42844: lemma same_INF: hoelzl@42844: assumes "ALL e:A. f e = g e" hoelzl@42844: shows "(INF e:A. f e) = (INF e:A. g e)" hoelzl@42844: proof- hoelzl@42844: have "f ` A = g ` A" unfolding image_def using assms by auto hoelzl@42844: thus ?thesis unfolding INFI_def by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42844: lemma same_SUP: hoelzl@42844: assumes "ALL e:A. f e = g e" hoelzl@42844: shows "(SUP e:A. f e) = (SUP e:A. g e)" hoelzl@42844: proof- hoelzl@42844: have "f ` A = g ` A" unfolding image_def using assms by auto hoelzl@42844: thus ?thesis unfolding SUPR_def by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42850: lemma SUPR_eq: hoelzl@42850: assumes "\i\A. \j\B. f i \ g j" hoelzl@42850: assumes "\j\B. \i\A. g j \ f i" hoelzl@42850: shows "(SUP i:A. f i) = (SUP j:B. g j)" hoelzl@42850: proof (intro antisym) hoelzl@42850: show "(SUP i:A. f i) \ (SUP j:B. g j)" hoelzl@42851: using assms by (metis SUP_leI le_SUPI2) hoelzl@42850: show "(SUP i:B. g i) \ (SUP j:A. f j)" hoelzl@42851: using assms by (metis SUP_leI le_SUPI2) hoelzl@42850: qed hoelzl@42850: hoelzl@44791: lemma SUP_ereal_le_addI: hoelzl@44794: fixes f :: "'i \ ereal" hoelzl@42849: assumes "\i. f i + y \ z" and "y \ -\" hoelzl@42849: shows "SUPR UNIV f + y \ z" hoelzl@42849: proof (cases y) hoelzl@42849: case (real r) hoelzl@44791: then have "\i. f i \ z - y" using assms by (simp add: ereal_le_minus_iff) hoelzl@42849: then have "SUPR UNIV f \ z - y" by (rule SUP_leI) hoelzl@44791: then show ?thesis using real by (simp add: ereal_le_minus_iff) hoelzl@42849: qed (insert assms, auto) hoelzl@42849: hoelzl@44791: lemma SUPR_ereal_add: hoelzl@44791: fixes f g :: "nat \ ereal" hoelzl@42850: assumes "incseq f" "incseq g" and pos: "\i. f i \ -\" "\i. g i \ -\" hoelzl@42849: shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g" hoelzl@44791: proof (rule ereal_SUPI) hoelzl@42849: fix y assume *: "\i. i \ UNIV \ f i + g i \ y" hoelzl@42849: have f: "SUPR UNIV f \ -\" using pos hoelzl@42849: unfolding SUPR_def Sup_eq_MInfty by (auto dest: image_eqD) hoelzl@42849: { fix j hoelzl@42849: { fix i hoelzl@42849: have "f i + g j \ f i + g (max i j)" hoelzl@42849: using `incseq g`[THEN incseqD] by (rule add_left_mono) auto hoelzl@42849: also have "\ \ f (max i j) + g (max i j)" hoelzl@42849: using `incseq f`[THEN incseqD] by (rule add_right_mono) auto hoelzl@42849: also have "\ \ y" using * by auto hoelzl@42849: finally have "f i + g j \ y" . } hoelzl@42849: then have "SUPR UNIV f + g j \ y" hoelzl@44791: using assms(4)[of j] by (intro SUP_ereal_le_addI) auto hoelzl@42849: then have "g j + SUPR UNIV f \ y" by (simp add: ac_simps) } hoelzl@42849: then have "SUPR UNIV g + SUPR UNIV f \ y" hoelzl@44791: using f by (rule SUP_ereal_le_addI) hoelzl@42849: then show "SUPR UNIV f + SUPR UNIV g \ y" by (simp add: ac_simps) hoelzl@42849: qed (auto intro!: add_mono le_SUPI) hoelzl@42849: hoelzl@44791: lemma SUPR_ereal_add_pos: hoelzl@44791: fixes f g :: "nat \ ereal" hoelzl@42850: assumes inc: "incseq f" "incseq g" and pos: "\i. 0 \ f i" "\i. 0 \ g i" hoelzl@42850: shows "(SUP i. f i + g i) = SUPR UNIV f + SUPR UNIV g" hoelzl@44791: proof (intro SUPR_ereal_add inc) hoelzl@42850: fix i show "f i \ -\" "g i \ -\" using pos[of i] by auto hoelzl@42850: qed hoelzl@42850: hoelzl@44791: lemma SUPR_ereal_setsum: hoelzl@44791: fixes f g :: "'a \ nat \ ereal" hoelzl@42850: assumes "\n. n \ A \ incseq (f n)" and pos: "\n i. n \ A \ 0 \ f n i" hoelzl@42850: shows "(SUP i. \n\A. f n i) = (\n\A. SUPR UNIV (f n))" hoelzl@42850: proof cases hoelzl@42850: assume "finite A" then show ?thesis using assms hoelzl@44791: by induct (auto simp: incseq_setsumI2 setsum_nonneg SUPR_ereal_add_pos) hoelzl@42850: qed simp hoelzl@42850: hoelzl@44791: lemma SUPR_ereal_cmult: hoelzl@44791: fixes f :: "nat \ ereal" assumes "\i. 0 \ f i" "0 \ c" hoelzl@42849: shows "(SUP i. c * f i) = c * SUPR UNIV f" hoelzl@44791: proof (rule ereal_SUPI) hoelzl@42849: fix i have "f i \ SUPR UNIV f" by (rule le_SUPI) auto hoelzl@42849: then show "c * f i \ c * SUPR UNIV f" hoelzl@44791: using `0 \ c` by (rule ereal_mult_left_mono) hoelzl@42849: next hoelzl@42849: fix y assume *: "\i. i \ UNIV \ c * f i \ y" hoelzl@42849: show "c * SUPR UNIV f \ y" hoelzl@42849: proof cases hoelzl@42849: assume c: "0 < c \ c \ \" hoelzl@42849: with * have "SUPR UNIV f \ y / c" hoelzl@44791: by (intro SUP_leI) (auto simp: ereal_le_divide_pos) hoelzl@42849: with c show ?thesis hoelzl@44791: by (auto simp: ereal_le_divide_pos) hoelzl@42849: next hoelzl@42849: { assume "c = \" have ?thesis hoelzl@42849: proof cases hoelzl@42849: assume "\i. f i = 0" hoelzl@42849: moreover then have "range f = {0}" by auto hoelzl@42849: ultimately show "c * SUPR UNIV f \ y" using * by (auto simp: SUPR_def) hoelzl@42849: next hoelzl@42849: assume "\ (\i. f i = 0)" hoelzl@42849: then obtain i where "f i \ 0" by auto hoelzl@42849: with *[of i] `c = \` `0 \ f i` show ?thesis by (auto split: split_if_asm) hoelzl@42849: qed } hoelzl@42849: moreover assume "\ (0 < c \ c \ \)" hoelzl@42849: ultimately show ?thesis using * `0 \ c` by auto hoelzl@42849: qed hoelzl@42849: qed hoelzl@42849: hoelzl@42850: lemma SUP_PInfty: hoelzl@44791: fixes f :: "'a \ ereal" hoelzl@44791: assumes "\n::nat. \i\A. ereal (real n) \ f i" hoelzl@42850: shows "(SUP i:A. f i) = \" hoelzl@44791: unfolding SUPR_def Sup_eq_top_iff[where 'a=ereal, unfolded top_ereal_def] hoelzl@42850: apply simp hoelzl@42850: proof safe hoelzl@44794: fix x :: ereal assume "x \ \" hoelzl@42850: show "\i\A. x < f i" hoelzl@42850: proof (cases x) hoelzl@42850: case PInf with `x \ \` show ?thesis by simp hoelzl@42850: next hoelzl@42850: case MInf with assms[of "0"] show ?thesis by force hoelzl@42850: next hoelzl@42850: case (real r) hoelzl@44791: with less_PInf_Ex_of_nat[of x] obtain n :: nat where "x < ereal (real n)" by auto hoelzl@42850: moreover from assms[of n] guess i .. hoelzl@42850: ultimately show ?thesis hoelzl@42850: by (auto intro!: bexI[of _ i]) hoelzl@42850: qed hoelzl@42850: qed hoelzl@42850: hoelzl@42850: lemma Sup_countable_SUPR: hoelzl@42850: assumes "A \ {}" hoelzl@44791: shows "\f::nat \ ereal. range f \ A \ Sup A = SUPR UNIV f" hoelzl@42850: proof (cases "Sup A") hoelzl@42850: case (real r) hoelzl@44791: have "\n::nat. \x. x \ A \ Sup A < x + 1 / ereal (real n)" hoelzl@42850: proof hoelzl@44791: fix n ::nat have "\x\A. Sup A - 1 / ereal (real n) < x" hoelzl@44791: using assms real by (intro Sup_ereal_close) (auto simp: one_ereal_def) hoelzl@42850: then guess x .. hoelzl@44791: then show "\x. x \ A \ Sup A < x + 1 / ereal (real n)" hoelzl@44791: by (auto intro!: exI[of _ x] simp: ereal_minus_less_iff) hoelzl@42850: qed hoelzl@42850: from choice[OF this] guess f .. note f = this hoelzl@42850: have "SUPR UNIV f = Sup A" hoelzl@44791: proof (rule ereal_SUPI) hoelzl@42850: fix i show "f i \ Sup A" using f hoelzl@42850: by (auto intro!: complete_lattice_class.Sup_upper) hoelzl@42850: next hoelzl@42850: fix y assume bound: "\i. i \ UNIV \ f i \ y" hoelzl@42850: show "Sup A \ y" hoelzl@44791: proof (rule ereal_le_epsilon, intro allI impI) hoelzl@44791: fix e :: ereal assume "0 < e" hoelzl@42850: show "Sup A \ y + e" hoelzl@42850: proof (cases e) hoelzl@42850: case (real r) hoelzl@42850: hence "0 < r" using `0 < e` by auto hoelzl@42850: then obtain n ::nat where *: "1 / real n < r" "0 < n" hoelzl@42850: using ex_inverse_of_nat_less by (auto simp: real_eq_of_nat inverse_eq_divide) hoelzl@44791: have "Sup A \ f n + 1 / ereal (real n)" using f[THEN spec, of n] by auto hoelzl@44791: also have "1 / ereal (real n) \ e" using real * by (auto simp: one_ereal_def ) hoelzl@44791: with bound have "f n + 1 / ereal (real n) \ y + e" by (rule add_mono) simp hoelzl@42850: finally show "Sup A \ y + e" . hoelzl@42850: qed (insert `0 < e`, auto) hoelzl@42850: qed hoelzl@42850: qed hoelzl@42850: with f show ?thesis by (auto intro!: exI[of _ f]) hoelzl@42850: next hoelzl@42850: case PInf hoelzl@42850: from `A \ {}` obtain x where "x \ A" by auto hoelzl@42850: show ?thesis hoelzl@42850: proof cases hoelzl@42850: assume "\ \ A" hoelzl@42850: moreover then have "\ \ Sup A" by (intro complete_lattice_class.Sup_upper) hoelzl@42850: ultimately show ?thesis by (auto intro!: exI[of _ "\x. \"]) hoelzl@42850: next hoelzl@42850: assume "\ \ A" hoelzl@42850: have "\x\A. 0 \ x" hoelzl@44791: by (metis Infty_neq_0 PInf complete_lattice_class.Sup_least ereal_infty_less_eq2 linorder_linear) hoelzl@42850: then obtain x where "x \ A" "0 \ x" by auto hoelzl@44791: have "\n::nat. \f. f \ A \ x + ereal (real n) \ f" hoelzl@42850: proof (rule ccontr) hoelzl@42850: assume "\ ?thesis" hoelzl@44791: then have "\n::nat. Sup A \ x + ereal (real n)" hoelzl@42850: by (simp add: Sup_le_iff not_le less_imp_le Ball_def) (metis less_imp_le) hoelzl@42850: then show False using `x \ A` `\ \ A` PInf hoelzl@42850: by(cases x) auto hoelzl@42850: qed hoelzl@42850: from choice[OF this] guess f .. note f = this hoelzl@42850: have "SUPR UNIV f = \" hoelzl@42850: proof (rule SUP_PInfty) hoelzl@44791: fix n :: nat show "\i\UNIV. ereal (real n) \ f i" hoelzl@42850: using f[THEN spec, of n] `0 \ x` hoelzl@44791: by (cases rule: ereal2_cases[of "f n" x]) (auto intro!: exI[of _ n]) hoelzl@42850: qed hoelzl@42850: then show ?thesis using f PInf by (auto intro!: exI[of _ f]) hoelzl@42850: qed hoelzl@42850: next hoelzl@42850: case MInf hoelzl@42850: with `A \ {}` have "A = {-\}" by (auto simp: Sup_eq_MInfty) hoelzl@42850: then show ?thesis using MInf by (auto intro!: exI[of _ "\x. -\"]) hoelzl@42850: qed hoelzl@42850: hoelzl@42850: lemma SUPR_countable_SUPR: hoelzl@44791: "A \ {} \ \f::nat \ ereal. range f \ g`A \ SUPR A g = SUPR UNIV f" hoelzl@42850: using Sup_countable_SUPR[of "g`A"] by (auto simp: SUPR_def) hoelzl@42850: hoelzl@44791: lemma Sup_ereal_cadd: hoelzl@44791: fixes A :: "ereal set" assumes "A \ {}" and "a \ -\" hoelzl@42850: shows "Sup ((\x. a + x) ` A) = a + Sup A" hoelzl@42850: proof (rule antisym) hoelzl@44791: have *: "\a::ereal. \A. Sup ((\x. a + x) ` A) \ a + Sup A" hoelzl@42850: by (auto intro!: add_mono complete_lattice_class.Sup_least complete_lattice_class.Sup_upper) hoelzl@42850: then show "Sup ((\x. a + x) ` A) \ a + Sup A" . hoelzl@42850: show "a + Sup A \ Sup ((\x. a + x) ` A)" hoelzl@42850: proof (cases a) hoelzl@42850: case PInf with `A \ {}` show ?thesis by (auto simp: image_constant) hoelzl@42850: next hoelzl@42850: case (real r) hoelzl@42850: then have **: "op + (- a) ` op + a ` A = A" hoelzl@44791: by (auto simp: image_iff ac_simps zero_ereal_def[symmetric]) hoelzl@42850: from *[of "-a" "(\x. a + x) ` A"] real show ?thesis unfolding ** hoelzl@44791: by (cases rule: ereal2_cases[of "Sup A" "Sup (op + a ` A)"]) auto hoelzl@42850: qed (insert `a \ -\`, auto) hoelzl@42850: qed hoelzl@42850: hoelzl@44791: lemma Sup_ereal_cminus: hoelzl@44791: fixes A :: "ereal set" assumes "A \ {}" and "a \ -\" hoelzl@42850: shows "Sup ((\x. a - x) ` A) = a - Inf A" hoelzl@44791: using Sup_ereal_cadd[of "uminus ` A" a] assms hoelzl@44791: by (simp add: comp_def image_image minus_ereal_def hoelzl@44791: ereal_Sup_uminus_image_eq) hoelzl@42850: hoelzl@44791: lemma SUPR_ereal_cminus: hoelzl@44794: fixes f :: "'i \ ereal" hoelzl@42850: fixes A assumes "A \ {}" and "a \ -\" hoelzl@42850: shows "(SUP x:A. a - f x) = a - (INF x:A. f x)" hoelzl@44791: using Sup_ereal_cminus[of "f`A" a] assms hoelzl@42850: unfolding SUPR_def INFI_def image_image by auto hoelzl@42850: hoelzl@44791: lemma Inf_ereal_cminus: hoelzl@44791: fixes A :: "ereal set" assumes "A \ {}" and "\a\ \ \" hoelzl@42850: shows "Inf ((\x. a - x) ` A) = a - Sup A" hoelzl@42850: proof - hoelzl@42850: { fix x have "-a - -x = -(a - x)" using assms by (cases x) auto } hoelzl@42850: moreover then have "(\x. -a - x)`uminus`A = uminus ` (\x. a - x) ` A" hoelzl@42850: by (auto simp: image_image) hoelzl@42850: ultimately show ?thesis hoelzl@44791: using Sup_ereal_cminus[of "uminus ` A" "-a"] assms hoelzl@44791: by (auto simp add: ereal_Sup_uminus_image_eq ereal_Inf_uminus_image_eq) hoelzl@42850: qed hoelzl@42850: hoelzl@44791: lemma INFI_ereal_cminus: hoelzl@44794: fixes a :: ereal assumes "A \ {}" and "\a\ \ \" hoelzl@42850: shows "(INF x:A. a - f x) = a - (SUP x:A. f x)" hoelzl@44791: using Inf_ereal_cminus[of "f`A" a] assms hoelzl@42850: unfolding SUPR_def INFI_def image_image hoelzl@42850: by auto hoelzl@42850: hoelzl@44791: lemma uminus_ereal_add_uminus_uminus: hoelzl@44791: fixes a b :: ereal shows "a \ \ \ b \ \ \ - (- a + - b) = a + b" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@43791: hoelzl@44791: lemma INFI_ereal_add: hoelzl@44794: fixes f :: "nat \ ereal" hoelzl@43791: assumes "decseq f" "decseq g" and fin: "\i. f i \ \" "\i. g i \ \" hoelzl@43791: shows "(INF i. f i + g i) = INFI UNIV f + INFI UNIV g" hoelzl@43791: proof - hoelzl@43791: have INF_less: "(INF i. f i) < \" "(INF i. g i) < \" hoelzl@43791: using assms unfolding INF_less_iff by auto hoelzl@43791: { fix i from fin[of i] have "- ((- f i) + (- g i)) = f i + g i" hoelzl@44791: by (rule uminus_ereal_add_uminus_uminus) } hoelzl@43791: then have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))" hoelzl@43791: by simp hoelzl@43791: also have "\ = INFI UNIV f + INFI UNIV g" hoelzl@44791: unfolding ereal_INFI_uminus hoelzl@43791: using assms INF_less hoelzl@44791: by (subst SUPR_ereal_add) hoelzl@44791: (auto simp: ereal_SUPR_uminus intro!: uminus_ereal_add_uminus_uminus) hoelzl@43791: finally show ?thesis . hoelzl@43791: qed hoelzl@43791: hoelzl@44791: subsection "Limits on @{typ ereal}" hoelzl@42844: hoelzl@42844: subsubsection "Topological space" hoelzl@42844: hoelzl@44791: instantiation ereal :: topological_space hoelzl@42844: begin hoelzl@42844: hoelzl@44791: definition "open A \ open (ereal -` A) hoelzl@44791: \ (\ \ A \ (\x. {ereal x <..} \ A)) hoelzl@44791: \ (-\ \ A \ (\x. {.. A))" hoelzl@42844: hoelzl@44791: lemma open_PInfty: "open A \ \ \ A \ (\x. {ereal x<..} \ A)" hoelzl@44791: unfolding open_ereal_def by auto hoelzl@42844: hoelzl@44791: lemma open_MInfty: "open A \ -\ \ A \ (\x. {.. A)" hoelzl@44791: unfolding open_ereal_def by auto hoelzl@42844: hoelzl@44791: lemma open_PInfty2: assumes "open A" "\ \ A" obtains x where "{ereal x<..} \ A" hoelzl@42844: using open_PInfty[OF assms] by auto hoelzl@42844: hoelzl@44791: lemma open_MInfty2: assumes "open A" "-\ \ A" obtains x where "{.. A" hoelzl@42844: using open_MInfty[OF assms] by auto hoelzl@42844: hoelzl@44791: lemma ereal_openE: assumes "open A" obtains x y where hoelzl@44791: "open (ereal -` A)" hoelzl@44791: "\ \ A \ {ereal x<..} \ A" hoelzl@44791: "-\ \ A \ {.. A" hoelzl@44791: using assms open_ereal_def by auto hoelzl@42844: hoelzl@42844: instance hoelzl@42844: proof hoelzl@44791: let ?U = "UNIV::ereal set" hoelzl@44791: show "open ?U" unfolding open_ereal_def hoelzl@42846: by (auto intro!: exI[of _ 0]) hoelzl@42844: next hoelzl@44791: fix S T::"ereal set" assume "open S" and "open T" hoelzl@44791: from `open S`[THEN ereal_openE] guess xS yS . hoelzl@44791: moreover from `open T`[THEN ereal_openE] guess xT yT . hoelzl@42846: ultimately have hoelzl@44791: "open (ereal -` (S \ T))" hoelzl@44791: "\ \ S \ T \ {ereal (max xS xT) <..} \ S \ T" hoelzl@44791: "-\ \ S \ T \ {..< ereal (min yS yT)} \ S \ T" hoelzl@42846: by auto hoelzl@44791: then show "open (S Int T)" unfolding open_ereal_def by blast hoelzl@42844: next hoelzl@44791: fix K :: "ereal set set" assume "\S\K. open S" hoelzl@44791: then have *: "\S. \x y. S \ K \ open (ereal -` S) \ hoelzl@44791: (\ \ S \ {ereal x <..} \ S) \ (-\ \ S \ {..< ereal y} \ S)" hoelzl@44791: by (auto simp: open_ereal_def) hoelzl@44791: then show "open (Union K)" unfolding open_ereal_def hoelzl@42846: proof (intro conjI impI) hoelzl@44791: show "open (ereal -` \K)" hoelzl@42851: using *[THEN choice] by (auto simp: vimage_Union) hoelzl@42846: qed ((metis UnionE Union_upper subset_trans *)+) hoelzl@42844: qed hoelzl@42844: end hoelzl@42844: hoelzl@44791: lemma open_ereal: "open S \ open (ereal ` S)" hoelzl@44791: by (auto simp: inj_vimage_image_eq open_ereal_def) hoelzl@42847: hoelzl@44791: lemma open_ereal_vimage: "open S \ open (ereal -` S)" hoelzl@44791: unfolding open_ereal_def by auto hoelzl@42847: hoelzl@44791: lemma open_ereal_lessThan[intro, simp]: "open {..< a :: ereal}" hoelzl@42846: proof - hoelzl@44791: have "\x. ereal -` {..} = UNIV" "ereal -` {..< -\} = {}" by auto hoelzl@44791: then show ?thesis by (cases a) (auto simp: open_ereal_def) hoelzl@42846: qed hoelzl@42846: hoelzl@44791: lemma open_ereal_greaterThan[intro, simp]: hoelzl@44791: "open {a :: ereal <..}" hoelzl@42846: proof - hoelzl@44791: have "\x. ereal -` {ereal x<..} = {x<..}" hoelzl@44791: "ereal -` {\<..} = {}" "ereal -` {-\<..} = UNIV" by auto hoelzl@44791: then show ?thesis by (cases a) (auto simp: open_ereal_def) hoelzl@42846: qed hoelzl@42846: hoelzl@44791: lemma ereal_open_greaterThanLessThan[intro, simp]: "open {a::ereal <..< b}" hoelzl@42844: unfolding greaterThanLessThan_def by auto hoelzl@42844: hoelzl@44791: lemma closed_ereal_atLeast[simp, intro]: "closed {a :: ereal ..}" hoelzl@42844: proof - hoelzl@42844: have "- {a ..} = {..< a}" by auto hoelzl@42844: then show "closed {a ..}" hoelzl@44791: unfolding closed_def using open_ereal_lessThan by auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma closed_ereal_atMost[simp, intro]: "closed {.. b :: ereal}" hoelzl@42844: proof - hoelzl@42844: have "- {.. b} = {b <..}" by auto hoelzl@42844: then show "closed {.. b}" hoelzl@44791: unfolding closed_def using open_ereal_greaterThan by auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma closed_ereal_atLeastAtMost[simp, intro]: hoelzl@44791: shows "closed {a :: ereal .. b}" hoelzl@42844: unfolding atLeastAtMost_def by auto hoelzl@42844: hoelzl@44791: lemma closed_ereal_singleton: hoelzl@44791: "closed {a :: ereal}" hoelzl@44791: by (metis atLeastAtMost_singleton closed_ereal_atLeastAtMost) hoelzl@42844: hoelzl@44791: lemma ereal_open_cont_interval: hoelzl@44794: fixes S :: "ereal set" hoelzl@42847: assumes "open S" "x \ S" "\x\ \ \" hoelzl@42844: obtains e where "e>0" "{x-e <..< x+e} \ S" hoelzl@42844: proof- hoelzl@44791: from `open S` have "open (ereal -` S)" by (rule ereal_openE) hoelzl@44791: then obtain e where "0 < e" and e: "\y. dist y (real x) < e \ ereal y \ S" hoelzl@42851: using assms unfolding open_dist by force hoelzl@42846: show thesis hoelzl@42846: proof (intro that subsetI) hoelzl@44791: show "0 < ereal e" using `0 < e` by auto hoelzl@44791: fix y assume "y \ {x - ereal e<.. S" using e[of t] by auto hoelzl@42846: qed hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_open_cont_interval2: hoelzl@44794: fixes S :: "ereal set" hoelzl@42847: assumes "open S" "x \ S" and x: "\x\ \ \" hoelzl@42844: obtains a b where "a < x" "x < b" "{a <..< b} \ S" hoelzl@42844: proof- hoelzl@44791: guess e using ereal_open_cont_interval[OF assms] . hoelzl@44791: with that[of "x-e" "x+e"] ereal_between[OF x, of e] hoelzl@42844: show thesis by auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: instance ereal :: t2_space hoelzl@42844: proof hoelzl@44791: fix x y :: ereal assume "x ~= y" hoelzl@44791: let "?P x (y::ereal)" = "EX U V. open U & open V & x : U & y : V & U Int V = {}" hoelzl@42844: hoelzl@44791: { fix x y :: ereal assume "x < y" hoelzl@44791: from ereal_dense[OF this] obtain z where z: "x < z" "z < y" by auto hoelzl@42844: have "?P x y" hoelzl@42844: apply (rule exI[of _ "{..n. ereal (f n)) ---> ereal x) net \ (f ---> x) net" (is "?l = ?r") hoelzl@42844: proof (intro iffI topological_tendstoI) hoelzl@42844: fix S assume "?l" "open S" "x \ S" hoelzl@42844: then show "eventually (\x. f x \ S) net" hoelzl@44791: using `?l`[THEN topological_tendstoD, OF open_ereal, OF `open S`] hoelzl@42844: by (simp add: inj_image_mem_iff) hoelzl@42844: next hoelzl@44791: fix S assume "?r" "open S" "ereal x \ S" hoelzl@44791: show "eventually (\x. ereal (f x) \ S) net" hoelzl@44791: using `?r`[THEN topological_tendstoD, OF open_ereal_vimage, OF `open S`] hoelzl@44791: using `ereal x \ S` by auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma lim_real_of_ereal[simp]: hoelzl@44791: assumes lim: "(f ---> ereal x) net" hoelzl@42844: shows "((\x. real (f x)) ---> x) net" hoelzl@42844: proof (intro topological_tendstoI) hoelzl@42844: fix S assume "open S" "x \ S" hoelzl@44791: then have S: "open S" "ereal x \ ereal ` S" hoelzl@42844: by (simp_all add: inj_image_mem_iff) hoelzl@44791: have "\x. f x \ ereal ` S \ real (f x) \ S" by auto hoelzl@44791: from this lim[THEN topological_tendstoD, OF open_ereal, OF S] hoelzl@42844: show "eventually (\x. real (f x) \ S) net" hoelzl@42844: by (rule eventually_mono) hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma Lim_PInfty: "f ----> \ <-> (ALL B. EX N. ALL n>=N. f n >= ereal B)" (is "?l = ?r") hoelzl@44794: proof hoelzl@44794: assume ?r hoelzl@44794: show ?l hoelzl@44794: apply(rule topological_tendstoI) hoelzl@42844: unfolding eventually_sequentially hoelzl@44794: proof- hoelzl@44794: fix S :: "ereal set" assume "open S" "\ : S" hoelzl@42844: from open_PInfty[OF this] guess B .. note B=this hoelzl@42844: from `?r`[rule_format,of "B+1"] guess N .. note N=this hoelzl@42844: show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI) hoelzl@42844: proof safe case goal1 hoelzl@44791: have "ereal B < ereal (B + 1)" by auto hoelzl@42844: also have "... <= f n" using goal1 N by auto hoelzl@42844: finally show ?case using B by fastsimp hoelzl@42844: qed hoelzl@42844: qed hoelzl@44794: next hoelzl@44794: assume ?l hoelzl@44794: show ?r hoelzl@44791: proof fix B::real have "open {ereal B<..}" "\ : {ereal B<..}" by auto hoelzl@42844: from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially] hoelzl@42844: guess N .. note N=this hoelzl@44791: show "EX N. ALL n>=N. ereal B <= f n" apply(rule_tac x=N in exI) using N by auto hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: hoelzl@42844: hoelzl@44791: lemma Lim_MInfty: "f ----> (-\) <-> (ALL B. EX N. ALL n>=N. f n <= ereal B)" (is "?l = ?r") hoelzl@44794: proof hoelzl@44794: assume ?r hoelzl@44794: show ?l hoelzl@44794: apply(rule topological_tendstoI) hoelzl@42844: unfolding eventually_sequentially hoelzl@44794: proof- hoelzl@44794: fix S :: "ereal set" hoelzl@44794: assume "open S" "(-\) : S" hoelzl@42844: from open_MInfty[OF this] guess B .. note B=this hoelzl@42844: from `?r`[rule_format,of "B-(1::real)"] guess N .. note N=this hoelzl@42844: show "EX N. ALL n>=N. f n : S" apply(rule_tac x=N in exI) hoelzl@42844: proof safe case goal1 hoelzl@44791: have "ereal (B - 1) >= f n" using goal1 N by auto hoelzl@44791: also have "... < ereal B" by auto hoelzl@42844: finally show ?case using B by fastsimp hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: next assume ?l show ?r hoelzl@44791: proof fix B::real have "open {..) : {..=N. ereal B >= f n" apply(rule_tac x=N in exI) using N by auto hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: hoelzl@42844: hoelzl@44791: lemma Lim_bounded_PInfty: assumes lim:"f ----> l" and "!!n. f n <= ereal B" shows "l ~= \" hoelzl@42844: proof(rule ccontr,unfold not_not) let ?B = "B + 1" assume as:"l=\" hoelzl@42844: from lim[unfolded this Lim_PInfty,rule_format,of "?B"] hoelzl@42844: guess N .. note N=this[rule_format,OF le_refl] hoelzl@44791: hence "ereal ?B <= ereal B" using assms(2)[of N] by(rule order_trans) hoelzl@44791: hence "ereal ?B < ereal ?B" apply (rule le_less_trans) by auto hoelzl@42844: thus False by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42844: hoelzl@44791: lemma Lim_bounded_MInfty: assumes lim:"f ----> l" and "!!n. f n >= ereal B" shows "l ~= (-\)" hoelzl@42844: proof(rule ccontr,unfold not_not) let ?B = "B - 1" assume as:"l=(-\)" hoelzl@42844: from lim[unfolded this Lim_MInfty,rule_format,of "?B"] hoelzl@42844: guess N .. note N=this[rule_format,OF le_refl] hoelzl@44791: hence "ereal B <= ereal ?B" using assms(2)[of N] order_trans[of "ereal B" "f N" "ereal(B - 1)"] by blast hoelzl@42844: thus False by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42844: hoelzl@42844: lemma tendsto_explicit: hoelzl@42844: "f ----> f0 <-> (ALL S. open S --> f0 : S --> (EX N. ALL n>=N. f n : S))" hoelzl@42844: unfolding tendsto_def eventually_sequentially by auto hoelzl@42844: hoelzl@42844: hoelzl@42844: lemma tendsto_obtains_N: hoelzl@42844: assumes "f ----> f0" hoelzl@42844: assumes "open S" "f0 : S" hoelzl@42844: obtains N where "ALL n>=N. f n : S" hoelzl@42844: using tendsto_explicit[of f f0] assms by auto hoelzl@42844: hoelzl@42844: hoelzl@42844: lemma tail_same_limit: hoelzl@42844: fixes X Y N hoelzl@42844: assumes "X ----> L" "ALL n>=N. X n = Y n" hoelzl@42844: shows "Y ----> L" hoelzl@42844: proof- hoelzl@42844: { fix S assume "open S" and "L:S" hoelzl@42844: from this obtain N1 where "ALL n>=N1. X n : S" hoelzl@42844: using assms unfolding tendsto_def eventually_sequentially by auto hoelzl@42844: hence "ALL n>=max N N1. Y n : S" using assms by auto hoelzl@42844: hence "EX N. ALL n>=N. Y n : S" apply(rule_tac x="max N N1" in exI) by auto hoelzl@42844: } hoelzl@42844: thus ?thesis using tendsto_explicit by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42844: hoelzl@42844: lemma Lim_bounded_PInfty2: hoelzl@44791: assumes lim:"f ----> l" and "ALL n>=N. f n <= ereal B" hoelzl@42844: shows "l ~= \" hoelzl@42844: proof- hoelzl@44791: def g == "(%n. if n>=N then f n else ereal B)" hoelzl@42844: hence "g ----> l" using tail_same_limit[of f l N g] lim by auto hoelzl@44791: moreover have "!!n. g n <= ereal B" using g_def assms by auto hoelzl@42844: ultimately show ?thesis using Lim_bounded_PInfty by auto hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma Lim_bounded_ereal: hoelzl@44791: assumes lim:"f ----> (l :: ereal)" hoelzl@42844: and "ALL n>=M. f n <= C" hoelzl@42844: shows "l<=C" hoelzl@42844: proof- hoelzl@42844: { assume "l=(-\)" hence ?thesis by auto } hoelzl@42844: moreover hoelzl@42844: { assume "~(l=(-\))" hoelzl@42844: { assume "C=\" hence ?thesis by auto } hoelzl@42844: moreover hoelzl@42844: { assume "C=(-\)" hence "ALL n>=M. f n = (-\)" using assms by auto hoelzl@42844: hence "l=(-\)" using assms hoelzl@42851: tendsto_unique[OF trivial_limit_sequentially] tail_same_limit[of "\n. -\" "-\" M f, OF tendsto_const] by auto hoelzl@42844: hence ?thesis by auto } hoelzl@42844: moreover hoelzl@44791: { assume "EX B. C = ereal B" hoelzl@44791: from this obtain B where B_def: "C=ereal B" by auto hoelzl@42844: hence "~(l=\)" using Lim_bounded_PInfty2 assms by auto hoelzl@44791: from this obtain m where m_def: "ereal m=l" using `~(l=(-\))` by (cases l) auto hoelzl@44791: from this obtain N where N_def: "ALL n>=N. f n : {ereal(m - 1) <..< ereal(m+1)}" hoelzl@44791: apply (subst tendsto_obtains_N[of f l "{ereal(m - 1) <..< ereal(m+1)}"]) using assms by auto hoelzl@42844: { fix n assume "n>=N" hoelzl@44791: hence "EX r. ereal r = f n" using N_def by (cases "f n") auto hoelzl@44791: } from this obtain g where g_def: "ALL n>=N. ereal (g n) = f n" by metis hoelzl@44791: hence "(%n. ereal (g n)) ----> l" using tail_same_limit[of f l N] assms by auto hoelzl@42844: hence *: "(%n. g n) ----> m" using m_def by auto hoelzl@42844: { fix n assume "n>=max N M" hoelzl@44791: hence "ereal (g n) <= ereal B" using assms g_def B_def by auto hoelzl@42844: hence "g n <= B" by auto hoelzl@42844: } hence "EX N. ALL n>=N. g n <= B" by blast hoelzl@42844: hence "m<=B" using * LIMSEQ_le_const2[of g m B] by auto hoelzl@42844: hence ?thesis using m_def B_def by auto hoelzl@42844: } ultimately have ?thesis by (cases C) auto hoelzl@42844: } ultimately show ?thesis by blast hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma real_of_ereal_mult[simp]: hoelzl@44791: fixes a b :: ereal shows "real (a * b) = real a * real b" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma real_of_ereal_eq_0: hoelzl@44794: fixes x :: ereal shows "real x = 0 \ x = \ \ x = -\ \ x = 0" hoelzl@42844: by (cases x) auto hoelzl@42844: hoelzl@44791: lemma tendsto_ereal_realD: hoelzl@44791: fixes f :: "'a \ ereal" hoelzl@44791: assumes "x \ 0" and tendsto: "((\x. ereal (real (f x))) ---> x) net" hoelzl@42844: shows "(f ---> x) net" hoelzl@42844: proof (intro topological_tendstoI) hoelzl@42844: fix S assume S: "open S" "x \ S" hoelzl@42844: with `x \ 0` have "open (S - {0})" "x \ S - {0}" by auto hoelzl@42844: from tendsto[THEN topological_tendstoD, OF this] hoelzl@42844: show "eventually (\x. f x \ S) net" hoelzl@44791: by (rule eventually_rev_mp) (auto simp: ereal_real real_of_ereal_0) hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma tendsto_ereal_realI: hoelzl@44791: fixes f :: "'a \ ereal" hoelzl@42847: assumes x: "\x\ \ \" and tendsto: "(f ---> x) net" hoelzl@44791: shows "((\x. ereal (real (f x))) ---> x) net" hoelzl@42844: proof (intro topological_tendstoI) hoelzl@42844: fix S assume "open S" "x \ S" hoelzl@42844: with x have "open (S - {\, -\})" "x \ S - {\, -\}" by auto hoelzl@42844: from tendsto[THEN topological_tendstoD, OF this] hoelzl@44791: show "eventually (\x. ereal (real (f x)) \ S) net" hoelzl@44791: by (elim eventually_elim1) (auto simp: ereal_real) hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_mult_cancel_left: hoelzl@44791: fixes a b c :: ereal shows "a * b = a * c \ hoelzl@42847: ((\a\ = \ \ 0 < b * c) \ a = 0 \ b = c)" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@42844: (simp_all add: zero_less_mult_iff) hoelzl@42844: hoelzl@44791: lemma ereal_inj_affinity: hoelzl@44794: fixes m t :: ereal hoelzl@42847: assumes "\m\ \ \" "m \ 0" "\t\ \ \" hoelzl@42844: shows "inj_on (\x. m * x + t) A" hoelzl@42844: using assms hoelzl@44791: by (cases rule: ereal2_cases[of m t]) hoelzl@44791: (auto intro!: inj_onI simp: ereal_add_cancel_right ereal_mult_cancel_left) hoelzl@42844: hoelzl@44791: lemma ereal_PInfty_eq_plus[simp]: hoelzl@44794: fixes a b :: ereal hoelzl@42844: shows "\ = a + b \ a = \ \ b = \" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_MInfty_eq_plus[simp]: hoelzl@44794: fixes a b :: ereal hoelzl@42844: shows "-\ = a + b \ (a = -\ \ b \ \) \ (b = -\ \ a \ \)" hoelzl@44791: by (cases rule: ereal2_cases[of a b]) auto hoelzl@42844: hoelzl@44791: lemma ereal_less_divide_pos: hoelzl@44794: fixes x y :: ereal hoelzl@44794: shows "x > 0 \ x \ \ \ y < z / x \ x * y < z" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_divide_less_pos: hoelzl@44794: fixes x y z :: ereal hoelzl@44794: shows "x > 0 \ x \ \ \ y / x < z \ y < x * z" hoelzl@44791: by (cases rule: ereal3_cases[of x y z]) (auto simp: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_divide_eq: hoelzl@44794: fixes a b c :: ereal hoelzl@44794: shows "b \ 0 \ \b\ \ \ \ a / b = c \ a = b * c" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@42844: (simp_all add: field_simps) hoelzl@42844: hoelzl@44794: lemma ereal_inverse_not_MInfty[simp]: "inverse (a::ereal) \ -\" hoelzl@42844: by (cases a) auto hoelzl@42844: hoelzl@44791: lemma ereal_mult_m1[simp]: "x * ereal (-1) = -x" hoelzl@42844: by (cases x) auto hoelzl@42844: hoelzl@44791: lemma ereal_LimI_finite: hoelzl@44794: fixes x :: ereal hoelzl@42847: assumes "\x\ \ \" hoelzl@42844: assumes "!!r. 0 < r ==> EX N. ALL n>=N. u n < x + r & x < u n + r" hoelzl@42844: shows "u ----> x" hoelzl@42844: proof (rule topological_tendstoI, unfold eventually_sequentially) hoelzl@44791: obtain rx where rx_def: "x=ereal rx" using assms by (cases x) auto hoelzl@42844: fix S assume "open S" "x : S" hoelzl@44791: then have "open (ereal -` S)" unfolding open_ereal_def by auto hoelzl@44791: with `x \ S` obtain r where "0 < r" and dist: "!!y. dist y rx < r ==> ereal y \ S" hoelzl@42846: unfolding open_real_def rx_def by auto hoelzl@42844: then obtain n where hoelzl@44791: upper: "!!N. n <= N ==> u N < x + ereal r" and hoelzl@44791: lower: "!!N. n <= N ==> x < u N + ereal r" using assms(2)[of "ereal r"] by auto hoelzl@42844: show "EX N. ALL n>=N. u n : S" hoelzl@42844: proof (safe intro!: exI[of _ n]) hoelzl@42844: fix N assume "n <= N" hoelzl@42844: from upper[OF this] lower[OF this] assms `0 < r` hoelzl@42844: have "u N ~: {\,(-\)}" by auto hoelzl@44791: from this obtain ra where ra_def: "(u N) = ereal ra" by (cases "u N") auto hoelzl@42844: hence "rx < ra + r" and "ra < rx + r" hoelzl@42844: using rx_def assms `0 < r` lower[OF `n <= N`] upper[OF `n <= N`] by auto hoelzl@42846: hence "dist (real (u N)) rx < r" hoelzl@42844: using rx_def ra_def hoelzl@42844: by (auto simp: dist_real_def abs_diff_less_iff field_simps) hoelzl@42847: from dist[OF this] show "u N : S" using `u N ~: {\, -\}` hoelzl@44791: by (auto simp: ereal_real split: split_if_asm) hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_LimI_finite_iff: hoelzl@44794: fixes x :: ereal hoelzl@42847: assumes "\x\ \ \" hoelzl@42844: shows "u ----> x <-> (ALL r. 0 < r --> (EX N. ALL n>=N. u n < x + r & x < u n + r))" hoelzl@42844: (is "?lhs <-> ?rhs") hoelzl@42847: proof hoelzl@42847: assume lim: "u ----> x" hoelzl@44791: { fix r assume "(r::ereal)>0" hoelzl@42844: from this obtain N where N_def: "ALL n>=N. u n : {x - r <..< x + r}" hoelzl@42844: apply (subst tendsto_obtains_N[of u x "{x - r <..< x + r}"]) hoelzl@44791: using lim ereal_between[of x r] assms `r>0` by auto hoelzl@42844: hence "EX N. ALL n>=N. u n < x + r & x < u n + r" hoelzl@44791: using ereal_minus_less[of r x] by (cases r) auto hoelzl@42847: } then show "?rhs" by auto hoelzl@42847: next hoelzl@42847: assume ?rhs then show "u ----> x" hoelzl@44791: using ereal_LimI_finite[of x] assms by auto hoelzl@42844: qed hoelzl@42844: hoelzl@42844: hoelzl@42844: subsubsection {* @{text Liminf} and @{text Limsup} *} hoelzl@42844: hoelzl@42844: definition hoelzl@42844: "Liminf net f = (GREATEST l. \yx. y < f x) net)" hoelzl@42844: hoelzl@42844: definition hoelzl@42844: "Limsup net f = (LEAST l. \y>l. eventually (\x. f x < y) net)" hoelzl@42844: hoelzl@42844: lemma Liminf_Sup: haftmann@44812: fixes f :: "'a => 'b::complete_linorder" hoelzl@42844: shows "Liminf net f = Sup {l. \yx. y < f x) net}" hoelzl@42844: by (auto intro!: Greatest_equality complete_lattice_class.Sup_upper simp: less_Sup_iff Liminf_def) hoelzl@42844: hoelzl@42844: lemma Limsup_Inf: haftmann@44812: fixes f :: "'a => 'b::complete_linorder" hoelzl@42844: shows "Limsup net f = Inf {l. \y>l. eventually (\x. f x < y) net}" hoelzl@42844: by (auto intro!: Least_equality complete_lattice_class.Inf_lower simp: Inf_less_iff Limsup_def) hoelzl@42844: hoelzl@44791: lemma ereal_SupI: hoelzl@44791: fixes x :: ereal hoelzl@42844: assumes "\y. y \ A \ y \ x" hoelzl@42844: assumes "\y. (\z. z \ A \ z \ y) \ x \ y" hoelzl@42844: shows "Sup A = x" hoelzl@44791: unfolding Sup_ereal_def hoelzl@42844: using assms by (auto intro!: Least_equality) hoelzl@42844: hoelzl@44791: lemma ereal_InfI: hoelzl@44791: fixes x :: ereal hoelzl@42844: assumes "\i. i \ A \ x \ i" hoelzl@42844: assumes "\y. (\i. i \ A \ y \ i) \ y \ x" hoelzl@42844: shows "Inf A = x" hoelzl@44791: unfolding Inf_ereal_def hoelzl@42844: using assms by (auto intro!: Greatest_equality) hoelzl@42844: hoelzl@42844: lemma Limsup_const: haftmann@44812: fixes c :: "'a::complete_linorder" hoelzl@42844: assumes ntriv: "\ trivial_limit net" hoelzl@42844: shows "Limsup net (\x. c) = c" hoelzl@42844: unfolding Limsup_Inf hoelzl@42844: proof (safe intro!: antisym complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower) hoelzl@42844: fix x assume *: "\y>x. eventually (\_. c < y) net" hoelzl@42844: show "c \ x" hoelzl@42844: proof (rule ccontr) hoelzl@42844: assume "\ c \ x" then have "x < c" by auto hoelzl@42844: then show False using ntriv * by (auto simp: trivial_limit_def) hoelzl@42844: qed hoelzl@42844: qed auto hoelzl@42844: hoelzl@42844: lemma Liminf_const: haftmann@44812: fixes c :: "'a::complete_linorder" hoelzl@42844: assumes ntriv: "\ trivial_limit net" hoelzl@42844: shows "Liminf net (\x. c) = c" hoelzl@42844: unfolding Liminf_Sup hoelzl@42844: proof (safe intro!: antisym complete_lattice_class.Sup_least complete_lattice_class.Sup_upper) hoelzl@42844: fix x assume *: "\y_. y < c) net" hoelzl@42844: show "x \ c" hoelzl@42844: proof (rule ccontr) hoelzl@42844: assume "\ x \ c" then have "c < x" by auto hoelzl@42844: then show False using ntriv * by (auto simp: trivial_limit_def) hoelzl@42844: qed hoelzl@42844: qed auto hoelzl@42844: haftmann@44812: lemma (in order) mono_set: haftmann@44812: "mono S \ (\x y. x \ y \ x \ S \ y \ S)" hoelzl@42844: by (auto simp: mono_def mem_def) hoelzl@42844: haftmann@44812: lemma (in order) mono_greaterThan [intro, simp]: "mono {B<..}" unfolding mono_set by auto haftmann@44812: lemma (in order) mono_atLeast [intro, simp]: "mono {B..}" unfolding mono_set by auto haftmann@44812: lemma (in order) mono_UNIV [intro, simp]: "mono UNIV" unfolding mono_set by auto haftmann@44812: lemma (in order) mono_empty [intro, simp]: "mono {}" unfolding mono_set by auto hoelzl@42844: haftmann@44812: lemma (in complete_linorder) mono_set_iff: haftmann@44812: fixes S :: "'a set" hoelzl@42844: defines "a \ Inf S" hoelzl@42844: shows "mono S \ (S = {a <..} \ S = {a..})" (is "_ = ?c") hoelzl@42844: proof hoelzl@42844: assume "mono S" hoelzl@42844: then have mono: "\x y. x \ y \ x \ S \ y \ S" by (auto simp: mono_set) hoelzl@42844: show ?c hoelzl@42844: proof cases hoelzl@42844: assume "a \ S" hoelzl@42844: show ?c hoelzl@42844: using mono[OF _ `a \ S`] haftmann@44812: by (auto intro: Inf_lower simp: a_def) hoelzl@42844: next hoelzl@42844: assume "a \ S" hoelzl@42844: have "S = {a <..}" hoelzl@42844: proof safe hoelzl@42844: fix x assume "x \ S" haftmann@44812: then have "a \ x" unfolding a_def by (rule Inf_lower) hoelzl@42844: then show "a < x" using `x \ S` `a \ S` by (cases "a = x") auto hoelzl@42844: next hoelzl@42844: fix x assume "a < x" hoelzl@42844: then obtain y where "y < x" "y \ S" unfolding a_def Inf_less_iff .. hoelzl@42844: with mono[of y x] show "x \ S" by auto hoelzl@42844: qed hoelzl@42844: then show ?c .. hoelzl@42844: qed hoelzl@42844: qed auto hoelzl@42844: hoelzl@42844: lemma lim_imp_Liminf: hoelzl@44791: fixes f :: "'a \ ereal" hoelzl@42844: assumes ntriv: "\ trivial_limit net" hoelzl@42844: assumes lim: "(f ---> f0) net" hoelzl@42844: shows "Liminf net f = f0" hoelzl@42844: unfolding Liminf_Sup hoelzl@44791: proof (safe intro!: ereal_SupI) hoelzl@42844: fix y assume *: "\y'x. y' < f x) net" hoelzl@42844: show "y \ f0" hoelzl@44791: proof (rule ereal_le_ereal) hoelzl@42844: fix B assume "B < y" hoelzl@42844: { assume "f0 < B" hoelzl@42844: then have "eventually (\x. f x < B \ B < f x) net" hoelzl@42844: using topological_tendstoD[OF lim, of "{..x. f x < B \ B < f x) = (\x. False)" by (auto simp: fun_eq_iff) hoelzl@42844: finally have False using ntriv[unfolded trivial_limit_def] by auto hoelzl@42844: } then show "B \ f0" by (metis linorder_le_less_linear) hoelzl@42844: qed hoelzl@42844: next hoelzl@42844: fix y assume *: "\z. z \ {l. \yx. y < f x) net} \ z \ y" hoelzl@42844: show "f0 \ y" hoelzl@42844: proof (safe intro!: *[rule_format]) hoelzl@42844: fix y assume "y < f0" then show "eventually (\x. y < f x) net" hoelzl@42844: using lim[THEN topological_tendstoD, of "{y <..}"] by auto hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_Liminf_le_Limsup: hoelzl@44791: fixes f :: "'a \ ereal" hoelzl@42844: assumes ntriv: "\ trivial_limit net" hoelzl@42844: shows "Liminf net f \ Limsup net f" hoelzl@42844: unfolding Limsup_Inf Liminf_Sup hoelzl@42844: proof (safe intro!: complete_lattice_class.Inf_greatest complete_lattice_class.Sup_least) hoelzl@42844: fix u v assume *: "\yx. y < f x) net" "\y>v. eventually (\x. f x < y) net" hoelzl@42844: show "u \ v" hoelzl@42844: proof (rule ccontr) hoelzl@42844: assume "\ u \ v" hoelzl@42844: then obtain t where "t < u" "v < t" hoelzl@44791: using ereal_dense[of v u] by (auto simp: not_le) hoelzl@42844: then have "eventually (\x. t < f x \ f x < t) net" hoelzl@42844: using * by (auto intro: eventually_conj) hoelzl@42844: also have "(\x. t < f x \ f x < t) = (\x. False)" by (auto simp: fun_eq_iff) hoelzl@42844: finally show False using ntriv by (auto simp: trivial_limit_def) hoelzl@42844: qed hoelzl@42844: qed hoelzl@42844: hoelzl@42844: lemma Liminf_mono: hoelzl@44791: fixes f g :: "'a => ereal" hoelzl@42844: assumes ev: "eventually (\x. f x \ g x) net" hoelzl@42844: shows "Liminf net f \ Liminf net g" hoelzl@42844: unfolding Liminf_Sup hoelzl@42844: proof (safe intro!: Sup_mono bexI) hoelzl@42844: fix a y assume "\yx. y < f x) net" and "y < a" hoelzl@42844: then have "eventually (\x. y < f x) net" by auto hoelzl@42844: then show "eventually (\x. y < g x) net" hoelzl@42844: by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto) hoelzl@42844: qed simp hoelzl@42844: hoelzl@42844: lemma Liminf_eq: hoelzl@44791: fixes f g :: "'a \ ereal" hoelzl@42844: assumes "eventually (\x. f x = g x) net" hoelzl@42844: shows "Liminf net f = Liminf net g" hoelzl@42844: by (intro antisym Liminf_mono eventually_mono[OF _ assms]) auto hoelzl@42844: hoelzl@42844: lemma Liminf_mono_all: hoelzl@44791: fixes f g :: "'a \ ereal" hoelzl@42844: assumes "\x. f x \ g x" hoelzl@42844: shows "Liminf net f \ Liminf net g" hoelzl@42844: using assms by (intro Liminf_mono always_eventually) auto hoelzl@42844: hoelzl@42844: lemma Limsup_mono: hoelzl@44791: fixes f g :: "'a \ ereal" hoelzl@42844: assumes ev: "eventually (\x. f x \ g x) net" hoelzl@42844: shows "Limsup net f \ Limsup net g" hoelzl@42844: unfolding Limsup_Inf hoelzl@42844: proof (safe intro!: Inf_mono bexI) hoelzl@42844: fix a y assume "\y>a. eventually (\x. g x < y) net" and "a < y" hoelzl@42844: then have "eventually (\x. g x < y) net" by auto hoelzl@42844: then show "eventually (\x. f x < y) net" hoelzl@42844: by (rule eventually_rev_mp) (rule eventually_mono[OF _ ev], auto) hoelzl@42844: qed simp hoelzl@42844: hoelzl@42844: lemma Limsup_mono_all: hoelzl@44791: fixes f g :: "'a \ ereal" hoelzl@42844: assumes "\x. f x \ g x" hoelzl@42844: shows "Limsup net f \ Limsup net g" hoelzl@42844: using assms by (intro Limsup_mono always_eventually) auto hoelzl@42844: hoelzl@42844: lemma Limsup_eq: hoelzl@44791: fixes f g :: "'a \ ereal" hoelzl@42844: assumes "eventually (\x. f x = g x) net" hoelzl@42844: shows "Limsup net f = Limsup net g" hoelzl@42844: by (intro antisym Limsup_mono eventually_mono[OF _ assms]) auto hoelzl@42844: hoelzl@42844: abbreviation "liminf \ Liminf sequentially" hoelzl@42844: hoelzl@42844: abbreviation "limsup \ Limsup sequentially" hoelzl@42844: hoelzl@42844: lemma liminf_SUPR_INFI: hoelzl@44791: fixes f :: "nat \ ereal" hoelzl@42844: shows "liminf f = (SUP n. INF m:{n..}. f m)" hoelzl@42844: unfolding Liminf_Sup eventually_sequentially hoelzl@42844: proof (safe intro!: antisym complete_lattice_class.Sup_least) hoelzl@42844: fix x assume *: "\yN. \n\N. y < f n" show "x \ (SUP n. INF m:{n..}. f m)" hoelzl@44791: proof (rule ereal_le_ereal) hoelzl@42844: fix y assume "y < x" hoelzl@42844: with * obtain N where "\n. N \ n \ y < f n" by auto hoelzl@42844: then have "y \ (INF m:{N..}. f m)" by (force simp: le_INF_iff) hoelzl@42844: also have "\ \ (SUP n. INF m:{n..}. f m)" by (intro le_SUPI) auto hoelzl@42844: finally show "y \ (SUP n. INF m:{n..}. f m)" . hoelzl@42844: qed hoelzl@42844: next hoelzl@42844: show "(SUP n. INF m:{n..}. f m) \ Sup {l. \yN. \n\N. y < f n}" hoelzl@42844: proof (unfold SUPR_def, safe intro!: Sup_mono bexI) hoelzl@42844: fix y n assume "y < INFI {n..} f" hoelzl@42844: from less_INFD[OF this] show "\N. \n\N. y < f n" by (intro exI[of _ n]) auto hoelzl@42844: qed (rule order_refl) hoelzl@42844: qed hoelzl@42844: hoelzl@42844: lemma tail_same_limsup: hoelzl@44791: fixes X Y :: "nat => ereal" hoelzl@42844: assumes "\n. N \ n \ X n = Y n" hoelzl@42844: shows "limsup X = limsup Y" hoelzl@42844: using Limsup_eq[of X Y sequentially] eventually_sequentially assms by auto hoelzl@42844: hoelzl@42844: lemma tail_same_liminf: hoelzl@44791: fixes X Y :: "nat => ereal" hoelzl@42844: assumes "\n. N \ n \ X n = Y n" hoelzl@42844: shows "liminf X = liminf Y" hoelzl@42844: using Liminf_eq[of X Y sequentially] eventually_sequentially assms by auto hoelzl@42844: hoelzl@42844: lemma liminf_mono: hoelzl@44791: fixes X Y :: "nat \ ereal" hoelzl@42844: assumes "\n. N \ n \ X n <= Y n" hoelzl@42844: shows "liminf X \ liminf Y" hoelzl@42844: using Liminf_mono[of X Y sequentially] eventually_sequentially assms by auto hoelzl@42844: hoelzl@42844: lemma limsup_mono: hoelzl@44791: fixes X Y :: "nat => ereal" hoelzl@42844: assumes "\n. N \ n \ X n <= Y n" hoelzl@42844: shows "limsup X \ limsup Y" hoelzl@42844: using Limsup_mono[of X Y sequentially] eventually_sequentially assms by auto hoelzl@42844: hoelzl@42844: declare trivial_limit_sequentially[simp] hoelzl@42844: hoelzl@42849: lemma hoelzl@44791: fixes X :: "nat \ ereal" hoelzl@44791: shows ereal_incseq_uminus[simp]: "incseq (\i. - X i) = decseq X" hoelzl@44791: and ereal_decseq_uminus[simp]: "decseq (\i. - X i) = incseq X" hoelzl@42849: unfolding incseq_def decseq_def by auto hoelzl@42849: hoelzl@42844: lemma liminf_bounded: hoelzl@44791: fixes X Y :: "nat \ ereal" hoelzl@42844: assumes "\n. N \ n \ C \ X n" hoelzl@42844: shows "C \ liminf X" hoelzl@42844: using liminf_mono[of N "\n. C" X] assms Liminf_const[of sequentially C] by simp hoelzl@42844: hoelzl@42844: lemma limsup_bounded: hoelzl@44791: fixes X Y :: "nat => ereal" hoelzl@42844: assumes "\n. N \ n \ X n <= C" hoelzl@42844: shows "limsup X \ C" hoelzl@42844: using limsup_mono[of N X "\n. C"] assms Limsup_const[of sequentially C] by simp hoelzl@42844: hoelzl@42844: lemma liminf_bounded_iff: hoelzl@44791: fixes x :: "nat \ ereal" hoelzl@42844: shows "C \ liminf x \ (\BN. \n\N. B < x n)" (is "?lhs <-> ?rhs") hoelzl@42844: proof safe hoelzl@42844: fix B assume "B < C" "C \ liminf x" hoelzl@42844: then have "B < liminf x" by auto hoelzl@42844: then obtain N where "B < (INF m:{N..}. x m)" hoelzl@42844: unfolding liminf_SUPR_INFI SUPR_def less_Sup_iff by auto hoelzl@42844: from less_INFD[OF this] show "\N. \n\N. B < x n" by auto hoelzl@42844: next hoelzl@42844: assume *: "\BN. \n\N. B < x n" hoelzl@42844: { fix B assume "Bn\N. B < x n" using `?rhs` by auto hoelzl@42844: hence "B \ (INF m:{N..}. x m)" by (intro le_INFI) auto hoelzl@42844: also have "... \ liminf x" unfolding liminf_SUPR_INFI by (intro le_SUPI) simp hoelzl@42844: finally have "B \ liminf x" . hoelzl@42844: } then show "?lhs" by (metis * leD liminf_bounded linorder_le_less_linear) hoelzl@42844: qed hoelzl@42844: hoelzl@42844: lemma liminf_subseq_mono: hoelzl@44791: fixes X :: "nat \ ereal" hoelzl@42844: assumes "subseq r" hoelzl@42844: shows "liminf X \ liminf (X \ r) " hoelzl@42844: proof- hoelzl@42844: have "\n. (INF m:{n..}. X m) \ (INF m:{n..}. (X \ r) m)" hoelzl@42844: proof (safe intro!: INF_mono) hoelzl@42844: fix n m :: nat assume "n \ m" then show "\ma\{n..}. X ma \ (X \ r) m" hoelzl@42844: using seq_suble[OF `subseq r`, of m] by (intro bexI[of _ "r m"]) auto hoelzl@42844: qed hoelzl@42844: then show ?thesis by (auto intro!: SUP_mono simp: liminf_SUPR_INFI comp_def) hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma ereal_real': assumes "\x\ \ \" shows "ereal (real x) = x" hoelzl@42847: using assms by auto hoelzl@42844: hoelzl@44791: lemma ereal_le_ereal_bounded: hoelzl@44791: fixes x y z :: ereal hoelzl@42849: assumes "z \ y" hoelzl@42849: assumes *: "\B. z < B \ B < x \ B \ y" hoelzl@42849: shows "x \ y" hoelzl@44791: proof (rule ereal_le_ereal) hoelzl@42849: fix B assume "B < x" hoelzl@42849: show "B \ y" hoelzl@42849: proof cases hoelzl@42849: assume "z < B" from *[OF this `B < x`] show "B \ y" . hoelzl@42847: next hoelzl@42849: assume "\ z < B" with `z \ y` show "B \ y" by auto hoelzl@42847: qed hoelzl@42844: qed hoelzl@42844: hoelzl@44791: lemma fixes x y :: ereal hoelzl@42849: shows Sup_atMost[simp]: "Sup {.. y} = y" hoelzl@42849: and Sup_lessThan[simp]: "Sup {..< y} = y" hoelzl@42849: and Sup_atLeastAtMost[simp]: "x \ y \ Sup { x .. y} = y" hoelzl@42849: and Sup_greaterThanAtMost[simp]: "x < y \ Sup { x <.. y} = y" hoelzl@42849: and Sup_atLeastLessThan[simp]: "x < y \ Sup { x ..< y} = y" hoelzl@44791: by (auto simp: Sup_ereal_def intro!: Least_equality hoelzl@44791: intro: ereal_le_ereal ereal_le_ereal_bounded[of x]) hoelzl@42849: hoelzl@42849: lemma Sup_greaterThanlessThan[simp]: hoelzl@44791: fixes x y :: ereal assumes "x < y" shows "Sup { x <..< y} = y" hoelzl@44791: unfolding Sup_ereal_def hoelzl@44791: proof (intro Least_equality ereal_le_ereal_bounded[of _ _ y]) hoelzl@42849: fix z assume z: "\u\{x<.. z" hoelzl@44791: from ereal_dense[OF `x < y`] guess w .. note w = this hoelzl@42849: with z[THEN bspec, of w] show "x \ z" by auto hoelzl@42849: qed auto hoelzl@42849: hoelzl@44791: lemma real_ereal_id: "real o ereal = id" hoelzl@42844: proof- hoelzl@44791: { fix x have "(real o ereal) x = id x" by auto } hoelzl@42844: from this show ?thesis using ext by blast hoelzl@42844: qed hoelzl@42844: hoelzl@44794: lemma open_image_ereal: "open(UNIV-{ \ , (-\ :: ereal)})" hoelzl@44791: by (metis range_ereal open_ereal open_UNIV) hoelzl@42844: hoelzl@44791: lemma ereal_le_distrib: hoelzl@44791: fixes a b c :: ereal shows "c * (a + b) \ c * a + c * b" hoelzl@44791: by (cases rule: ereal3_cases[of a b c]) hoelzl@42844: (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) hoelzl@42844: hoelzl@44791: lemma ereal_pos_distrib: hoelzl@44791: fixes a b c :: ereal assumes "0 \ c" "c \ \" shows "c * (a + b) = c * a + c * b" hoelzl@44791: using assms by (cases rule: ereal3_cases[of a b c]) hoelzl@42844: (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) hoelzl@42844: hoelzl@44791: lemma ereal_pos_le_distrib: hoelzl@44791: fixes a b c :: ereal hoelzl@42844: assumes "c>=0" hoelzl@42844: shows "c * (a + b) <= c * a + c * b" hoelzl@44791: using assms by (cases rule: ereal3_cases[of a b c]) hoelzl@42844: (auto simp add: field_simps) hoelzl@42844: hoelzl@44791: lemma ereal_max_mono: hoelzl@44791: "[| (a::ereal) <= b; c <= d |] ==> max a c <= max b d" hoelzl@44791: by (metis sup_ereal_def sup_mono) hoelzl@42844: hoelzl@42844: hoelzl@44791: lemma ereal_max_least: hoelzl@44791: "[| (a::ereal) <= x; c <= x |] ==> max a c <= x" hoelzl@44791: by (metis sup_ereal_def sup_least) hoelzl@42844: hoelzl@44804: subsubsection {* Tests for code generator *} hoelzl@44804: hoelzl@44804: (* A small list of simple arithmetic expressions *) hoelzl@44804: hoelzl@44804: value [code] "- \ :: ereal" hoelzl@44804: value [code] "\-\\ :: ereal" hoelzl@44804: value [code] "4 + 5 / 4 - ereal 2 :: ereal" hoelzl@44804: value [code] "ereal 3 < \" hoelzl@44804: value [code] "real (\::ereal) = 0" hoelzl@44804: hoelzl@42844: end