# HG changeset patch # User nipkow # Date 1531295028 -7200 # Node ID f221bc388ad019211212ec9b93d02f2248cb8d6a # Parent 640386ab99f3dec93edfa0662825c0e6b16a6fa2 (re)moved lemmas diff -r 640386ab99f3 -r f221bc388ad0 src/HOL/Analysis/Extended_Real_Limits.thy --- a/src/HOL/Analysis/Extended_Real_Limits.thy Wed Jul 11 11:16:26 2018 +0200 +++ b/src/HOL/Analysis/Extended_Real_Limits.thy Wed Jul 11 09:43:48 2018 +0200 @@ -361,59 +361,6 @@ subsection \Extended-Real.thy\ -text\The proof of this one is copied from \verb+ereal_add_mono+.\ -lemma ereal_add_strict_mono2: - fixes a b c d :: ereal - assumes "a < b" - and "c < d" - shows "a + c < b + d" -using assms -apply (cases a) -apply (cases rule: ereal3_cases[of b c d], auto) -apply (cases rule: ereal3_cases[of b c d], auto) -done - -text \The next ones are analogues of \verb+mult_mono+ and \verb+mult_mono'+ in ereal.\ - -lemma ereal_mult_mono: - fixes a b c d::ereal - assumes "b \ 0" "c \ 0" "a \ b" "c \ d" - shows "a * c \ b * d" -by (metis ereal_mult_right_mono mult.commute order_trans assms) - -lemma ereal_mult_mono': - fixes a b c d::ereal - assumes "a \ 0" "c \ 0" "a \ b" "c \ d" - shows "a * c \ b * d" -by (metis ereal_mult_right_mono mult.commute order_trans assms) - -lemma ereal_mult_mono_strict: - fixes a b c d::ereal - assumes "b > 0" "c > 0" "a < b" "c < d" - shows "a * c < b * d" -proof - - have "c < \" using \c < d\ by auto - then have "a * c < b * c" by (metis ereal_mult_strict_left_mono[OF assms(3) assms(2)] mult.commute) - moreover have "b * c \ b * d" using assms(2) assms(4) by (simp add: assms(1) ereal_mult_left_mono less_imp_le) - ultimately show ?thesis by simp -qed - -lemma ereal_mult_mono_strict': - fixes a b c d::ereal - assumes "a > 0" "c > 0" "a < b" "c < d" - shows "a * c < b * d" -apply (rule ereal_mult_mono_strict, auto simp add: assms) using assms by auto - -lemma ereal_abs_add: - fixes a b::ereal - shows "abs(a+b) \ abs a + abs b" -by (cases rule: ereal2_cases[of a b]) (auto) - -lemma ereal_abs_diff: - fixes a b::ereal - shows "abs(a-b) \ abs a + abs b" -by (cases rule: ereal2_cases[of a b]) (auto) - lemma sum_constant_ereal: fixes a::ereal shows "(\i\I. a) = a * card I" @@ -1569,7 +1516,7 @@ ultimately have "(w o a) \ limsup (u o r) + limsup (v o r o s)" by simp then have "limsup w = limsup (u o r) + limsup (v o r o s)" using l(1) LIMSEQ_unique by blast then have "limsup w \ limsup u + limsup v" - using \limsup (u o r) \ limsup u\ \limsup (v o r o s) \ limsup v\ ereal_add_mono by simp + using \limsup (u o r) \ limsup u\ \limsup (v o r o s) \ limsup v\ add_mono by simp then show ?thesis unfolding w_def by simp qed @@ -1616,7 +1563,7 @@ ultimately have "(w o a) \ liminf (u o r) + liminf (v o r o s)" by simp then have "liminf w = liminf (u o r) + liminf (v o r o s)" using l(1) LIMSEQ_unique by blast then have "liminf w \ liminf u + liminf v" - using \liminf (u o r) \ liminf u\ \liminf (v o r o s) \ liminf v\ ereal_add_mono by simp + using \liminf (u o r) \ liminf u\ \liminf (v o r o s) \ liminf v\ add_mono by simp then show ?thesis unfolding w_def by simp qed diff -r 640386ab99f3 -r f221bc388ad0 src/HOL/Library/Extended_Nonnegative_Real.thy --- a/src/HOL/Library/Extended_Nonnegative_Real.thy Wed Jul 11 11:16:26 2018 +0200 +++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Wed Jul 11 09:43:48 2018 +0200 @@ -1629,7 +1629,7 @@ apply transfer apply (simp add: SUP_ereal_add_left) apply (subst (1 2) max.absorb2) - apply (auto intro: SUP_upper2 ereal_add_nonneg_nonneg) + apply (auto intro: SUP_upper2 add_nonneg_nonneg) done lemma ennreal_SUP_const_minus: (* TODO: rename: ennreal_SUP_const_minus *) @@ -1814,7 +1814,7 @@ shows "z \ y \ x + (y - z) = x + y - z" including ennreal.lifting by transfer - (insert ereal_add_mono[of 0], auto simp add: ereal_diff_positive max.absorb2 add_diff_eq_ereal) + (insert add_mono[of "0::ereal"], auto simp add: ereal_diff_positive max.absorb2 add_diff_eq_ereal) lemma add_diff_inverse_ennreal: fixes x y :: ennreal shows "x \ y \ x + (y - x) = y" diff -r 640386ab99f3 -r f221bc388ad0 src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Wed Jul 11 11:16:26 2018 +0200 +++ b/src/HOL/Library/Extended_Real.thy Wed Jul 11 09:43:48 2018 +0200 @@ -593,16 +593,16 @@ using reals_Archimedean2[of r] by simp qed simp_all -lemma ereal_add_mono: +lemma ereal_add_strict_mono2: fixes a b c d :: ereal - assumes "a \ b" - and "c \ d" - shows "a + c \ b + d" - using assms - apply (cases a) - apply (cases rule: ereal3_cases[of b c d], auto) - apply (cases rule: ereal3_cases[of b c d], auto) - done + assumes "a < b" + and "c < d" + shows "a + c < b + d" +using assms +apply (cases a) +apply (cases rule: ereal3_cases[of b c d], auto) +apply (cases rule: ereal3_cases[of b c d], auto) +done lemma ereal_minus_le_minus[simp]: fixes a b :: ereal @@ -674,6 +674,11 @@ shows "\ x \ y; -x \ y \ \ \x\ \ y" by(cases x y rule: ereal2_cases)(simp_all) +lemma ereal_abs_add: + fixes a b::ereal + shows "abs(a+b) \ abs a + abs b" +by (cases rule: ereal2_cases[of a b]) (auto) + lemma real_of_ereal_le_0[simp]: "real_of_ereal (x :: ereal) \ 0 \ x \ 0 \ x = \" by (cases x) auto @@ -783,11 +788,6 @@ lemma incseq_ereal: "incseq f \ incseq (\x. ereal (f x))" unfolding incseq_def by auto -lemma ereal_add_nonneg_nonneg[simp]: - fixes a b :: ereal - shows "0 \ a \ 0 \ b \ 0 \ a + b" - using add_mono[of 0 a 0 b] by simp - lemma sum_ereal[simp]: "(\x\A. ereal (f x)) = ereal (\x\A. f x)" proof (cases "finite A") case True @@ -1050,6 +1050,35 @@ using ereal_mult_right_mono by (simp add: mult.commute[of c]) +lemma ereal_mult_mono: + fixes a b c d::ereal + assumes "b \ 0" "c \ 0" "a \ b" "c \ d" + shows "a * c \ b * d" +by (metis ereal_mult_right_mono mult.commute order_trans assms) + +lemma ereal_mult_mono': + fixes a b c d::ereal + assumes "a \ 0" "c \ 0" "a \ b" "c \ d" + shows "a * c \ b * d" +by (metis ereal_mult_right_mono mult.commute order_trans assms) + +lemma ereal_mult_mono_strict: + fixes a b c d::ereal + assumes "b > 0" "c > 0" "a < b" "c < d" + shows "a * c < b * d" +proof - + have "c < \" using \c < d\ by auto + then have "a * c < b * c" by (metis ereal_mult_strict_left_mono[OF assms(3) assms(2)] mult.commute) + moreover have "b * c \ b * d" using assms(2) assms(4) by (simp add: assms(1) ereal_mult_left_mono less_imp_le) + ultimately show ?thesis by simp +qed + +lemma ereal_mult_mono_strict': + fixes a b c d::ereal + assumes "a > 0" "c > 0" "a < b" "c < d" + shows "a * c < b * d" +apply (rule ereal_mult_mono_strict, auto simp add: assms) using assms by auto + lemma zero_less_one_ereal[simp]: "0 \ (1::ereal)" by (simp add: one_ereal_def zero_ereal_def) @@ -1484,6 +1513,12 @@ lemma ediff_le_self [simp]: "x - y \ (x :: enat)" by(cases x y rule: enat.exhaust[case_product enat.exhaust]) simp_all +lemma ereal_abs_diff: + fixes a b::ereal + shows "abs(a-b) \ abs a + abs b" +by (cases rule: ereal2_cases[of a b]) (auto) + + subsubsection \Division\ instantiation ereal :: inverse @@ -2263,7 +2298,7 @@ case False then show ?thesis by (subst continuous_at_Sup_mono[where f="\x. x + c"]) - (auto simp: continuous_at_imp_continuous_at_within continuous_at mono_def ereal_add_mono \I \ {}\ \c \ -\\) + (auto simp: continuous_at_imp_continuous_at_within continuous_at mono_def add_mono \I \ {}\ \c \ -\\) qed lemma SUP_ereal_add_right: @@ -2322,7 +2357,7 @@ apply (subst (2) add.commute) apply (subst SUP_ereal_add_left[symmetric, OF UNIV_not_empty assms(3)]) apply (subst (2) add.commute) - apply (rule SUP_combine[symmetric] ereal_add_mono inc[THEN monoD] | assumption)+ + apply (rule SUP_combine[symmetric] add_mono inc[THEN monoD] | assumption)+ done lemma INF_eq_minf: "(INF i:I. f i::ereal) \ -\ \ (\b>-\. \i\I. b \ f i)" @@ -2336,7 +2371,7 @@ unfolding INF_eq_minf using assms by (intro exI[of _ 0]) auto then show ?thesis by (subst continuous_at_Inf_mono[where f="\x. x + c"]) - (auto simp: mono_def ereal_add_mono \I \ {}\ \c \ -\\ continuous_at_imp_continuous_at_within continuous_at) + (auto simp: mono_def add_mono \I \ {}\ \c \ -\\ continuous_at_imp_continuous_at_within continuous_at) qed lemma INF_ereal_add_right: @@ -2357,7 +2392,7 @@ show ?thesis proof (rule antisym) show "(INF i:I. f i) + (INF i:I. g i) \ (INF i:I. f i + g i)" - by (rule INF_greatest; intro ereal_add_mono INF_lower) + by (rule INF_greatest; intro add_mono INF_lower) next have "(INF i:I. f i + g i) \ (INF i:I. (INF j:I. f i + g j))" using directed by (intro INF_greatest) (blast intro: INF_lower2) @@ -3093,12 +3128,6 @@ by (cases rule: ereal3_cases[of a b c]) (auto simp add: field_simps not_le mult_le_0_iff mult_less_0_iff) -lemma ereal_max_mono: "(a::ereal) \ b \ c \ d \ max a c \ max b d" - by (metis sup_ereal_def sup_mono) - -lemma ereal_max_least: "(a::ereal) \ x \ c \ x \ max a c \ x" - by (metis sup_ereal_def sup_least) - lemma ereal_LimI_finite: fixes x :: ereal assumes "\x\ \ \" @@ -3206,7 +3235,7 @@ with ev show "eventually ?P' F" by eventually_elim auto have "?INF P f + (SUP P:?F. ?INF P g) \ ?INF ?P' f + (SUP P:?F. ?INF P g)" - by (intro ereal_add_mono INF_mono) auto + by (intro add_mono INF_mono) auto also have "\ = (SUP P':?F. ?INF ?P' f + ?INF P' g)" proof (rule SUP_ereal_add_right[symmetric]) show "INFIMUM {x. P x \ 0 \ f x} f \ - \" @@ -3223,7 +3252,7 @@ show "(\x. P x \ Q x) \ ?F" using * by (auto simp: eventually_conj) show "?INF P f + ?INF Q g \ (INF x:{x. P x \ Q x}. f x + g x)" - by (intro INF_greatest ereal_add_mono) (auto intro: INF_lower) + by (intro INF_greatest add_mono) (auto intro: INF_lower) qed qed finally show "(SUP P:?F. ?INF P f + (SUP P:?F. ?INF P g)) \ (SUP P:?F. INF x:Collect P. f x + g x)" . @@ -3300,7 +3329,7 @@ shows "f sums (SUP n. \ii. \j=0..i. f i + g i) = suminf f + suminf g" apply (subst (1 2 3) suminf_ereal_eq_SUP) unfolding sum.distrib - apply (intro assms ereal_add_nonneg_nonneg SUP_ereal_add_pos incseq_sumI sum_nonneg ballI)+ + apply (intro assms add_nonneg_nonneg SUP_ereal_add_pos incseq_sumI sum_nonneg ballI)+ done lemma suminf_cmult_ereal: @@ -3656,7 +3685,7 @@ show ?thesis proof (rule antisym) show "(SUP i:I. f i + g i) \ (SUP i:I. f i) + (SUP i:I. g i)" - by (rule SUP_least; intro ereal_add_mono SUP_upper) + by (rule SUP_least; intro add_mono SUP_upper) next have "bot < (SUP i:I. g i)" using \I \ {}\ nonneg(2) by (auto simp: bot_ereal_def less_SUP_iff) @@ -3688,7 +3717,7 @@ fix i j assume "i \ I" "j \ I" from directed[OF \insert n N \ A\ this] guess k .. then show "\k\I. f n i + (\l\N. f l j) \ f n k + (\l\N. f l k)" - by (intro bexI[of _ k]) (auto intro!: ereal_add_mono sum_mono) + by (intro bexI[of _ k]) (auto intro!: add_mono sum_mono) qed ultimately show ?case by simp @@ -4104,7 +4133,7 @@ lemma Limsup_add_ereal_right: "F \ bot \ abs c \ \ \ Limsup F (\n. g n + (c :: ereal)) = Limsup F g + c" - by (rule Limsup_compose_continuous_mono) (auto simp: mono_def ereal_add_mono continuous_on_def) + by (rule Limsup_compose_continuous_mono) (auto simp: mono_def add_mono continuous_on_def) lemma Limsup_add_ereal_left: "F \ bot \ abs c \ \ \ Limsup F (\n. (c :: ereal) + g n) = c + Limsup F g" @@ -4112,7 +4141,7 @@ lemma Liminf_add_ereal_right: "F \ bot \ abs c \ \ \ Liminf F (\n. g n + (c :: ereal)) = Liminf F g + c" - by (rule Liminf_compose_continuous_mono) (auto simp: mono_def ereal_add_mono continuous_on_def) + by (rule Liminf_compose_continuous_mono) (auto simp: mono_def add_mono continuous_on_def) lemma Liminf_add_ereal_left: "F \ bot \ abs c \ \ \ Liminf F (\n. (c :: ereal) + g n) = c + Liminf F g"