# HG changeset patch # User paulson # Date 1456151876 0 # Node ID 340738057c8cf868f7a228e56da4910c43f179d8 # Parent 85ed00c1fe7c94571194be4bd56d4da28d9c1f2e An assortment of useful lemmas about sums, norm, etc. Also: norm_conv_dist [symmetric] is now a simprule! diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Complex.thy --- a/src/HOL/Complex.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Complex.thy Mon Feb 22 14:37:56 2016 +0000 @@ -167,10 +167,10 @@ lemma Im_divide_numeral [simp]: "Im (z / numeral w) = Im z / numeral w" by (simp add: Im_divide sqr_conv_mult) -lemma Re_divide_of_nat: "Re (z / of_nat n) = Re z / of_nat n" +lemma Re_divide_of_nat [simp]: "Re (z / of_nat n) = Re z / of_nat n" by (cases n) (simp_all add: Re_divide divide_simps power2_eq_square del: of_nat_Suc) -lemma Im_divide_of_nat: "Im (z / of_nat n) = Im z / of_nat n" +lemma Im_divide_of_nat [simp]: "Im (z / of_nat n) = Im z / of_nat n" by (cases n) (simp_all add: Im_divide divide_simps power2_eq_square del: of_nat_Suc) lemma of_real_Re [simp]: @@ -330,6 +330,12 @@ lemma cmod_plus_Re_le_0_iff: "cmod z + Re z \ 0 \ Re z = - cmod z" using abs_Re_le_cmod[of z] by auto +lemma cmod_Re_le_iff: "Im x = Im y \ cmod x \ cmod y \ abs (Re x) \ abs (Re y)" + by (metis add.commute add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) + +lemma cmod_Im_le_iff: "Re x = Re y \ cmod x \ cmod y \ abs (Im x) \ abs (Im y)" + by (metis add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) + lemma Im_eq_0: "\Re z\ = cmod z \ Im z = 0" by (subst (asm) power_eq_iff_eq_base[symmetric, where n=2]) (auto simp add: norm_complex_def) diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Conditionally_Complete_Lattices.thy --- a/src/HOL/Conditionally_Complete_Lattices.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Conditionally_Complete_Lattices.thy Mon Feb 22 14:37:56 2016 +0000 @@ -646,4 +646,44 @@ apply (metis eq cSUP_upper) done +lemma cSUP_UNION: + fixes f :: "_ \ 'b::conditionally_complete_lattice" + assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" + and bdd_UN: "bdd_above (\x\A. f ` B x)" + shows "(SUP z : \x\A. B x. f z) = (SUP x:A. SUP z:B x. f z)" +proof - + have bdd: "\x. x \ A \ bdd_above (f ` B x)" + using bdd_UN by (meson UN_upper bdd_above_mono) + obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" + using bdd_UN by (auto simp: bdd_above_def) + then have bdd2: "bdd_above ((\x. SUP z:B x. f z) ` A)" + unfolding bdd_above_def by (force simp: bdd cSUP_le_iff ne(2)) + have "(SUP z:\x\A. B x. f z) \ (SUP x:A. SUP z:B x. f z)" + using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper2 simp: bdd2 bdd) + moreover have "(SUP x:A. SUP z:B x. f z) \ (SUP z:\x\A. B x. f z)" + using assms by (fastforce simp add: intro!: cSUP_least intro: cSUP_upper simp: image_UN bdd_UN) + ultimately show ?thesis + by (rule order_antisym) +qed + +lemma cINF_UNION: + fixes f :: "_ \ 'b::conditionally_complete_lattice" + assumes ne: "A \ {}" "\x. x \ A \ B(x) \ {}" + and bdd_UN: "bdd_below (\x\A. f ` B x)" + shows "(INF z : \x\A. B x. f z) = (INF x:A. INF z:B x. f z)" +proof - + have bdd: "\x. x \ A \ bdd_below (f ` B x)" + using bdd_UN by (meson UN_upper bdd_below_mono) + obtain M where "\x y. x \ A \ y \ B(x) \ f y \ M" + using bdd_UN by (auto simp: bdd_below_def) + then have bdd2: "bdd_below ((\x. INF z:B x. f z) ` A)" + unfolding bdd_below_def by (force simp: bdd le_cINF_iff ne(2)) + have "(INF z:\x\A. B x. f z) \ (INF x:A. INF z:B x. f z)" + using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower simp: bdd2 bdd) + moreover have "(INF x:A. INF z:B x. f z) \ (INF z:\x\A. B x. f z)" + using assms by (fastforce simp add: intro!: cINF_greatest intro: cINF_lower2 simp: bdd bdd_UN bdd2) + ultimately show ?thesis + by (rule order_antisym) +qed + end diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Filter.thy --- a/src/HOL/Filter.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Filter.thy Mon Feb 22 14:37:56 2016 +0000 @@ -990,6 +990,9 @@ translations "LIM x F1. f :> F2" == "CONST filterlim (\x. f) F2 F1" +lemma filterlim_top [simp]: "filterlim f top F" + by (simp add: filterlim_def) + lemma filterlim_iff: "(LIM x F1. f x :> F2) \ (\P. eventually P F2 \ eventually (\x. P (f x)) F1)" unfolding filterlim_def le_filter_def eventually_filtermap .. diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Groups.thy --- a/src/HOL/Groups.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Groups.thy Mon Feb 22 14:37:56 2016 +0000 @@ -1203,6 +1203,16 @@ thus "\a\ \ 0" by simp qed +lemma abs_le_self_iff [simp]: "\a\ \ a \ 0 \ a" +proof - + have "\a. (0::'a) \ \a\" + using abs_ge_zero by blast + then have "\a\ \ a \ 0 \ a" + using order.trans by blast + then show ?thesis + using abs_of_nonneg eq_refl by blast +qed + lemma zero_less_abs_iff [simp]: "0 < \a\ \ a \ 0" by (simp add: less_le) diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Limits.thy --- a/src/HOL/Limits.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Limits.thy Mon Feb 22 14:37:56 2016 +0000 @@ -21,6 +21,14 @@ by (subst eventually_INF_base) (auto simp: subset_eq eventually_principal intro!: exI[of _ "max a b" for a b]) +corollary eventually_at_infinity_pos: + "eventually p at_infinity \ (\b. 0 < b \ (\x. norm x \ b \ p x))" +apply (simp add: eventually_at_infinity, auto) +apply (case_tac "b \ 0") +using norm_ge_zero order_trans zero_less_one apply blast +apply (force simp:) +done + lemma at_infinity_eq_at_top_bot: "(at_infinity :: real filter) = sup at_top at_bot" apply (simp add: filter_eq_iff eventually_sup eventually_at_infinity @@ -67,14 +75,14 @@ "Bfun f F \ (\K>0. eventually (\x. norm (f x) \ K) F)" unfolding Bfun_metric_def norm_conv_dist proof safe - fix y K assume "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" + fix y K assume K: "0 < K" and *: "eventually (\x. dist (f x) y \ K) F" moreover have "eventually (\x. dist (f x) 0 \ dist (f x) y + dist 0 y) F" by (intro always_eventually) (metis dist_commute dist_triangle) with * have "eventually (\x. dist (f x) 0 \ K + dist 0 y) F" by eventually_elim auto with \0 < K\ show "\K>0. eventually (\x. dist (f x) 0 \ K) F" by (intro exI[of _ "K + dist 0 y"] add_pos_nonneg conjI zero_le_dist) auto -qed auto +qed (force simp del: norm_conv_dist [symmetric]) lemma BfunI: assumes K: "eventually (\x. norm (f x) \ K) F" shows "Bfun f F" @@ -1036,10 +1044,9 @@ proof eventually_elim fix x assume A: "dist (f x) c < 1/2" and B: "norm (f x) \ norm c + 1" note B - also have "norm (f x) = dist (f x) 0" by (simp add: norm_conv_dist) + also have "norm (f x) = dist (f x) 0" by simp also have "... \ dist (f x) c + dist c 0" by (rule dist_triangle) - also note A - finally show False by (simp add: norm_conv_dist) + finally show False using A by simp qed with assms show False by simp qed @@ -1541,6 +1548,11 @@ shows "\(f \ a) F; ((\x. f x - g x) \ 0) F\ \ (g \ a) F" by (erule Lim_transform) (simp add: tendsto_minus_cancel) +proposition Lim_transform_eq: + fixes a :: "'a::real_normed_vector" + shows "((\x. f x - g x) \ 0) F \ (f \ a) F \ (g \ a) F" +using Lim_transform Lim_transform2 by blast + lemma Lim_transform_eventually: "eventually (\x. f x = g x) net \ (f \ l) net \ (g \ l) net" apply (rule topological_tendstoI) diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy --- a/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Bounded_Continuous_Function.thy Mon Feb 22 14:37:56 2016 +0000 @@ -259,7 +259,7 @@ by (rule dist_triangle_add) also have "\ \ dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" unfolding zero_bcontfun_def using assms - by (auto intro!: add_mono dist_bounded_Abs const_bcontfun) + by (metis add_mono const_bcontfun dist_bounded_Abs) finally show "dist (f x + g x) 0 \ dist (Abs_bcontfun f) 0 + dist (Abs_bcontfun g) 0" . qed (simp add: continuous_on_add) qed @@ -341,7 +341,7 @@ fix f g :: "('a, 'b) bcontfun" show "dist f g = norm (f - g)" by (simp add: norm_bcontfun_def dist_bcontfun_def zero_bcontfun_def - Abs_bcontfun_inverse const_bcontfun norm_conv_dist[symmetric] dist_norm) + Abs_bcontfun_inverse const_bcontfun dist_norm) show "norm (f + g) \ norm f + norm g" unfolding norm_bcontfun_def proof (subst dist_bcontfun_def, safe intro!: cSUP_least) @@ -364,31 +364,30 @@ proof (intro continuous_at_Sup_mono bdd_aboveI2) fix x show "dist (Rep_bcontfun f x) 0 \ norm f" using dist_bounded[of f x 0] - by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def + by (simp add: norm_bcontfun_def Abs_bcontfun_inverse zero_bcontfun_def const_bcontfun) qed (auto intro!: monoI mult_left_mono continuous_intros) moreover have "range (\x. dist (Rep_bcontfun (a *\<^sub>R f) x) 0) = (\x. \a\ * x) ` (range (\x. dist (Rep_bcontfun f x) 0))" - by (auto simp: norm_conv_dist[symmetric]) + by auto ultimately show "norm (a *\<^sub>R f) = \a\ * norm f" - by (simp add: norm_bcontfun_def dist_bcontfun_def norm_conv_dist Abs_bcontfun_inverse - zero_bcontfun_def const_bcontfun image_image) presburger + by (simp add: norm_bcontfun_def dist_bcontfun_def Abs_bcontfun_inverse + zero_bcontfun_def const_bcontfun image_image) qed qed (auto simp: norm_bcontfun_def sgn_bcontfun_def) end lemma bcontfun_normI: "continuous_on UNIV f \ (\x. norm (f x) \ b) \ f \ bcontfun" - unfolding norm_conv_dist - by (auto intro: bcontfunI) + by (metis bcontfunI dist_0_norm dist_commute) lemma norm_bounded: fixes f :: "('a::topological_space, 'b::real_normed_vector) bcontfun" shows "norm (Rep_bcontfun f x) \ norm f" using dist_bounded[of f x 0] - by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def + by (simp add: norm_bcontfun_def Abs_bcontfun_inverse zero_bcontfun_def const_bcontfun) lemma norm_bound: @@ -396,8 +395,7 @@ assumes "\x. norm (Rep_bcontfun f x) \ b" shows "norm f \ b" using dist_bound[of f 0 b] assms - by (simp add: norm_bcontfun_def norm_conv_dist Abs_bcontfun_inverse zero_bcontfun_def - const_bcontfun) + by (simp add: norm_bcontfun_def Abs_bcontfun_inverse zero_bcontfun_def const_bcontfun) subsection \Continuously Extended Functions\ diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Mon Feb 22 14:37:56 2016 +0000 @@ -6489,7 +6489,7 @@ if "z \ s" for z proof - obtain d h N where "0 < d" "summable h" and le_h: "\n y. \N \ n; y \ ball z d\ \ norm(f n y) \ h n" - using to_g \z \ s\ by blast + using to_g \z \ s\ by meson then obtain r where "r>0" and r: "ball z r \ ball z d \ s" using \z \ s\ s by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) have 1: "open (ball z d \ s)" diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Probability/Bochner_Integration.thy Mon Feb 22 14:37:56 2016 +0000 @@ -1229,7 +1229,7 @@ s: "\i. simple_function M (s i)" and pointwise: "\x. x \ space M \ (\i. s i x) \ f x" and bound: "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" - by (simp add: norm_conv_dist) metis + by simp metis show ?thesis proof (rule integrableI_sequence) @@ -2545,7 +2545,7 @@ then have s: "\i. simple_function (N \\<^sub>M M) (s i)" "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) \ f x y" "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" - by (auto simp: space_pair_measure norm_conv_dist) + by (auto simp: space_pair_measure) have [measurable]: "\i. s i \ borel_measurable (N \\<^sub>M M)" by (rule borel_measurable_simple_function) fact diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Product_Type.thy Mon Feb 22 14:37:56 2016 +0000 @@ -1184,6 +1184,14 @@ "snd ` (A \ B) = (if A = {} then {} else B)" by force +lemma fst_image_Sigma: + "fst ` (Sigma A B) = {x \ A. B(x) \ {}}" + by force + +lemma snd_image_Sigma: + "snd ` (Sigma A B) = (\ x \ A. B x)" + by force + lemma vimage_fst: "fst -` A = A \ UNIV" by auto diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Real_Vector_Spaces.thy --- a/src/HOL/Real_Vector_Spaces.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Real_Vector_Spaces.thy Mon Feb 22 14:37:56 2016 +0000 @@ -862,6 +862,14 @@ "norm (- numeral w::'a::real_normed_algebra_1) = numeral w" by (subst of_real_neg_numeral [symmetric], subst norm_of_real, simp) +lemma norm_of_real_add1 [simp]: + "norm (of_real x + 1 :: 'a :: real_normed_div_algebra) = abs (x + 1)" + by (metis norm_of_real of_real_1 of_real_add) + +lemma norm_of_real_addn [simp]: + "norm (of_real x + numeral b :: 'a :: real_normed_div_algebra) = abs (x + numeral b)" + by (metis norm_of_real of_real_add of_real_numeral) + lemma norm_of_int [simp]: "norm (of_int z::'a::real_normed_algebra_1) = \of_int z\" by (subst of_real_of_int_eq [symmetric], rule norm_of_real) @@ -1298,6 +1306,8 @@ lemma norm_conv_dist: "norm x = dist x 0" unfolding dist_norm by simp +declare norm_conv_dist [symmetric, simp] + lemma dist_diff [simp]: "dist a (a - b) = norm b" "dist (a - b) a = norm b" by (simp_all add: dist_norm) @@ -1868,7 +1878,7 @@ have "inverse e < of_nat (nat \inverse e + 1\)" by linarith also note n finally show "dist (1 / of_nat n :: 'a) 0 < e" using e - by (simp add: divide_simps mult.commute norm_conv_dist[symmetric] norm_divide) + by (simp add: divide_simps mult.commute norm_divide) qed lemma (in metric_space) complete_def: diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Series.thy --- a/src/HOL/Series.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Series.thy Mon Feb 22 14:37:56 2016 +0000 @@ -390,6 +390,9 @@ by (simp add: ac_simps) qed simp +corollary sums_iff_shift': "(\i. f (i + n)) sums (s - (\i f sums s" + by (simp add: sums_iff_shift) + lemma summable_iff_shift: "summable (\n. f (n + k)) \ summable f" by (metis diff_add_cancel summable_def sums_iff_shift[abs_def]) @@ -535,6 +538,12 @@ lemma suminf_divide: "summable f \ suminf (\n. f n / c) = suminf f / c" by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric]) +lemma sums_mult_D: "\(\n. c * f n) sums a; c \ 0\ \ f sums (a/c)" + using sums_mult_iff by fastforce + +lemma summable_mult_D: "\summable (\n. c * f n); c \ 0\ \ summable f" + by (auto dest: summable_divide) + text\Sum of a geometric progression.\ lemma geometric_sums: "norm c < 1 \ (\n. c^n) sums (1 / (1 - c))" diff -r 85ed00c1fe7c -r 340738057c8c src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Fri Feb 19 13:40:50 2016 +0100 +++ b/src/HOL/Set_Interval.thy Mon Feb 22 14:37:56 2016 +0000 @@ -1640,6 +1640,11 @@ finally show ?case . qed +lemma setsum_atMost_shift: + fixes f :: "nat \ 'a::comm_monoid_add" + shows "(\i\n. f i) = f 0 + (\i (\i = m..n. f i) = f n + (\i = m..Properties of Power Series\ -lemma powser_zero: +lemma powser_zero [simp]: fixes f :: "nat \ 'a::real_normed_algebra_1" shows "(\n. f n * 0 ^ n) = f 0" proof - @@ -89,6 +89,11 @@ using sums_finite [of "{0}" "\n. a n * 0 ^ n"] by simp +lemma powser_sums_zero_iff [simp]: + fixes a :: "nat \ 'a::real_normed_div_algebra" + shows "(\n. a n * 0^n) sums x \ a 0 = x" +using powser_sums_zero sums_unique2 by blast + text\Power series has a circle or radius of convergence: if it sums for @{term x}, then it sums absolutely for @{term z} with @{term "\z\ < \x\"}.\ @@ -1342,6 +1347,31 @@ lemma exp_setsum: "finite I \ exp(setsum f I) = setprod (\x. exp(f x)) I" by (induction I rule: finite_induct) (auto simp: exp_add_commuting mult.commute) +lemma exp_divide_power_eq: + fixes x:: "'a::{real_normed_field,banach}" + assumes "n>0" shows "exp (x / of_nat n) ^ n = exp x" +using assms +proof (induction n arbitrary: x) + case 0 then show ?case by simp +next + case (Suc n) + show ?case + proof (cases "n=0") + case True then show ?thesis by simp + next + case False + then have [simp]: "x * of_nat n / (1 + of_nat n) / of_nat n = x / (1 + of_nat n)" + by simp + have [simp]: "x / (1 + of_nat n) + x * of_nat n / (1 + of_nat n) = x" + apply (simp add: divide_simps) + using of_nat_eq_0_iff apply (fastforce simp: distrib_left) + done + show ?thesis + using Suc.IH [of "x * of_nat n / (1 + of_nat n)"] False + by (simp add: exp_add [symmetric]) + qed +qed + subsubsection \Properties of the Exponential Function on Reals\