diff --git a/thys/Fourier/Fourier.thy b/thys/Fourier/Fourier.thy --- a/thys/Fourier/Fourier.thy +++ b/thys/Fourier/Fourier.thy @@ -1,2740 +1,2742 @@ section\The basics of Fourier series\ text\Ported from HOL Light; thanks to Manuel Eberl for help with the real asymp proof methods\ theory "Fourier" - imports Periodic Square_Integrable "HOL-Real_Asymp.Real_Asymp" Confine + imports Periodic Square_Integrable "HOL-Real_Asymp.Real_Asymp" Confine Fourier_Aux2 begin subsection\Orthonormal system of L2 functions and their Fourier coefficients\ definition orthonormal_system :: "'a::euclidean_space set \ ('b \ 'a \ real) \ bool" where "orthonormal_system S w \ \m n. l2product S (w m) (w n) = (if m = n then 1 else 0)" definition orthonormal_coeff :: "'a::euclidean_space set \ (nat \ 'a \ real) \ ('a \ real) \ nat \ real" where "orthonormal_coeff S w f n = l2product S (w n) f" lemma orthonormal_system_eq: "orthonormal_system S w \ l2product S (w m) (w n) = (if m = n then 1 else 0)" by (simp add: orthonormal_system_def) lemma orthonormal_system_l2norm: "orthonormal_system S w \ l2norm S (w i) = 1" by (simp add: l2norm_def orthonormal_system_def) lemma orthonormal_partial_sum_diff: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x - (\i\I. a i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 + (\i\I. (a i)\<^sup>2) - 2 * (\i\I. a i * orthonormal_coeff S w f i)" proof - have "S \ sets lebesgue" using f square_integrable_def by blast then have "(\x. \i\I. a i * w i x) square_integrable S" by (intro square_integrable_sum square_integrable_lmult w \finite I\) with assms show ?thesis apply (simp add: square_integrable_diff l2norm_pow_2 orthonormal_coeff_def orthonormal_system_def) apply (simp add: square_integrable_diff square_integrable_lmult l2product_rdiff l2product_sym l2product_rsum l2product_rmult algebra_simps power2_eq_square if_distrib [of "\x. _ * x"]) done qed lemma orthonormal_optimal_partial_sum: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "l2norm S (\x. f x - (\i\I. orthonormal_coeff S w f i * w i x)) \ l2norm S (\x. f x - (\i\I. a i * w i x))" proof - have "2 * (a i * l2product S f(w i)) \ (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2" for i using sum_squares_bound [of "a i" "l2product S f(w i)"] by (force simp: algebra_simps) then have *: "2 * (\i\I. a i * l2product S f(w i)) \ (\i\I. (a i)\<^sup>2 + (l2product S f(w i))\<^sup>2)" by (simp add: sum_distrib_left sum_mono) have S: "S \ sets lebesgue" using assms square_integrable_def by blast with assms * show ?thesis apply (simp add: l2norm_le square_integrable_sum square_integrable_diff square_integrable_lmult flip: l2norm_pow_2) apply (simp add: orthonormal_coeff_def orthonormal_partial_sum_diff) apply (simp add: l2product_sym power2_eq_square sum.distrib) done qed lemma Bessel_inequality: assumes "orthonormal_system S w" "\i. (w i) square_integrable S" "f square_integrable S" "finite I" shows "(\i\I. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" using orthonormal_partial_sum_diff [OF assms, of "orthonormal_coeff S w f"] apply (simp add: algebra_simps flip: power2_eq_square) by (metis (lifting) add_diff_cancel_left' diff_ge_0_iff_ge zero_le_power2) lemma Fourier_series_square_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" shows "summable (confine (\i. (orthonormal_coeff S w f i) ^ 2) I)" proof - have "incseq (\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2)" by (auto simp: incseq_def intro: sum_mono2) moreover have "\i. (\i\I \ {..i}. (orthonormal_coeff S w f i)\<^sup>2) \ (l2norm S f)\<^sup>2" by (simp add: Bessel_inequality assms) ultimately obtain L where "(\n. \i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) \ L" using incseq_convergent by blast then have "\r>0. \no. \n\no. \(\i\{..n} \ I. (orthonormal_coeff S w f i)\<^sup>2) - L\ < r" by (simp add: LIMSEQ_iff Int_commute) then show ?thesis by (auto simp: summable_def sums_confine_le LIMSEQ_iff) qed lemma orthonormal_Fourier_partial_sum_diff_squared: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" and "finite I" shows "(l2norm S (\x. f x -(\i\I. orthonormal_coeff S w f i * w i x)))\<^sup>2 = (l2norm S f)\<^sup>2 - (\i\I. (orthonormal_coeff S w f i)\<^sup>2)" using orthonormal_partial_sum_diff [OF assms, where a = "orthonormal_coeff S w f"] by (simp add: power2_eq_square) lemma Fourier_series_l2_summable: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?\ = "confine (\i. (orthonormal_coeff S w f i)\<^sup>2) I" have si: "?\ A square_integrable S" if "finite A" for A by (force intro: square_integrable_lmult w square_integrable_sum S that) have "\N. \m\N. \n\N. l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" if "e > 0" for e proof - have "summable ?\" by (rule Fourier_series_square_summable [OF os w f]) then have "Cauchy (\n. sum ?\ {..n})" by (simp add: summable_def sums_def_le convergent_eq_Cauchy) then obtain M where M: "\m n. \m\M; n\M\ \ dist (sum ?\ {..m}) (sum ?\ {..n}) < e\<^sup>2" using \e > 0\ unfolding Cauchy_def by (drule_tac x="e^2" in spec) auto have "\m\M; n\M\ \ l2norm S (\x. ?\ {..m} x - ?\ {..n} x) < e" for m n proof (induct m n rule: linorder_class.linorder_wlog) case (le m n) have sum_diff: "sum f(I \ {..n}) - sum f(I \ {..m}) = sum f(I \ {Suc m..n})" for f :: "nat \ real" proof - have "(I \ {..n}) = (I \ {..m} \ I \ {Suc m..n})" "(I \ {..m}) \ (I \ {Suc m..n}) = {}" using le by auto then show ?thesis by (simp add: algebra_simps flip: sum.union_disjoint) qed have "(l2norm S (\x. ?\ {..n} x - ?\ {..m} x))^2 \ \(\i\I \ {..n}. (orthonormal_coeff S w f i)\<^sup>2) - (\i\I \ {..m}. (orthonormal_coeff S w f i)\<^sup>2)\" proof (simp add: sum_diff) have "(l2norm S (?\ {Suc m..n}))\<^sup>2 = (\i\I \ {Suc m..n}. l2product S (\x. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * w j x) (\x. orthonormal_coeff S w f i * w i x))" by (simp add: l2norm_pow_2 l2product_rsum si w) also have "\ = (\i\I \ {Suc m..n}. \j\I \ {Suc m..n}. orthonormal_coeff S w f j * orthonormal_coeff S w f i * l2product S (w j) (w i))" by (simp add: l2product_lsum l2product_lmult l2product_rmult si w flip: mult.assoc) also have "\ \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" by (simp add: orthonormal_system_eq [OF os] power2_eq_square if_distrib [of "(*)_"] cong: if_cong) finally show "(l2norm S (?\ {Suc m..n}))\<^sup>2 \ \\i\I \ {Suc m..n}. (orthonormal_coeff S w f i)\<^sup>2\" . qed also have "\ < e\<^sup>2" using M [OF \n\M\ \m\M\] by (simp add: dist_real_def sum_confine_eq_Int Int_commute) finally have "(l2norm S (\x. ?\ {..m} x - ?\ {..n} x))^2 < e^2" using l2norm_diff [OF si si] by simp with \e > 0\ show ?case by (simp add: power2_less_imp_less) next case (sym a b) then show ?case by (subst l2norm_diff) (auto simp: si) qed then show ?thesis by blast qed with that show thesis by (blast intro: si [of "{.._}"] l2_complete [of "\n. ?\ {..n}"]) qed lemma Fourier_series_l2_summable_strong: assumes os: "orthonormal_system S w" and w: "\i. (w i) square_integrable S" and f: "f square_integrable S" obtains g where "g square_integrable S" "\i. i \ I \ orthonormal_coeff S w (\x. f x - g x) i = 0" "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" proof - have S: "S \ sets lebesgue" using f square_integrable_def by blast obtain g where g: "g square_integrable S" and teng: "(\n. l2norm S (\x. (\i\I \ {..n}. orthonormal_coeff S w f i * w i x) - g x)) \ 0" using Fourier_series_l2_summable [OF assms] . show thesis proof show "orthonormal_coeff S w (\x. f x - g x) i = 0" if "i \ I" for i proof (rule tendsto_unique) let ?\ = "\A x. (\i\I \ A. orthonormal_coeff S w f i * w i x)" let ?f = "\n. l2product S (w i) (\x. (f x - ?\ {..n} x) + (?\ {..n} x - g x))" show "?f \ orthonormal_coeff S w (\x. f x - g x) i" by (simp add: orthonormal_coeff_def) have 1: "(\n. l2product S (w i) (\x. f x - ?\ {..n} x)) \ 0" proof (rule tendsto_eventually) have "l2product S (w i) (\x. f x - ?\ {..j} x) = 0" if "j \ i" for j using that \i \ I\ apply (simp add: l2product_rdiff l2product_rsum l2product_rmult orthonormal_coeff_def w f S) apply (simp add: orthonormal_system_eq [OF os] if_distrib [of "(*)_"] cong: if_cong) done then show "\\<^sub>F n in sequentially. l2product S (w i) (\x. f x - ?\ {..n} x) = 0" using eventually_at_top_linorder by blast qed have 2: "(\n. l2product S (w i) (\x. ?\ {..n} x - g x)) \ 0" proof (intro Lim_null_comparison [OF _ teng] eventuallyI) show "norm (l2product S (w i) (\x. ?\ {..n} x - g x)) \ l2norm S (\x. ?\ {..n} x - g x)" for n using Schwartz_inequality_abs [of "w i" S "(\x. ?\ {..n} x - g x)"] by (simp add: S g f w orthonormal_system_l2norm [OF os]) qed show "?f \ 0" using that tendsto_add [OF 1 2] by (subst l2product_radd) (simp_all add: l2product_radd w f g S) qed auto qed (auto simp: g teng) qed - subsection\Actual trigonometric orthogonality relations\ lemma integrable_sin_cx: "integrable (lebesgue_on {-pi..pi}) (\x. sin(x * c))" by (intro continuous_imp_integrable_real continuous_intros) lemma integrable_cos_cx: "integrable (lebesgue_on {-pi..pi}) (\x. cos(x * c))" by (intro continuous_imp_integrable_real continuous_intros) +lemma integral_cos_Z' [simp]: + assumes "n \ \" + shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(n * x)) = (if n = 0 then 2 * pi else 0)" + by (metis assms integral_cos_Z mult_commute_abs) + lemma integral_sin_and_cos: fixes m n::int shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\m \ 0; n \ 0\ \ integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" by (simp_all add: abs_if sin_times_sin cos_times_sin sin_times_cos cos_times_cos integrable_sin_cx integrable_cos_cx mult_ac flip: distrib_left distrib_right left_diff_distrib right_diff_distrib) lemma integral_sin_and_cos_Z [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * cos(n * x)) = (if \m\ = \n\ then if n = 0 then 2 * pi else pi else 0)" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(m * x) * sin(n * x)) = 0" "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(m * x) * cos(n * x)) = 0" "\integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x))\ = (if \m\ = \n\ \ n \ 0 then pi else 0)" using assms unfolding Ints_def apply safe unfolding integral_sin_and_cos apply auto done lemma integral_sin_and_cos_N [simp]: fixes m n::real assumes "m \ \" "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin (m * x) * sin (n * x)) = (if m = n \ n \ 0 then pi else 0)" using assms unfolding Nats_altdef1 by (auto simp: integral_sin_and_cos) lemma integrable_sin_and_cos: fixes m n::int shows "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. cos(x * m) * sin(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * cos(x * n))" "integrable (lebesgue_on {a..b}) (\x. sin(x * m) * sin(x * n))" by (intro continuous_imp_integrable_real continuous_intros)+ lemma sqrt_pi_ge1: "sqrt pi \ 1" using pi_gt3 by auto definition trigonometric_set :: "nat \ real \ real" where "trigonometric_set n \ if n = 0 then \x. 1 / sqrt(2 * pi) else if odd n then \x. sin(real(Suc (n div 2)) * x) / sqrt(pi) else (\x. cos((n div 2) * x) / sqrt pi)" lemma trigonometric_set: "trigonometric_set 0 x = 1 / sqrt(2 * pi)" "trigonometric_set (Suc (2 * n)) x = sin(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (2 * n + 2) x = cos(real(Suc n) * x) / sqrt(pi)" "trigonometric_set (Suc (Suc (2 * n))) x = cos(real(Suc n) * x) / sqrt(pi)" by (simp_all add: trigonometric_set_def algebra_simps add_divide_distrib) lemma trigonometric_set_even: "trigonometric_set(2*k) = (if k = 0 then (\x. 1 / sqrt(2 * pi)) else (\x. cos(k * x) / sqrt pi))" by (induction k) (auto simp: trigonometric_set_def add_divide_distrib split: if_split_asm) - lemma orthonormal_system_trigonometric_set: "orthonormal_system {-pi..pi} trigonometric_set" proof - have "l2product {-pi..pi} (trigonometric_set m) (trigonometric_set n) = (if m = n then 1 else 0)" for m n proof (induction m rule: odd_even_cases) case 0 show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def) next case (odd m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def double_not_eq_Suc_double) next case (even m) show ?case by (induction n rule: odd_even_cases) (auto simp: trigonometric_set l2product_def Suc_double_not_eq_double) qed then show ?thesis unfolding orthonormal_system_def by auto qed lemma square_integrable_trigonometric_set: "(trigonometric_set i) square_integrable {-pi..pi}" proof - have "continuous_on {-pi..pi} (\x. sin ((1 + real n) * x) / sqrt pi)" for n by (intro continuous_intros) auto moreover have "continuous_on {-pi..pi} (\x. cos (real i * x / 2) / sqrt pi) " by (intro continuous_intros) auto ultimately show ?thesis by (simp add: trigonometric_set_def) qed subsection\Weierstrass for trigonometric polynomials\ lemma Weierstrass_trig_1: fixes g :: "real \ real" assumes contf: "continuous_on UNIV g" and periodic: "\x. g(x + 2 * pi) = g x" and 1: "norm z = 1" shows "continuous (at z within (sphere 0 1)) (g \ Im \ Ln)" proof (cases "Re z < 0") let ?f = "complex_of_real \ g \ (\x. x + pi) \ Im \ Ln \ uminus" case True have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" proof (rule continuous_transform_within) have [simp]: "z \ \\<^sub>\\<^sub>0" using True complex_nonneg_Reals_iff by auto show "continuous (at z within (sphere 0 1)) ?f" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto show "?f x' = (complex_of_real \ g \ Im \ Ln) x'" if "x' \ (sphere 0 1)" and "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto with that show ?thesis by (auto simp: Ln_minus add.commute periodic) qed qed (use 1 in auto) then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) next case False let ?f = "complex_of_real \ g \ Im \ Ln \ uminus" have "z \ 0" using 1 by auto with False have "z \ \\<^sub>\\<^sub>0" by (auto simp: not_less nonpos_Reals_def) then have "continuous (at z within (sphere 0 1)) (complex_of_real \ g \ Im \ Ln)" by (intro continuous_within_Ln continuous_intros continuous_on_imp_continuous_within [OF contf]) auto then have "continuous (at z within (sphere 0 1)) (Re \ complex_of_real \ g \ Im \ Ln)" unfolding o_def by (rule continuous_Re) then show ?thesis by (simp add: o_def) qed inductive_set cx_poly :: "(complex \ real) set" where Re: "Re \ cx_poly" | Im: "Im \ cx_poly" | const: "(\x. c) \ cx_poly" | add: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x + g x) \ cx_poly" | mult: "\f \ cx_poly; g \ cx_poly\ \ (\x. f x * g x) \ cx_poly" declare cx_poly.intros [intro] lemma Weierstrass_trig_polynomial: assumes contf: "continuous_on {-pi..pi} f" and fpi: "f(-pi) = f pi" and "0 < e" obtains n::nat and a b where "\x::real. x \ {-pi..pi} \ \f x - (\k\n. a k * sin (k * x) + b k * cos (k * x))\ < e" proof - interpret CR: function_ring_on "cx_poly" "sphere 0 1" proof show "continuous_on (sphere 0 1) f" if "f \ cx_poly" for f using that by induction (assumption | intro continuous_intros)+ fix x y::complex assume "x \ sphere 0 1" and "y \ sphere 0 1" and "x \ y" then consider "Re x \ Re y" | "Im x \ Im y" using complex_eqI by blast then show "\f\cx_poly. f x \ f y" by (metis cx_poly.Im cx_poly.Re) qed (auto simp: cx_poly.intros) have "continuous (at z within (sphere 0 1)) (f \ Im \ Ln)" if "norm z = 1" for z proof - obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ {-pi..pi} \ g x = f x" and periodic: "\x. g(x + 2*pi) = g x" using Tietze_periodic_interval [OF contf fpi] by auto let ?f = "(g \ Im \ Ln)" show ?thesis proof (rule continuous_transform_within) show "continuous (at z within (sphere 0 1)) ?f" using Weierstrass_trig_1 [OF contg periodic that] by (simp add: sphere_def) show "?f x' = (f \ Im \ Ln) x'" if "x' \ sphere 0 1" "dist x' z < 1" for x' proof - have "x' \ 0" using that by auto then have "Im (Ln x') \ {-pi..pi}" using Im_Ln_le_pi [of x'] mpi_less_Im_Ln [of x'] by simp then show ?thesis using gf by simp qed qed (use that in auto) qed then have "continuous_on (sphere 0 1) (f \ Im \ Ln)" using continuous_on_eq_continuous_within mem_sphere_0 by blast then obtain g where "g \ cx_poly" and g: "\x. x \ sphere 0 1 \ \(f \ Im \ Ln) x - g x\ < e" using CR.Stone_Weierstrass_basic [of "f \ Im \ Ln"] \e > 0\ by meson define trigpoly where "trigpoly \ \f. \n a b. f = (\x. (\k\n. a k * sin(real k * x) + b k * cos(real k * x)))" have tp_const: "trigpoly(\x. c)" for c unfolding trigpoly_def by (rule_tac x=0 in exI) auto have tp_add: "trigpoly(\x. f x + g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast let ?a = "\n. (if n \ n1 then a1 n else 0) + (if n \ n2 then a2 n else 0)" let ?b = "\n. (if n \ n1 then b1 n else 0) + (if n \ n2 then b2 n else 0)" have eq: "{k. k \ max a b \ k \ a} = {..a}" "{k. k \ max a b \ k \ b} = {..b}" for a b::nat by auto have "(\x. f x + g x) = (\x. \k \ max n1 n2. ?a k * sin (real k * x) + ?b k * cos (real k * x))" by (simp add: feq geq algebra_simps eq sum.distrib if_distrib [of "\u. _ * u"] cong: if_cong flip: sum.inter_filter) then show ?thesis unfolding trigpoly_def by meson qed have tp_sum: "trigpoly(\x. \i\S. f i x)" if "finite S" "\i. i \ S \ trigpoly(f i)" for f and S :: "nat set" using that by induction (auto simp: tp_const tp_add) have tp_cmul: "trigpoly(\x. c * f x)" if "trigpoly f" for f c proof - obtain n a b where feq: "f = (\x. \k\n. a k * sin (real k * x) + b k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast have "(\x. c * f x) = (\x. \k \ n. (c * a k) * sin (real k * x) + (c * b k) * cos (real k * x))" by (simp add: feq algebra_simps sum_distrib_left) then show ?thesis unfolding trigpoly_def by meson qed have tp_cdiv: "trigpoly(\x. f x / c)" if "trigpoly f" for f c using tp_cmul [OF \trigpoly f\, of "inverse c"] by (simp add: divide_inverse_commute) have tp_diff: "trigpoly(\x. f x - g x)" if "trigpoly f" "trigpoly g" for f g using tp_add [OF \trigpoly f\ tp_cmul [OF \trigpoly g\, of "-1"]] by auto have tp_sin: "trigpoly(\x. sin (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. sin (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (rule_tac x="\i. 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_cos: "trigpoly(\x. cos (n * x))" if "n \ \" for n unfolding trigpoly_def proof obtain k where "n = real k" using Nats_cases \n \ \\ by blast then show "\a b. (\x. cos (n * x)) = (\x. \i \ nat\n\. a i * sin (real i * x) + b i * cos (real i * x))" apply (rule_tac x="\i. 0" in exI) apply (rule_tac x="\i. if i = k then 1 else 0" in exI) apply (simp add: if_distrib [of "\u. u * _"] cong: if_cong) done qed have tp_sincos: "trigpoly(\x. sin(i * x) * sin(j * x)) \ trigpoly(\x. sin(i * x) * cos(j * x)) \ trigpoly(\x. cos(i * x) * sin(j * x)) \ trigpoly(\x. cos(i * x) * cos(j * x))" (is "?P i j") for i j::nat proof (rule linorder_wlog [of _ j i]) show "?P j i" if "i \ j" for j i using that by (auto simp add: sin_times_sin cos_times_cos sin_times_cos cos_times_sin diff_divide_distrib simp flip: of_nat_diff left_diff_distrib distrib_right intro!: tp_add tp_diff tp_cdiv tp_cos tp_sin) qed (simp add: mult_ac) have tp_mult: "trigpoly(\x. f x * g x)" if "trigpoly f" "trigpoly g" for f g proof - obtain n1 a1 b1 where feq: "f = (\x. \k\n1. a1 k * sin (real k * x) + b1 k * cos (real k * x))" using \trigpoly f\ trigpoly_def by blast obtain n2 a2 b2 where geq: "g = (\x. \k\n2. a2 k * sin (real k * x) + b2 k * cos (real k * x))" using \trigpoly g\ trigpoly_def by blast then show ?thesis unfolding feq geq by (simp add: algebra_simps sum_product tp_sum tp_add tp_cmul tp_sincos del: mult.commute) qed have *: "trigpoly(\x. f(exp(\ * of_real x)))" if "f \ cx_poly" for f using that proof induction case Re then show ?case using tp_cos [of 1] by (auto simp: Re_exp) next case Im then show ?case using tp_sin [of 1] by (auto simp: Im_exp) qed (auto simp: tp_const tp_add tp_mult) obtain n a b where eq: "(g (iexp x)) = (\k\n. a k * sin (real k * x) + b k * cos (real k * x))" for x using * [OF \g \ cx_poly\] trigpoly_def by meson show thesis proof show "\f \ - (\k\n. a k * sin (real k * \) + b k * cos (real k * \))\ < e" if "\ \ {-pi..pi}" for \ proof - have "f \ - g (iexp \) = (f \ Im \ Ln) (iexp \) - g (iexp \)" proof (cases "\ = -pi") case True then show ?thesis by (simp add: exp_minus fpi) next case False then show ?thesis using that by auto qed then show ?thesis using g [of "exp(\ * of_real \)"] by (simp flip: eq) qed qed qed subsection\A bit of extra hacking round so that the ends of a function are OK\ lemma integral_tweak_ends: fixes a b :: real assumes "a < b" "e > 0" obtains f where "continuous_on {a..b} f" "f a = d" "f b = 0" "l2norm {a..b} f < e" proof - have cont: "continuous_on {a..b} (\x. if x \ a+1 / real(Suc n) then ((Suc n) * d) * ((a+1 / (Suc n)) - x) else 0)" for n proof (cases "a+1/(Suc n) \ b") case True have *: "1 / (1 + real n) > 0" by auto have abeq: "{a..b} = {a..a+1/(Suc n)} \ {a+1/(Suc n)..b}" using \a < b\ True apply auto using "*" by linarith show ?thesis unfolding abeq proof (rule continuous_on_cases) show "continuous_on {a..a+1 / real (Suc n)} (\x. real (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed auto next case False show ?thesis proof (rule continuous_on_eq [where f = "\x. ((Suc n) * d) * ((a+1/(Suc n)) - x)"]) show " continuous_on {a..b} (\x. (Suc n) * d * (a+1 / real (Suc n) - x))" by (intro continuous_intros) qed (use False in auto) qed let ?f = "\k x. (if x \ a+1 / (Suc k) then (Suc k) * d * (a+1 / (Suc k) - x) else 0)\<^sup>2" let ?g = "\x. if x = a then d\<^sup>2 else 0" have bmg: "?g \ borel_measurable (lebesgue_on {a..b})" apply (rule measurable_restrict_space1) using borel_measurable_if_I [of _ "{a}", OF borel_measurable_const] by auto have bmf: "?f k \ borel_measurable (lebesgue_on {a..b})" for k proof - have bm: "(\x. (Suc k) * d * (a+1 / real (Suc k) - x)) \ borel_measurable (lebesgue_on {..a+1 / (Suc k)})" by (intro measurable) (auto simp: measurable_completion measurable_restrict_space1) show ?thesis apply (intro borel_measurable_power measurable_restrict_space1) using borel_measurable_if_I [of _ "{.. a+1 / (Suc k)}", OF bm] apply auto done qed have int_d2: "integrable (lebesgue_on {a..b}) (\x. d\<^sup>2)" by (intro continuous_imp_integrable_real continuous_intros) have "(\k. ?f k x) \ ?g x" if x: "x \ {a..b}" for x proof (cases "x = a") case False then have "x > a" using x by auto with x obtain N where "N > 0" and N: "1 / real N < x-a" using real_arch_invD [of "x-a"] by (force simp: inverse_eq_divide) then have "x > a+1 / (1 + real k)" if "k \ N" for k proof - have "a+1 / (1 + real k) < a+1 / real N" using that \0 < N\ by (simp add: field_simps) also have "\ < x" using N by linarith finally show ?thesis . qed then show ?thesis apply (intro tendsto_eventually eventually_sequentiallyI [where c=N]) by (fastforce simp: False) qed auto then have tends: "AE x in (lebesgue_on {a..b}). (\k. ?f k x) \ ?g x" by force have le_d2: "\k. AE x in (lebesgue_on {a..b}). norm (?f k x) \ d\<^sup>2" proof show "norm ((if x \ a+1 / real (Suc k) then real (Suc k) * d * (a+1 / real (Suc k) - x) else 0)\<^sup>2) \ d\<^sup>2" if "x \ space (lebesgue_on {a..b})" for k x using that apply (simp add: abs_mult divide_simps flip: abs_le_square_iff) apply (auto simp: abs_if zero_less_mult_iff mult_left_le) done qed have integ: "integrable (lebesgue_on {a..b}) ?g" using integrable_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto have int: "(\k. integral\<^sup>L (lebesgue_on {a..b}) (?f k)) \ integral\<^sup>L (lebesgue_on {a..b}) ?g" using integral_dominated_convergence [OF bmg bmf int_d2 tends le_d2] by auto then obtain N where N: "\k. k \ N \ \integral\<^sup>L (lebesgue_on {a..b}) (?f k) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" apply (simp add: lim_sequentially dist_real_def) apply (drule_tac x="e^2" in spec) using \e > 0\ by auto obtain M where "M > 0" and M: "1 / real M < b-a" using real_arch_invD [of "b-a"] by (metis \a < b\ diff_gt_0_iff_gt inverse_eq_divide neq0_conv) have *: "\integral\<^sup>L (lebesgue_on {a..b}) (?f (max M N)) - integral\<^sup>L (lebesgue_on {a..b}) ?g\ < e\<^sup>2" using N by force let ?\ = "\x. if x \ a+1/(Suc (max M N)) then ((Suc (max M N)) * d) * ((a+1/(Suc (max M N))) - x) else 0" show ?thesis proof show "continuous_on {a..b} ?\" by (rule cont) have "1 / (1 + real (max M N)) \ 1 / (real M)" by (simp add: \0 < M\ frac_le) then have "\ (b \ a+1 / (1 + real (max M N)))" using M \a < b\ \M > 0\ max.cobounded1 [of M N] by linarith then show "?\ b = 0" by simp have null_a: "{a} \ null_sets (lebesgue_on {a..b})" using \a < b\ by (simp add: null_sets_restrict_space) have "LINT x|lebesgue_on {a..b}. ?g x = 0" - using null_AE_impI [OF null_a] by (force intro: integral_eq_zero_AE) + by (force intro: integral_eq_zero_AE AE_I' [OF null_a]) then have "l2norm {a..b} ?\ < sqrt (e^2)" unfolding l2norm_def l2product_def power2_eq_square [symmetric] apply (intro real_sqrt_less_mono) using "*" by linarith then show "l2norm {a..b} ?\ < e" using \e > 0\ by auto qed auto qed lemma square_integrable_approximate_continuous_ends: assumes f: "f square_integrable {a..b}" and "a < b" "0 < e" obtains g where "continuous_on {a..b} g" "g b = g a" "g square_integrable {a..b}" "l2norm {a..b} (\x. f x - g x) < e" proof - obtain g where contg: "continuous_on UNIV g" and "g square_integrable {a..b}" and lg: "l2norm {a..b} (\x. f x - g x) < e/2" using f \e > 0\ square_integrable_approximate_continuous by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) obtain h where conth: "continuous_on {a..b} h" and "h a = g b - g a" "h b = 0" and lh: "l2norm {a..b} h < e/2" using integral_tweak_ends \e > 0\ by (metis \a < b\ zero_less_divide_iff zero_less_numeral) have "h square_integrable {a..b}" using \continuous_on {a..b} h\ continuous_imp_square_integrable by blast show thesis proof show "continuous_on {a..b} (\x. g x + h x)" by (blast intro: continuous_on_subset [OF contg] conth continuous_intros) then show "(\x. g x + h x) square_integrable {a..b}" using continuous_imp_square_integrable by blast show "g b + h b = g a + h a" by (simp add: \h a = g b - g a\ \h b = 0\) have "l2norm {a..b} (\x. (f x - g x) + - h x) < e" proof (rule le_less_trans [OF l2norm_triangle [of "\x. f x - g x" "{a..b}" "\x. - (h x)"]]) show "(\x. f x - g x) square_integrable {a..b}" using \g square_integrable {a..b}\ f square_integrable_diff by blast show "(\x. - h x) square_integrable {a..b}" by (simp add: \h square_integrable {a..b}\) show "l2norm {a..b} (\x. f x - g x) + l2norm {a..b} (\x. - h x) < e" using \h square_integrable {a..b}\ l2norm_neg lg lh by auto qed then show "l2norm {a..b} (\x. f x - (g x + h x)) < e" by (simp add: field_simps) qed qed subsection\Hence the main approximation result\ lemma Weierstrass_l2_trig_polynomial: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a b where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" proof - let ?\ = "\n a b x. \k\n. a k * sin(real k * x) + b k * cos(real k * x)" obtain g where contg: "continuous_on {-pi..pi} g" and geq: "g(-pi) = g pi" and g: "g square_integrable {-pi..pi}" and norm_fg: "l2norm {-pi..pi} (\x. f x - g x) < e/2" using \e > 0\ by (auto intro: square_integrable_approximate_continuous_ends [OF f, of "e/2"]) then obtain n a b where g_phi_less: "\x. x \ {-pi..pi} \ \g x - (?\ n a b x)\ < e/6" using \e > 0\ Weierstrass_trig_polynomial [OF contg geq, of "e/6"] by (meson zero_less_divide_iff zero_less_numeral) show thesis proof have si: "(?\ n u v) square_integrable {-pi..pi}" for n::nat and u v proof (intro square_integrable_sum continuous_imp_square_integrable) show "continuous_on {-pi..pi} (\x. u k * sin (real k * x) + v k * cos (real k * x))" if "k \ {..n}" for k using that by (intro continuous_intros) qed auto have "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) = l2norm {-pi..pi} (\x. (f x - g x) + (g x - (?\ n a b x)))" by simp also have "\ \ l2norm {-pi..pi} (\x. f x - g x) + l2norm {-pi..pi} (\x. g x - (?\ n a b x))" proof (rule l2norm_triangle) show "(\x. f x - g x) square_integrable {-pi..pi}" using f g square_integrable_diff by blast show "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast qed also have "\ < e" proof - have g_phi: "(\x. g x - (?\ n a b x)) square_integrable {-pi..pi}" using g si square_integrable_diff by blast have "l2norm {-pi..pi} (\x. g x - (?\ n a b x)) \ e/2" unfolding l2norm_def l2product_def power2_eq_square [symmetric] proof (rule real_le_lsqrt) show "0 \ LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2" by (rule integral_nonneg_AE) auto have "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ LINT x|lebesgue_on {-pi..pi}. (e / 6) ^ 2" proof (rule integral_mono) show "integrable (lebesgue_on {-pi..pi}) (\x. (g x - (?\ n a b x))\<^sup>2)" using g_phi square_integrable_def by auto show "integrable (lebesgue_on {-pi..pi}) (\x. (e / 6)\<^sup>2)" by (intro continuous_intros continuous_imp_integrable_real) show "(g x - (?\ n a b x))\<^sup>2 \ (e / 6)\<^sup>2" if "x \ space (lebesgue_on {-pi..pi})" for x using \e > 0\ g_phi_less that apply (simp add: abs_le_square_iff [symmetric]) using less_eq_real_def by blast qed also have "\ \ (e / 2)\<^sup>2" using \e > 0\ pi_less_4 by (auto simp: power2_eq_square) finally show "LINT x|lebesgue_on {-pi..pi}. (g x - (?\ n a b x))\<^sup>2 \ (e / 2)\<^sup>2" . qed (use \e > 0\ in auto) with norm_fg show ?thesis by auto qed finally show "l2norm {-pi..pi} (\x. f x - (?\ n a b x)) < e" . qed qed proposition Weierstrass_l2_trigonometric_set: assumes f: "f square_integrable {-pi..pi}" and "0 < e" obtains n a where "l2norm {-pi..pi} (\x. f x - (\k\n. a k * trigonometric_set k x)) < e" proof - obtain n a b where lee: "l2norm {-pi..pi} (\x. f x - (\k\n. a k * sin(real k * x) + b k * cos(real k * x))) < e" using Weierstrass_l2_trig_polynomial [OF assms] . let ?a = "\k. if k = 0 then sqrt(2 * pi) * b 0 else if even k then sqrt pi * b(k div 2) else if k \ 2 * n then sqrt pi * a((Suc k) div 2) else 0" show thesis proof have [simp]: "Suc (i * 2) \ n * 2 \ i {..k\n. b k * cos (real k * x)) = (\i\n. if i = 0 then b 0 else b i * cos (real i * x))" for x by (rule sum.cong) auto moreover have "(\k\n. a k * sin (real k * x)) = (\i\n. (if Suc (2 * i) \ 2 * n then sqrt pi * a (Suc i) * sin ((1 + real i) * x) else 0) / sqrt pi)" (is "?lhs = ?rhs") for x proof (cases "n=0") case False then obtain n' where n': "n = Suc n'" using not0_implies_Suc by blast have "?lhs = (\k = Suc 0..n. a k * sin (real k * x))" by (simp add: atMost_atLeast0 sum_shift_lb_Suc0_0) also have "\ = (\ij\n'. a (Suc j) * sin (real (Suc j) * x)) = (\i = ?rhs" by (simp add: field_simps if_distrib [of "\x. x/_"] sum.inter_restrict [where B = "{..k\n. a k * sin(real k * x) + b k * cos(real k * x)) = (\k \ Suc(2*n). ?a k * trigonometric_set k x)" for x unfolding sum.in_pairs_0 trigonometric_set_def by (simp add: sum.distrib if_distrib [of "\x. x*_"] cong: if_cong) with lee show "l2norm {-pi..pi} (\x. f x - (\k \ Suc(2*n). ?a k * trigonometric_set k x)) < e" by auto qed qed subsection\Convergence wrt the L2 norm of trigonometric Fourier series\ definition Fourier_coefficient where "Fourier_coefficient \ orthonormal_coeff {-pi..pi} trigonometric_set" lemma Fourier_series_l2: assumes "f square_integrable {-pi..pi}" shows "(\n. l2norm {-pi..pi} (\x. f x - (\i\n. Fourier_coefficient f i * trigonometric_set i x))) \ 0" proof (clarsimp simp add: lim_sequentially dist_real_def Fourier_coefficient_def) let ?h = "\n x. (\i\n. orthonormal_coeff {-pi..pi} trigonometric_set f i * trigonometric_set i x)" show "\N. \n\N. \l2norm {-pi..pi} (\x. f x - ?h n x)\ < e" if "0 < e" for e proof - obtain N a where lte: "l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x)) < e" using Weierstrass_l2_trigonometric_set by (meson \0 < e\ assms) show ?thesis proof (intro exI allI impI) show "\l2norm {-pi..pi} (\x. f x - ?h m x)\ < e" if "N \ m" for m proof - have "\l2norm {-pi..pi} (\x. f x - ?h m x)\ = l2norm {-pi..pi} (\x. f x - ?h m x)" proof (rule abs_of_nonneg) show "0 \ l2norm {-pi..pi} (\x. f x - ?h m x)" apply (intro l2norm_pos_le square_integrable_diff square_integrable_sum square_integrable_lmult square_integrable_trigonometric_set assms, auto) done qed also have "\ \ l2norm {-pi..pi} (\x. f x - (\k\N. a k * trigonometric_set k x))" proof - have "(\i\m. (if i \ N then a i else 0) * trigonometric_set i x) = (\i\N. a i * trigonometric_set i x)" for x using sum.inter_restrict [where A = "{..m}" and B = "{..N}", symmetric] that by (force simp: if_distrib [of "\x. x * _"] min_absorb2 cong: if_cong) moreover have "l2norm {-pi..pi} (\x. f x - ?h m x) \ l2norm {-pi..pi} (\x. f x - (\i\m. (if i \ N then a i else 0) * trigonometric_set i x))" using orthonormal_optimal_partial_sum [OF orthonormal_system_trigonometric_set square_integrable_trigonometric_set assms] by simp ultimately show ?thesis by simp qed also have "\ < e" by (rule lte) finally show ?thesis . qed qed qed qed subsection\Fourier coefficients go to 0 (weak form of Riemann-Lebesgue)\ lemma trigonometric_set_mul_absolutely_integrable: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. trigonometric_set n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "trigonometric_set n \ borel_measurable (lebesgue_on {-pi..pi})" using square_integrable_def square_integrable_trigonometric_set by blast show "bounded (trigonometric_set n ` {-pi..pi})" unfolding bounded_iff using pi_gt3 sqrt_pi_ge1 by (rule_tac x=1 in exI) (auto simp: trigonometric_set_def dist_real_def intro: order_trans [OF abs_sin_le_one] order_trans [OF abs_cos_le_one]) qed (auto simp: assms) lemma trigonometric_set_mul_integrable: "f absolutely_integrable_on {-pi..pi} \ integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using trigonometric_set_mul_absolutely_integrable by (simp add: integrable_restrict_space set_integrable_def) lemma trigonometric_set_integrable [simp]: "integrable (lebesgue_on {-pi..pi}) (trigonometric_set n)" using trigonometric_set_mul_integrable [where f = id] by simp (metis absolutely_integrable_imp_integrable fmeasurableD interval_cbox lmeasurable_cbox square_integrable_imp_absolutely_integrable square_integrable_trigonometric_set) lemma absolutely_integrable_sin_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. sin(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_sin_cx mult_commute_abs) show "bounded ((\x. sin (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_sin_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma absolutely_integrable_cos_product: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. cos(k * x) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (k * x)) \ borel_measurable (lebesgue_on {-pi..pi})" by (metis borel_measurable_integrable integrable_cos_cx mult_commute_abs) show "bounded ((\x. cos (k * x)) ` {-pi..pi})" by (metis (mono_tags, lifting) abs_cos_le_one bounded_iff imageE real_norm_def) qed (auto simp: assms) lemma assumes "f absolutely_integrable_on {-pi..pi}" shows Fourier_products_integrable_cos: "integrable (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x)" and Fourier_products_integrable_sin: "integrable (lebesgue_on {-pi..pi}) (\x. sin(k * x) * f x)" using absolutely_integrable_cos_product absolutely_integrable_sin_product assms by (auto simp: integrable_restrict_space set_integrable_def) lemma Riemann_lebesgue_square_integrable: assumes "orthonormal_system S w" "\i. w i square_integrable S" "f square_integrable S" shows "orthonormal_coeff S w f \ 0" using Fourier_series_square_summable [OF assms, of UNIV] summable_LIMSEQ_zero by force proposition Riemann_lebesgue: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient f \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain g where "continuous_on UNIV g" and gabs: "g absolutely_integrable_on {-pi..pi}" and fg_e2: "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. \f x - g x\) < e/2" using absolutely_integrable_approximate_continuous [OF assms, of "e/2"] by (metis (full_types) box_real(2) half_gt_zero_iff lmeasurable_cbox) have "g square_integrable {-pi..pi}" using \continuous_on UNIV g\ continuous_imp_square_integrable continuous_on_subset by blast then have "orthonormal_coeff {-pi..pi} trigonometric_set g \ 0" using Riemann_lebesgue_square_integrable orthonormal_system_trigonometric_set square_integrable_trigonometric_set by blast with \e > 0\ obtain N where N: "\n. n \ N \ \orthonormal_coeff {-pi..pi} trigonometric_set g n\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def) have "\Fourier_coefficient f n\ < e" if "n \ N" for n proof - have "\LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x\ < e/2" using N [OF \n \ N\] by (auto simp: orthonormal_coeff_def l2product_def) have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * g x)" using gabs trigonometric_set_mul_integrable by blast moreover have "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * f x)" using assms trigonometric_set_mul_integrable by blast ultimately have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ = \(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * (g x - f x))\" by (simp add: algebra_simps flip: Bochner_Integration.integral_diff) also have "\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. trigonometric_set n x * (g x - f x))" by (simp add: gabs assms trigonometric_set_mul_integrable) have "(\x. f x - g x) absolutely_integrable_on {-pi..pi}" using gabs assms by blast then show "integrable (lebesgue_on {-pi..pi}) (\x. \f x - g x\)" by (simp add: absolutely_integrable_imp_integrable) fix x assume "x \ space (lebesgue_on {-pi..pi})" then have "-pi \ x" "x \ pi" by auto have "\trigonometric_set n x\ \ 1" using pi_ge_two apply (simp add: trigonometric_set_def) using sqrt_pi_ge1 abs_sin_le_one order_trans abs_cos_le_one by metis then show "\trigonometric_set n x * (g x - f x)\ \ \f x - g x\" using abs_ge_zero mult_right_mono by (fastforce simp add: abs_mult abs_minus_commute) qed finally have "\(LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * g x) - (LINT x|lebesgue_on {-pi..pi}. trigonometric_set n x * f x)\ \ LINT x|lebesgue_on {-pi..pi}. \f x - g x\" . then show ?thesis using N [OF \n \ N\] fg_e2 unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by linarith qed then show "\no. \n\no. dist (Fourier_coefficient f n) 0 < e" by auto qed lemma Riemann_lebesgue_sin: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have "\Fourier_coefficient f(Suc (2 * n))\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. sin ((1 + real n) * x) * f x\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by simp qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. sin (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_cos: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(real n * x) * f x)) \ 0" unfolding lim_sequentially proof (intro allI impI) fix e::real assume "e > 0" then obtain N where N: "\n. n \ N \ \Fourier_coefficient f n\ < e/4" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) have "\LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x\ < e" if "n > N" for n using that proof (induction n) case (Suc n) have eq: "(x * 2 + x * (real n * 2)) / 2 = x + x * (real n)" for x by simp have "\Fourier_coefficient f(2*n + 2)\ < e/4" using N Suc.prems by auto then have "\LINT x|lebesgue_on {-pi..pi}. f x * cos (x + x * (real n))\ < sqrt pi * e / 4" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def field_simps eq) also have "\ \ e" using \0 < e\ pi_less_4 real_sqrt_less_mono by (fastforce simp add: field_simps) finally show ?case by (simp add: field_simps) qed auto then show "\no. \n\no. dist (LINT x|lebesgue_on {-pi..pi}. cos (real n * x) * f x) 0 < e" by (metis (full_types) le_neq_implies_less less_add_same_cancel2 less_trans norm_conv_dist real_norm_def zero_less_one) qed lemma Riemann_lebesgue_sin_half: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. LINT x|lebesgue_on {-pi..pi}. sin ((real n + 1/2) * x) * f x) \ 0" proof (simp add: algebra_simps sin_add) let ?INT = "integral\<^sup>L (lebesgue_on {-pi..pi})" let ?f = "(\n. ?INT (\x. sin(n * x) * cos(1/2 * x) * f x) + ?INT (\x. cos(n * x) * sin(1/2 * x) * f x))" show "(\n. ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))) \ 0" proof (rule Lim_transform_eventually) have sin: "(\x. sin (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_sin_product assms) have cos: "(\x. cos (1/2 * x) * f x) absolutely_integrable_on {-pi..pi}" by (intro absolutely_integrable_cos_product assms) show "\\<^sub>F n in sequentially. ?f n = ?INT (\x. f x * (cos (x * real n) * sin (x/2)) + f x * (sin (x * real n) * cos (x/2)))" unfolding mult.assoc apply (rule eventuallyI) apply (subst Bochner_Integration.integral_add [symmetric]) apply (safe intro!: cos absolutely_integrable_sin_product sin absolutely_integrable_cos_product absolutely_integrable_imp_integrable) apply (auto simp: algebra_simps) done have "?f \ 0 + 0" unfolding mult.assoc by (intro tendsto_add Riemann_lebesgue_sin Riemann_lebesgue_cos sin cos) then show "?f \ (0::real)" by simp qed qed lemma Fourier_sum_limit_pair: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\n. \k\2 * n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e::real assume "e > 0" then obtain N1 where N1: "\n. n \ N1 \ \Fourier_coefficient f n\ < e/2" using Riemann_lebesgue [OF assms] unfolding lim_sequentially by (metis norm_conv_dist real_norm_def zero_less_divide_iff zero_less_numeral) obtain N2 where N2: "\n. n \ N2 \ \(\k\2 * n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using L unfolding lim_sequentially dist_real_def by (meson \0 < e\ half_gt_zero_iff) show "\no. \n\no. \(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof (intro exI allI impI) fix n assume n: "N1 + 2*N2 + 1 \ n" then have "n \ N1" "n \ N2" "n div 2 \ N2" by linarith+ consider "n = 2 * (n div 2)" | "n = Suc(2 * (n div 2))" by linarith then show "\(\k\n. Fourier_coefficient f k * trigonometric_set k t) - l\ < e" proof cases case 1 show ?thesis apply (subst 1) using N2 [OF \n div 2 \ N2\] by linarith next case 2 have "\(\k\2 * (n div 2). Fourier_coefficient f k * trigonometric_set k t) - l\ < e/2" using N2 [OF \n div 2 \ N2\] by linarith moreover have "\Fourier_coefficient f(Suc (2 * (n div 2))) * trigonometric_set (Suc (2 * (n div 2))) t\ < (e/2) * 1" proof - have "\trigonometric_set (Suc (2 * (n div 2))) t\ \ 1" apply (simp add: trigonometric_set) using abs_sin_le_one sqrt_pi_ge1 by (blast intro: order_trans) moreover have "\Fourier_coefficient f(Suc (2 * (n div 2)))\ < e/2" using N1 \N1 \ n\ by auto ultimately show ?thesis unfolding abs_mult by (meson abs_ge_zero le_less_trans mult_left_mono real_mult_less_iff1 zero_less_one) qed ultimately show ?thesis apply (subst 2) unfolding sum.atMost_Suc by linarith qed qed qed next assume ?rhs then show ?lhs unfolding lim_sequentially by (auto simp: elim!: imp_forward ex_forward) qed subsection\Express Fourier sum in terms of the special expansion at the origin\ lemma Fourier_sum_0: "(\k \ n. Fourier_coefficient f k * trigonometric_set k 0) = (\k \ n div 2. Fourier_coefficient f(2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k = 2 * 0.. Suc (2 * (n div 2)). Fourier_coefficient f k * trigonometric_set k 0)" proof (rule sum.mono_neutral_left) show "\i\{2 * 0..Suc (2 * (n div 2))} - {..n}. Fourier_coefficient f i * trigonometric_set i 0 = 0" proof clarsimp fix i assume "i \ Suc (2 * (n div 2))" "\ i \ n" then have "i = Suc n" "even n" by presburger+ moreover assume "trigonometric_set i 0 \ 0" ultimately show "Fourier_coefficient f i = 0" by (simp add: trigonometric_set_def) qed qed auto also have "\ = ?rhs" unfolding sum.in_pairs by (simp add: trigonometric_set atMost_atLeast0) finally show ?thesis . qed lemma Fourier_sum_0_explicit: "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (Fourier_coefficient f 0 / sqrt 2 + (\k = 1..n div 2. Fourier_coefficient f(2*k))) / sqrt pi" (is "?lhs = ?rhs") proof - have "(\k=0..n. Fourier_coefficient f k * trigonometric_set k 0) = Fourier_coefficient f 0 * trigonometric_set 0 0 + (\k = 1..n. Fourier_coefficient f k * trigonometric_set k 0)" by (simp add: Fourier_sum_0 sum.atLeast_Suc_atMost) also have "\ = ?rhs" proof - have "Fourier_coefficient f 0 * trigonometric_set 0 0 = Fourier_coefficient f 0 / (sqrt 2 * sqrt pi)" by (simp add: real_sqrt_mult trigonometric_set(1)) moreover have "(\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k 0) = (\k = Suc 0..n div 2. Fourier_coefficient f(2*k)) / sqrt pi" proof (induction n) case (Suc n) show ?case by (simp add: Suc) (auto simp: trigonometric_set_def field_simps) qed auto ultimately show ?thesis by (simp add: add_divide_distrib) qed finally show ?thesis by (simp add: atMost_atLeast0) qed lemma Fourier_sum_0_integrals: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = (integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(k * x) * f x))) / pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / (sqrt (2 * pi) * sqrt 2 * sqrt pi) = integral\<^sup>L (lebesgue_on {-pi..pi}) f / (2 * pi)" by (simp add: algebra_simps real_sqrt_mult) moreover have "(\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. trigonometric_set (2*k) x * f x) / sqrt pi = (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. cos (k * x) * f x) / pi" by (simp add: trigonometric_set_def sum_divide_distrib) ultimately show ?thesis unfolding Fourier_sum_0_explicit by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set add_divide_distrib) qed lemma Fourier_sum_0_integral: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\k\n. Fourier_coefficient f k * trigonometric_set k 0) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. (1/2 + (\k = Suc 0..n div 2. cos(k * x))) * f x) / pi" proof - note * = Fourier_products_integrable_cos [OF assms] have "integrable (lebesgue_on {-pi..pi}) (\x. \n = Suc 0..n div 2. f x * cos (x * real n))" using * by (force simp: mult_ac) moreover have "integral\<^sup>L (lebesgue_on {-pi..pi}) f / 2 + (\k = Suc 0..n div 2. LINT x|lebesgue_on {-pi..pi}. f x * cos (x * real k)) = (LINT x|lebesgue_on {-pi..pi}. f x / 2) + (LINT x|lebesgue_on {-pi..pi}. (\n = Suc 0..n div 2. f x * cos (x * real n)))" proof (subst Bochner_Integration.integral_sum) show "integrable (lebesgue_on {-pi..pi}) (\x. f x * cos (x * real i))" if "i \ {Suc 0..n div 2}" for i using that * [of i] by (auto simp: Bochner_Integration.integral_sum mult_ac) qed auto ultimately show ?thesis using Fourier_products_integrable_cos [OF assms] absolutely_integrable_imp_integrable [OF assms] by (simp add: Fourier_sum_0_integrals sum_distrib_left assms algebra_simps) qed subsection\How Fourier coefficients behave under addition etc\ lemma Fourier_coefficient_add: assumes "f absolutely_integrable_on {-pi..pi}" "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x + g x) i = Fourier_coefficient f i + Fourier_coefficient g i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def distrib_left) lemma Fourier_coefficient_minus: assumes "f absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. - f x) i = - Fourier_coefficient f i" using assms trigonometric_set_mul_integrable by (simp add: Fourier_coefficient_def orthonormal_coeff_def l2product_def) lemma Fourier_coefficient_diff: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" shows "Fourier_coefficient (\x. f x - g x) i = Fourier_coefficient f i - Fourier_coefficient g i" proof - have mg: "(\x. - g x) absolutely_integrable_on {-pi..pi}" using g by (simp add: absolutely_integrable_measurable_real) show ?thesis using Fourier_coefficient_add [OF f mg] Fourier_coefficient_minus [OF g] by simp qed lemma Fourier_coefficient_const: "Fourier_coefficient (\x. c) i = (if i = 0 then c * sqrt(2 * pi) else 0)" by (auto simp: Fourier_coefficient_def orthonormal_coeff_def l2product_def trigonometric_set_def divide_simps) lemma Fourier_offset_term: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "Fourier_coefficient (\x. f(x+t)) (2 * n + 2) * trigonometric_set (2 * n + 2) 0 = Fourier_coefficient f(2 * n+1) * trigonometric_set (2 * n+1) t + Fourier_coefficient f(2 * n + 2) * trigonometric_set (2 * n + 2) t" proof - have eq: "(2 + 2 * real n) * x / 2 = (1 + real n) * x" for x by (simp add: divide_simps) have 1: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (cos (x + x * n) * cos (t + t * n)))" using Fourier_products_integrable_cos [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have 2: "integrable (lebesgue_on {-pi..pi}) (\x. f x * (sin (x + x * n) * sin (t + t * n)))" using Fourier_products_integrable_sin [of f "Suc n"] f by (simp add: algebra_simps) (simp add: mult.assoc [symmetric]) have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * (x + t - t)) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" proof (rule has_integral_periodic_offset) have "(\x. cos (real (Suc n) * (x - t)) * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. cos (real (Suc n) * (x - t))) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. cos (real (Suc n) * (x - t))) ` {-pi..pi})" by (rule boundedI [where B=1]) auto qed (use f in auto) then show "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos ((Suc n) * (x - t)) * f x) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by (simp add: has_bochner_integral_integrable integrable_restrict_space set_integrable_def) next fix x show "cos (real (Suc n) * (x + (pi - - pi) - t)) * f(x + (pi - - pi)) = cos (real (Suc n) * (x - t)) * f x" using periodic cos.plus_of_nat [of "(Suc n) * (x - t)" "Suc n"] by (simp add: algebra_simps) qed then have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. cos (real (Suc n) * x) * f(x + t)) (LINT x|lebesgue_on {-pi..pi}. cos (real (Suc n) * (x - t)) * f x)" by simp then have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * (x - t)) * f x" using has_bochner_integral_integral_eq by blast also have "\ = LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x - ((Suc n) * t)) * f x" by (simp add: algebra_simps) also have "\ = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" by (simp add: cos_diff algebra_simps 1 2 flip: integral_mult_left_zero integral_mult_right_zero) finally have "LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f(x+t) = cos ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. cos ((Suc n) * x) * f x) + sin ((Suc n) * t) * (LINT x|lebesgue_on {-pi..pi}. sin ((Suc n) * x) * f x)" . then show ?thesis unfolding Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def by (simp add: eq) (simp add: divide_simps algebra_simps l2product_def) qed lemma Fourier_sum_offset: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" (is "?lhs = ?rhs") proof - have "?lhs = Fourier_coefficient f 0 * trigonometric_set 0 t + (\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "(\k = Suc 0..2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" proof (cases n) case (Suc n') have "(\k = Suc 0..2 * n. Fourier_coefficient f k * trigonometric_set k t) = (\k = Suc 0.. Suc (Suc (2 * n')). Fourier_coefficient f k * trigonometric_set k t)" by (simp add: Suc) also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient f(Suc k) * trigonometric_set (Suc k) t)" unfolding atMost_atLeast0 using sum.shift_bounds_cl_Suc_ivl by blast also have "\ = (\k \ Suc (2 * n'). Fourier_coefficient (\x. f(x+t)) (Suc k) * trigonometric_set (Suc k) 0)" unfolding sum.in_pairs_0 proof (rule sum.cong [OF refl]) show "Fourier_coefficient f(Suc (2 * k)) * trigonometric_set (Suc (2 * k)) t + Fourier_coefficient f(Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) t = Fourier_coefficient (\x. f(x + t)) (Suc (2 * k)) * trigonometric_set (Suc (2 * k)) 0 + Fourier_coefficient (\x. f(x + t)) (Suc (Suc (2 * k))) * trigonometric_set (Suc (Suc (2 * k))) 0" if "k \ {..n'}" for k using that Fourier_offset_term [OF assms, of t ] by (auto simp: trigonometric_set_def) qed also have "\ = (\k = Suc 0..Suc (Suc (2 * n')). Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp only: atMost_atLeast0 sum.shift_bounds_cl_Suc_ivl) also have "\ = (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: Suc) finally show ?thesis . qed auto moreover have "?rhs = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0 + (\k = Suc 0..2*n. Fourier_coefficient (\x. f(x+t)) k * trigonometric_set k 0)" by (simp add: atMost_atLeast0 sum.atLeast_Suc_atMost) moreover have "Fourier_coefficient f 0 * trigonometric_set 0 t = Fourier_coefficient (\x. f(x+t)) 0 * trigonometric_set 0 0" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def integral_periodic_offset periodic) ultimately show ?thesis by simp qed lemma Fourier_sum_offset_unpaired: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" (is "?lhs = ?rhs") proof - have "?lhs = (\k\n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0 + Fourier_coefficient (\x. f(x+t)) (Suc (2*k)) * trigonometric_set (Suc (2*k)) 0)" apply (simp add: assms Fourier_sum_offset) apply (subst sum.in_pairs_0 [symmetric]) by (simp add: trigonometric_set) also have "\ = ?rhs" by (auto simp: trigonometric_set) finally show ?thesis . qed subsection\Express partial sums using Dirichlet kernel\ definition Dirichlet_kernel where "Dirichlet_kernel \ \n x. if x = 0 then real n + 1/2 else sin((real n + 1/2) * x) / (2 * sin(x/2))" lemma Dirichlet_kernel_0 [simp]: "\x\ < 2 * pi \ Dirichlet_kernel 0 x = 1/2" by (auto simp: Dirichlet_kernel_def sin_zero_iff) lemma Dirichlet_kernel_minus [simp]: "Dirichlet_kernel n (-x) = Dirichlet_kernel n x" by (auto simp: Dirichlet_kernel_def) lemma Dirichlet_kernel_continuous_strong: "continuous_on {-(2 * pi)<..<2 * pi} (Dirichlet_kernel n)" proof - have "isCont (Dirichlet_kernel n) z" if "-(2 * pi) < z" "z < 2 * pi" for z proof (cases "z=0") case True have *: "(\x. sin ((n + 1/2) * x) / (2 * sin (x/2))) \0 \ real n + 1/2" by real_asymp show ?thesis unfolding Dirichlet_kernel_def continuous_at True apply (rule Lim_transform_eventually [where f = "\x. sin((n + 1/2) * x) / (2 * sin(x/2))"]) apply (auto simp: * eventually_at_filter) done next case False have "continuous (at z) (\x. sin((real n + 1/2) * x) / (2 * sin(x/2)))" using that False by (intro continuous_intros) (auto simp: sin_zero_iff) moreover have "\\<^sub>F x in nhds z. x \ 0" using False t1_space_nhds by blast ultimately show ?thesis using False by (force simp: Dirichlet_kernel_def continuous_at eventually_at_filter elim: Lim_transform_eventually) qed then show ?thesis by (simp add: continuous_on_eq_continuous_at) qed lemma Dirichlet_kernel_continuous: "continuous_on {-pi..pi} (Dirichlet_kernel n)" apply (rule continuous_on_subset [OF Dirichlet_kernel_continuous_strong], clarsimp) using pi_gt_zero by linarith lemma absolutely_integrable_mult_Dirichlet_kernel: assumes "f absolutely_integrable_on {-pi..pi}" shows "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Dirichlet_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Dirichlet_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) have "compact (Dirichlet_kernel n ` {-pi..pi})" by (auto simp: compact_continuous_image [OF Dirichlet_kernel_continuous]) then show "bounded (Dirichlet_kernel n ` {-pi..pi})" using compact_imp_bounded by blast qed (use assms in auto) lemma cosine_sum_lemma: "(1/2 + (\k = Suc 0..n. cos(real k * x))) * sin(x/2) = sin((real n + 1/2) * x) / 2" proof - consider "n=0" | "n \ 1" by linarith then show ?thesis proof cases case 2 then have "(\k = Suc 0..n. (sin (real k * x + x/2) * 2 - sin (real k * x - x/2) * 2) / 2) = sin (real n * x + x/2) - sin (x/2)" proof (induction n) case (Suc n) show ?case proof (cases "n=0") case False with Suc show ?thesis by (simp add: divide_simps algebra_simps) qed auto qed auto then show ?thesis by (simp add: distrib_right sum_distrib_right cos_times_sin) qed auto qed lemma Dirichlet_kernel_cosine_sum: assumes "\x\ < 2 * pi" shows "Dirichlet_kernel n x = 1/2 + (\k = Suc 0..n. cos(real k * x))" proof - have "sin ((real n + 1/2) * x) / (2 * sin (x/2)) = 1/2 + (\k = Suc 0..n. cos (real k * x))" if "x \ 0" proof - have "sin(x/2) \ 0" using assms that by (auto simp: sin_zero_iff) then show ?thesis using cosine_sum_lemma [of x n] by (simp add: divide_simps) qed then show ?thesis by (auto simp: Dirichlet_kernel_def) qed lemma integrable_Dirichlet_kernel: "integrable (lebesgue_on {-pi..pi}) (Dirichlet_kernel n)" using Dirichlet_kernel_continuous continuous_imp_integrable_real by blast lemma integral_Dirichlet_kernel [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = pi" proof - have "integral\<^sup>L (lebesgue_on {-pi..pi}) (Dirichlet_kernel n) = LINT x|lebesgue_on {-pi..pi}. 1/2 + (\k = Suc 0..n. cos (k * x))" using pi_ge_two by (force intro: Bochner_Integration.integral_cong [OF refl Dirichlet_kernel_cosine_sum]) also have "\ = (LINT x|lebesgue_on {-pi..pi}. 1/2) + (LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (k * x)))" proof (rule Bochner_Integration.integral_add) show "integrable (lebesgue_on {-pi..pi}) (\x. \k = Suc 0..n. cos (real k * x))" by (rule Bochner_Integration.integrable_sum) (metis integrable_cos_cx mult_commute_abs) qed auto also have "\ = pi" proof - have "\i. \Suc 0 \ i; i \ n\ \ integrable (lebesgue_on {-pi..pi}) (\x. cos (real i * x))" by (metis integrable_cos_cx mult_commute_abs) then have "LINT x|lebesgue_on {-pi..pi}. (\k = Suc 0..n. cos (real k * x)) = 0" by (simp add: Bochner_Integration.integral_sum) then show ?thesis by auto qed finally show ?thesis . qed lemma integral_Dirichlet_kernel_half [simp]: "integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi/2" proof - have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) + integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n) = pi" using integral_combine [OF integrable_Dirichlet_kernel] integral_Dirichlet_kernel by simp moreover have "integral\<^sup>L (lebesgue_on {- pi..0}) (Dirichlet_kernel n) = integral\<^sup>L (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using integral_reflect_real [of pi 0 "Dirichlet_kernel n"] by simp ultimately show ?thesis by simp qed lemma Fourier_sum_offset_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f(x+t)) / pi" (is "?lhs = ?rhs") proof - have ft: "(\x. f(x+t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp have "?lhs = (\k=0..n. Fourier_coefficient (\x. f(x+t)) (2*k) * trigonometric_set (2*k) 0)" using Fourier_sum_offset_unpaired assms atMost_atLeast0 by auto also have "\ = Fourier_coefficient (\x. f(x+t)) 0 / sqrt (2 * pi) + (\k = Suc 0..n. Fourier_coefficient (\x. f(x+t)) (2*k) / sqrt pi)" by (simp add: sum.atLeast_Suc_atMost trigonometric_set_def) also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t)) / (2 * pi) + (\k = Suc 0..n. (LINT x|lebesgue_on {-pi..pi}. cos (real k * x) * f(x+t)) / pi)" by (simp add: Fourier_coefficient_def orthonormal_coeff_def trigonometric_set_def l2product_def) also have "\ = LINT x|lebesgue_on {-pi..pi}. f(x+t) / (2 * pi) + (\k = Suc 0..n. (cos (real k * x) * f(x+t)) / pi)" using Fourier_products_integrable_cos [OF ft] absolutely_integrable_imp_integrable [OF ft] by simp also have "\ = (LINT x|lebesgue_on {-pi..pi}. f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x+t))) / pi" by (simp add: divide_simps sum_distrib_right mult.assoc) also have "\ = ?rhs" proof - have "LINT x|lebesgue_on {-pi..pi}. f(x + t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)" proof - have eq: "f(x+t) / 2 + (\k = Suc 0..n. cos (real k * x) * f(x + t)) = Dirichlet_kernel n x * f(x + t)" if "- pi \ x" "x \ pi" for x proof (cases "x = 0") case False then have "sin (x/2) \ 0" using that by (auto simp: sin_zero_iff) then have "f(x + t) * (1/2 + (\k = Suc 0..n. cos(real k * x))) = f(x + t) * sin((real n + 1/2) * x) / 2 / sin(x/2)" using cosine_sum_lemma [of x n] by (simp add: divide_simps) then show ?thesis by (simp add: Dirichlet_kernel_def False field_simps sum_distrib_left) qed (simp add: Dirichlet_kernel_def algebra_simps) show ?thesis by (rule Bochner_Integration.integral_cong [OF refl]) (simp add: eq) qed then show ?thesis by simp qed finally show ?thesis . qed lemma Fourier_sum_limit_Dirichlet_kernel: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "((\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l) \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * l" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) / pi) \ l" by (simp add: Fourier_sum_limit_pair [OF f, symmetric] Fourier_sum_offset_Dirichlet_kernel assms) also have "\ \ (\k. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel k x * f(x + t)) * inverse pi) \ pi * l * inverse pi" by (auto simp: divide_simps) also have "\ \ ?rhs" using real_tendsto_mult_right_iff [of "inverse pi", symmetric] by auto finally show ?thesis . qed subsection\A directly deduced sufficient condition for convergence at a point\ lemma simple_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and ft: "(\x. (f(x+t) - f t) / sin(x/2)) absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - let ?\ = "\n. \k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t" let ?\ = "\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (f(x + t) - f t)" have fft: "(\x. f x - f t) absolutely_integrable_on {-pi..pi}" by (simp add: f absolutely_integrable_measurable_real) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have *: "?\ \ 0 \ ?\ \ 0" by (simp add: Fourier_sum_limit_Dirichlet_kernel fft periodic) moreover have "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t) \ 0" if "?\ \ 0" proof (rule Lim_transform_eventually [OF that eventually_sequentiallyI]) show "(\k\n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k\n. Fourier_coefficient f k * trigonometric_set k t) - f t" if "Suc 0 \ n" for n proof - have "(\k = Suc 0..n. Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t) = (\k = Suc 0..n. Fourier_coefficient f k * trigonometric_set k t)" proof (rule sum.cong [OF refl]) fix k assume k: "k \ {Suc 0..n}" then have [simp]: "integral\<^sup>L (lebesgue_on {-pi..pi}) (trigonometric_set k) = 0" by (auto simp: trigonometric_set_def) show "Fourier_coefficient (\x. f x - f t) k * trigonometric_set k t = Fourier_coefficient f k * trigonometric_set k t" using k unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (simp add: right_diff_distrib f absolutely_integrable_measurable_real trigonometric_set_mul_integrable) qed moreover have "Fourier_coefficient (\x. f x - f t) 0 * trigonometric_set 0 t = Fourier_coefficient f 0 * trigonometric_set 0 t - f t" using f unfolding Fourier_coefficient_def orthonormal_coeff_def l2product_def by (auto simp: divide_simps trigonometric_set absolutely_integrable_imp_integrable) ultimately show ?thesis by (simp add: sum.atLeast_Suc_atMost atMost_atLeast0) qed qed moreover have "?\ \ 0" proof - have eq: "\n. ?\ n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n + 1/2) * x) * ((f(x+t) - f t) / sin(x/2) / 2))" unfolding Dirichlet_kernel_def by (rule Bochner_Integration.integral_cong [OF refl]) (auto simp: divide_simps sin_zero_iff) show ?thesis unfolding eq by (intro ft set_integrable_divide Riemann_lebesgue_sin_half) qed ultimately show ?thesis by (simp add: LIM_zero_cancel) qed subsection\A more natural sufficient Hoelder condition at a point\ lemma bounded_inverse_sin_half: assumes "d > 0" obtains B where "B>0" "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" proof - have "bounded ((\x. inverse (sin (x/2))) ` ({-pi..pi} - {- d<..x \ {-pi..pi}; x \ 0\ \ sin (x/2) \ 0" for x using \0 < d\ by (auto simp: sin_zero_iff) then show "continuous_on ({-pi..pi} - {-d<..x. inverse (sin (x/2)))" using \0 < d\ by (intro continuous_intros) auto show "compact ({-pi..pi} - {-d<.. 0" "a > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\ powr a" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof (rule simple_Fourier_convergence_periodic [OF f]) have half: "((\x. sin(x/2)) has_real_derivative 1/2) (at 0)" by (rule derivative_eq_intros refl | force)+ then obtain e0::real where "e0 > 0" and e0: "\x. \x \ 0; \x\ < e0\ \ \sin (x/2) / x - 1/2\ < 1/4" apply (simp add: DERIV_def Lim_at dist_real_def) apply (drule_tac x="1/4" in spec, auto) done obtain e where "e > 0" and e: "\x. \x\ < e \ \(f(x+t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a-1))" proof show "min d e0 > 0" using \d > 0\ \e0 > 0\ by auto next fix x assume x: "\x\ < min d e0" show "\(f(x + t) - f t) / sin (x/2)\ \ 4 * (\M\ * \x\ powr (a - 1))" proof (cases "sin(x/2) = 0 \ x=0") case False have eq: "\(f(x + t) - f t) / sin (x/2)\ = \inverse (sin (x/2) / x)\ * (\f(x + t) - f t\ / \x\)" by simp show ?thesis unfolding eq proof (rule mult_mono) have "\sin (x/2) / x - 1/2\ < 1/4" using e0 [of x] x False by force then have "1/4 \ \sin (x/2) / x\" by (simp add: abs_if field_simps split: if_split_asm) then show "\inverse (sin (x/2) / x)\ \ 4" using False by (simp add: field_simps) have "\f(x + t) - f t\ / \x\ \ M * \x + t - t\ powr (a - 1)" using ft [of "x+t"] x by (auto simp: divide_simps algebra_simps Transcendental.powr_mult_base) also have "\ \ \M\ * \x\ powr (a - 1)" by (simp add: mult_mono) finally show "\f(x + t) - f t\ / \x\ \ \M\ * \x\ powr (a - 1)" . qed auto qed auto qed obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {- e<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \e > 0\] by metis let ?g = "\x. max (B * \f(x + t) - f t\) ((4 * \M\) * \x\ powr (a - 1))" show "(\x. (f(x + t) - f t) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto show "(\x. (f(x + t) - f t) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "(\x. f(x + t)) \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def fxt integrable_imp_measurable by blast show "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto qed auto have "(\x. f(x + t) - f t) absolutely_integrable_on {-pi..pi}" by (intro set_integral_diff fxt) (simp add: set_integrable_def) moreover have "(\x. \x\ powr (a - 1)) absolutely_integrable_on {-pi..pi}" proof - have "((\x. x powr (a - 1)) has_integral inverse a * pi powr a - inverse a * 0 powr a) {0..pi}" proof (rule fundamental_theorem_of_calculus_interior) show "continuous_on {0..pi} (\x. inverse a * x powr a)" using \a > 0\ by (intro continuous_on_powr' continuous_intros) auto show "((\b. inverse a * b powr a) has_vector_derivative x powr (a - 1)) (at x)" if "x \ {0<..a > 0\ apply (simp flip: has_field_derivative_iff_has_vector_derivative) apply (rule derivative_eq_intros | simp)+ done qed auto then have int: "(\x. x powr (a - 1)) integrable_on {0..pi}" by blast show ?thesis proof (rule nonnegative_absolutely_integrable_1) show "(\x. \x\ powr (a - 1)) integrable_on {-pi..pi}" proof (rule Henstock_Kurzweil_Integration.integrable_combine) show "(\x. \x\ powr (a - 1)) integrable_on {0..pi}" using int integrable_eq by fastforce then show "(\x. \x\ powr (a - 1)) integrable_on {- pi..0}" using Henstock_Kurzweil_Integration.integrable_reflect_real by fastforce qed auto qed auto qed ultimately show "?g integrable_on {-pi..pi}" apply (intro set_lebesgue_integral_eq_integral absolutely_integrable_max_1 absolutely_integrable_bounded_measurable_product_real set_integrable_abs measurable) apply (auto simp: image_constant_conv) done show "norm ((f(x + t) - f t) / sin (x/2)) \ ?g x" if "x \ {-pi..pi}" for x proof (cases "\x\ < e") case True then show ?thesis using e [OF True] by (simp add: max_def) next case False then have "\inverse (sin (x/2))\ \ B" using B that by force then have "\inverse (sin (x/2))\ * \f(x + t) - f t\ \ B * \f(x + t) - f t\" by (simp add: mult_right_mono) then have "\f(x + t) - f t\ / \sin (x/2)\ \ B * \f(x + t) - f t\" by (simp add: divide_simps split: if_split_asm) then show ?thesis by auto qed qed auto qed (auto simp: periodic) text\In particular, a Lipschitz condition at the point\ corollary Lipschitz_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and "d > 0" and ft: "\x. \x-t\ < d \ \f x - f t\ \ M * \x-t\" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" using Hoelder_Fourier_convergence_periodic [OF f \d > 0\, of 1] assms by (fastforce simp add:) text\In particular, if left and right derivatives both exist\ proposition bi_differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and f_lt: "f differentiable at_left t" and f_gt: "f differentiable at_right t" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" proof - obtain D_lt where D_lt: "\e. e > 0 \ \d>0. \x dist x t < d \ dist ((f x - f t) / (x - t)) D_lt < e" using f_lt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson lessThan_iff) then obtain d_lt where "d_lt > 0" and d_lt: "\x. \x < t; 0 < \x - t\; \x - t\ < d_lt\ \ \(f x - f t) / (x - t) - D_lt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) obtain D_gt where D_gt: "\e. e > 0 \ \d>0. \x>t. 0 < dist x t \ dist x t < d \ dist ((f x - f t) / (x - t)) D_gt < e" using f_gt unfolding has_field_derivative_iff real_differentiable_def Lim_within by (meson greaterThan_iff) then obtain d_gt where "d_gt > 0" and d_gt: "\x. \x > t; 0 < \x - t\; \x - t\ < d_gt\ \ \(f x - f t) / (x - t) - D_gt\ < 1" unfolding dist_real_def by (metis dist_commute dist_real_def zero_less_one) show ?thesis proof (rule Lipschitz_Fourier_convergence_periodic [OF f]) fix x assume "\x - t\ < min d_lt d_gt" then have *: "\x - t\ < d_lt" "\x - t\ < d_gt" by auto consider "xt" by linarith then show "\f x - f t\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" proof cases case 1 then have "\(f x - f t) / (x - t) - D_lt\ < 1" using d_lt [OF 1] * by auto then have "\(f x - f t) / (x - t)\ - \D_lt\ < 1" by linarith then have "\f x - f t\ \ (\D_lt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . next case 3 then have "\(f x - f t) / (x - t) - D_gt\ < 1" using d_gt [OF 3] * by auto then have "\(f x - f t) / (x - t)\ - \D_gt\ < 1" by linarith then have "\f x - f t\ \ (\D_gt\ + 1) * \x - t\" by (simp add: comm_semiring_class.distrib divide_simps split: if_split_asm) also have "\ \ (\D_lt\ + \D_gt\ + 1) * \x - t\" by (simp add: mult_right_mono) finally show ?thesis . qed auto qed (auto simp: \0 < d_gt\ \0 < d_lt\ periodic) qed text\And in particular at points where the function is differentiable\ lemma differentiable_Fourier_convergence_periodic: assumes f: "f absolutely_integrable_on {-pi..pi}" and fdif: "f differentiable (at t)" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ f t" by (rule bi_differentiable_Fourier_convergence_periodic) (auto simp: differentiable_at_withinI assms) text\Use reflection to halve the region of integration\ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" proof - show "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {-pi..pi}" apply (rule absolutely_integrable_mult_Dirichlet_kernel) using absolutely_integrable_periodic_offset [OF f, of t] periodic apply simp done then show "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {-pi..pi}" by (subst absolutely_integrable_reflect_real [symmetric]) simp show "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {-pi..pi}" by (simp add: absolutely_integrable_measurable_real borel_measurable_integrable integrable_Dirichlet_kernel) qed lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * f(t+x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * f(t-x)) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * c) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected [OF f periodic, of n] \d \ pi\ by (force intro: absolutely_integrable_on_subinterval)+ lemma absolutely_integrable_mult_Dirichlet_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Dirichlet_kernel n x * (f(t+x) + f(t-x))) absolutely_integrable_on {0..d}" "(\x. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - c)) absolutely_integrable_on {0..d}" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part assms by (simp_all add: distrib_left right_diff_distrib) - lemma integral_reflect_and_add: fixes f :: "real \ 'b::euclidean_space" assumes "integrable (lebesgue_on {-a..a}) f" shows "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" proof (cases "a \ 0") case True have f: "integrable (lebesgue_on {0..a}) f" "integrable (lebesgue_on {-a..0}) f" - using integrable_subinterval_real assms by fastforce+ + using integrable_subinterval assms by fastforce+ then have fm: "integrable (lebesgue_on {0..a}) (\x. f(-x))" using integrable_reflect_real by fastforce have "integral\<^sup>L (lebesgue_on {-a..a}) f = integral\<^sup>L (lebesgue_on {-a..0}) f + integral\<^sup>L (lebesgue_on {0..a}) f" by (simp add: True assms integral_combine) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f(-x)) + integral\<^sup>L (lebesgue_on {0..a}) f" by (metis (no_types) add.inverse_inverse add.inverse_neutral integral_reflect_real) also have "\ = integral\<^sup>L (lebesgue_on {0..a}) (\x. f x + f(-x))" by (simp add: fm f) finally show ?thesis . qed (use assms in auto) lemma Fourier_sum_offset_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\k\2*n. Fourier_coefficient f k * trigonometric_set k t) - l = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) / pi" proof - have fxt: "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto have int: "integrable (lebesgue_on {0..pi}) (Dirichlet_kernel n)" using not_integrable_integral_eq by fastforce have "LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t) = LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * f(x + t) + Dirichlet_kernel n (- x) * f(- x + t)" by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel fxt) also have "\ = (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) + pi * l" using absolutely_integrable_mult_Dirichlet_kernel_reflected_part [OF f periodic order_refl, of n t] apply (simp add: algebra_simps absolutely_integrable_imp_integrable int) done finally show ?thesis by (simp add: Fourier_sum_offset_Dirichlet_kernel assms divide_simps) qed lemma Fourier_sum_limit_Dirichlet_kernel_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l))) \ 0" apply (simp flip: Fourier_sum_limit_pair [OF f]) apply (simp add: Lim_null [where l=l] Fourier_sum_offset_Dirichlet_kernel_half assms) done subsection\Localization principle: convergence only depends on values "nearby"\ proposition Riemann_localization_integral: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and "d > 0" and d: "\x. \x\ < d \ f x = g x" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * g x)) \ 0" (is "?a \ 0") proof - let ?f = "\x. (if \x\ < d then 0 else f x - g x) / sin(x/2) / 2" have eq: "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" for n proof - have "?a n = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * (if \x\ < d then 0 else f x - g x))" apply (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel f g flip: Bochner_Integration.integral_diff) apply (auto simp: d algebra_simps intro: Bochner_Integration.integral_cong) done also have "\ = integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin((n+1/2) * x) * ?f x)" using \d > 0\ by (auto simp: Dirichlet_kernel_def intro: Bochner_Integration.integral_cong) finally show ?thesis . qed show ?thesis unfolding eq proof (rule Riemann_lebesgue_sin_half) obtain B where "B>0" and B: "\x. x \ ({-pi..pi} - {-d<.. \inverse (sin (x/2))\ \ B" using bounded_inverse_sin_half [OF \d > 0\] by metis have "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) absolutely_integrable_on {-pi..pi}" proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" proof (intro measurable) show "f \ borel_measurable (lebesgue_on {-pi..pi})" "g \ borel_measurable (lebesgue_on {-pi..pi})" using absolutely_integrable_on_def f g integrable_imp_measurable by blast+ show "(\x. x) \ borel_measurable (lebesgue_on {-pi..pi})" "(\x. sin (x/2)) \ borel_measurable (lebesgue_on {-pi..pi})" by (intro continuous_intros continuous_imp_measurable_on_sets_lebesgue | force)+ qed auto have "(\x. B * \f x - g x\) absolutely_integrable_on {-pi..pi}" using \B > 0\ by (simp add: f g set_integrable_abs) then show "(\x. B * \f x - g x\) integrable_on {-pi..pi}" using absolutely_integrable_on_def by blast next fix x assume x: "x \ {-pi..pi}" then have [simp]: "sin (x/2) = 0 \ x=0" using \0 < d\ by (force simp: sin_zero_iff) show "norm ((if \x\ < d then 0 else f x - g x) / sin (x/2)) \ B * \f x - g x\" proof (cases "\x\ < d") case False then have "1 \ B * \sin (x/2)\" using \B > 0\ B [of x] x \d > 0\ by (auto simp: abs_less_iff divide_simps split: if_split_asm) then have "1 * \f x - g x\ \ B * \sin (x/2)\ * \f x - g x\" by (metis (full_types) abs_ge_zero mult.commute mult_left_mono) then show ?thesis using False by (auto simp: divide_simps algebra_simps) qed (simp add: d) qed auto then show "(\x. (if \x\ < d then 0 else f x - g x) / sin (x/2) / 2) absolutely_integrable_on {-pi..pi}" using set_integrable_divide by blast qed qed lemma Riemann_localization_integral_range: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. Dirichlet_kernel n x * f x) - integral\<^sup>L (lebesgue_on {-d..d}) (\x. Dirichlet_kernel n x * f x)) \ 0" proof - have *: "(\n. (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f x) - (LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * (if x \ {-d..d} then f x else 0))) \ 0" proof (rule Riemann_localization_integral [OF f _ \0 < d\]) have "f absolutely_integrable_on {-d..d}" using f absolutely_integrable_on_subinterval \d \ pi\ by fastforce moreover have "(\x. if - pi \ x \ x \ pi then if x \ {-d..d} then f x else 0 else 0) = (\x. if x \ {-d..d} then f x else 0)" using \d \ pi\ by force ultimately show "(\x. if x \ {-d..d} then f x else 0) absolutely_integrable_on {-pi..pi}" apply (subst absolutely_integrable_restrict_UNIV [symmetric]) apply (simp flip: absolutely_integrable_restrict_UNIV [of "{-d..d}" f]) done qed auto then show ?thesis using integral_restrict [of "{-d..d}" "{-pi..pi}" "\x. Dirichlet_kernel _ x * f x"] assms by (simp add: if_distrib cong: if_cong) qed lemma Riemann_localization: assumes f: "f absolutely_integrable_on {-pi..pi}" and g: "g absolutely_integrable_on {-pi..pi}" and perf: "\x. f(x + 2*pi) = f x" and perg: "\x. g(x + 2*pi) = g x" and "d > 0" and d: "\x. \x-t\ < d \ f x = g x" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ c \ (\n. \k\n. Fourier_coefficient g k * trigonometric_set k t) \ c" proof - have "(\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * f(x + t)) \ pi * c \ (\n. LINT x|lebesgue_on {-pi..pi}. Dirichlet_kernel n x * g (x + t)) \ pi * c" proof (intro Lim_transform_eq Riemann_localization_integral) show "(\x. f(x + t)) absolutely_integrable_on {-pi..pi}" "(\x. g (x + t)) absolutely_integrable_on {-pi..pi}" using assms by (auto intro: absolutely_integrable_periodic_offset) qed (use assms in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel assms) qed subsection\Localize the earlier integral\ lemma Riemann_localization_integral_range_half: assumes f: "f absolutely_integrable_on {-pi..pi}" and "0 < d" "d \ pi" shows "(\n. (LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f x + f(-x))) - (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f x + f(-x)))) \ 0" proof - have *: "(\x. Dirichlet_kernel n x * f x) absolutely_integrable_on {-d..d}" for n by (metis (full_types) absolutely_integrable_mult_Dirichlet_kernel absolutely_integrable_on_subinterval \d \ pi\ atLeastatMost_subset_iff f neg_le_iff_le) show ?thesis using absolutely_integrable_mult_Dirichlet_kernel [OF f] using Riemann_localization_integral_range [OF assms] apply (simp add: "*" absolutely_integrable_imp_integrable integral_reflect_and_add) apply (simp add: algebra_simps) done qed lemma Fourier_sum_limit_Dirichlet_kernel_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. \k\n. Fourier_coefficient f k * trigonometric_set k t) \ l \ (\n. (LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l))) \ 0" proof - have "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset [OF f, of t] periodic by simp then have int: "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" by auto have "(\n. LINT x|lebesgue_on {0..pi}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0 \ (\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * (f(t+x) + f(t-x) - 2*l)) \ 0" by (rule Lim_transform_eq) (use Riemann_localization_integral_range_half [OF int d] in auto) then show ?thesis by (simp add: Fourier_sum_limit_Dirichlet_kernel_half assms) qed subsection\Make a harmless simplifying tweak to the Dirichlet kernel\ lemma inte_Dirichlet_kernel_mul_expand: assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x = LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2))) \ (integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2))))" proof (cases "0 \ S") case True have *: "{x. x = 0 \ x \ S} \ sets (lebesgue_on S)" using True by (simp add: S sets_restrict_space_iff cong: conj_cong) have bm1: "(\x. Dirichlet_kernel n x * f x) \ borel_measurable (lebesgue_on S)" unfolding Dirichlet_kernel_def by (force intro: * assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) have bm2: "(\x. sin ((n + 1/2) * x) * f x / (2 * sin (x/2))) \ borel_measurable (lebesgue_on S)" by (intro assms measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto have 0: "{0} \ null_sets (lebesgue_on S)" using True by (simp add: S null_sets_restrict_space) show ?thesis apply (intro conjI integral_cong_AE integrable_cong_AE bm1 bm2 AE_I' [OF 0]) unfolding Dirichlet_kernel_def by auto next case False then show ?thesis unfolding Dirichlet_kernel_def by (auto intro!: Bochner_Integration.integral_cong Bochner_Integration.integrable_cong) qed lemma assumes f: "f \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" shows integral_Dirichlet_kernel_mul_expand: "(LINT x|lebesgue_on S. Dirichlet_kernel n x * f x) = (LINT x|lebesgue_on S. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th1") and integrable_Dirichlet_kernel_mul_expand: "integrable (lebesgue_on S) (\x. Dirichlet_kernel n x * f x) \ integrable (lebesgue_on S) (\x. sin((n+1/2) * x) * f x / (2 * sin(x/2)))" (is "?th2") using inte_Dirichlet_kernel_mul_expand [OF assms] by auto proposition Fourier_sum_limit_sine_part: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and d: "0 < d" "d \ pi" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l \ (\n. LINT x|lebesgue_on {0..d}. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)) \ 0" (is "?lhs \ ?\ \ 0") proof - have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have fbm: "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {-pi..pi}" by (force intro: ftmx ftx) let ?\ = "\n. LINT x|lebesgue_on {0..d}. Dirichlet_kernel n x * ((f(t+x) + f(t-x)) - 2*l)" have "?lhs \ ?\ \ 0" by (intro Fourier_sum_limit_Dirichlet_kernel_part assms) moreover have "?\ \ 0 \ ?\ \ 0" proof (rule Lim_transform_eq [OF Lim_transform_eventually]) let ?f = "\n. LINT x|lebesgue_on {0..d}. sin((real n + 1/2) * x) * (1 / (2 * sin(x/2)) - 1/x) * (f(t+x) + f(t-x) - 2*l)" have "?f n = (LINT x|lebesgue_on {-pi..pi}. sin ((n + 1/2) * x) * ((if x \ {0..d} then 1 / (2 * sin (x/2)) - 1/x else 0) * (f(t+x) + f(t-x) - 2*l)))" for n using d by (simp add: integral_restrict if_distrib [of "\u. _ * (u * _)"] mult.assoc flip: atLeastAtMost_iff cong: if_cong) moreover have "\ \ 0" proof (intro Riemann_lebesgue_sin_half absolutely_integrable_bounded_measurable_product_real) have "(\x. 1 / (2 * sin(x/2)) - 1/x) \ borel_measurable (lebesgue_on {0..d})" by (intro measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto then show "(\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) \ borel_measurable (lebesgue_on {-pi..pi})" using d by (simp add: borel_measurable_if_lebesgue_on flip: atLeastAtMost_iff) let ?g = "\x::real. 1 / (2 * sin(x/2)) - 1/x" have limg: "(?g \ ?g a) (at a within {0..d})" \\thanks to Manuel Eberl\ if a: "0 \ a" "a \ d" for a proof - have "(?g \ 0) (at_right 0)" and "(?g \ ?g d) (at_left d)" using d sin_gt_zero[of "d/2"] by (real_asymp simp: field_simps)+ moreover have "(?g \ ?g a) (at a)" if "a > 0" using a that d sin_gt_zero[of "a/2"] that by (real_asymp simp: field_simps) ultimately show ?thesis using that d by (cases "a = 0 \ a = d") (auto simp: at_within_Icc_at at_within_Icc_at_right at_within_Icc_at_left) qed have "((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi}) = ((\x. 1 / (2 * sin(x/2)) - 1/x) ` {0..d})" using d by (auto intro: image_eqI [where x=0]) moreover have "bounded \" by (intro compact_imp_bounded compact_continuous_image) (auto simp: continuous_on limg) ultimately show "bounded ((\x. if x \ {0..d} then 1 / (2 * sin(x/2)) - 1/x else 0) ` {-pi..pi})" by simp qed (auto simp: fbm) ultimately show "?f \ (0::real)" by simp show "\\<^sub>F n in sequentially. ?f n = ?\ n - ?\ n" proof (rule eventually_sequentiallyI [where c = "Suc 0"]) fix n assume "n \ Suc 0" have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2)))" using d apply (subst integrable_Dirichlet_kernel_mul_expand [symmetric]) apply (intro assms absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF fbm] absolutely_integrable_imp_integrable absolutely_integrable_mult_Dirichlet_kernel_reflected_part2 | force)+ done moreover have "integrable (lebesgue_on {0..d}) (\x. sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" proof (rule integrable_cong_AE_imp) let ?g = "\x. Dirichlet_kernel n x * (2 * sin(x/2) / x * (f(t+x) + f(t-x) - 2*l))" have *: "\2 * sin (x/2) / x\ \ 1" for x::real using abs_sin_x_le_abs_x [of "x/2"] by (simp add: divide_simps mult.commute abs_mult) have "bounded ((\x. 1 + (x/2)\<^sup>2) ` {-pi..pi})" by (intro compact_imp_bounded compact_continuous_image continuous_intros) auto then have bo: "bounded ((\x. 2 * sin (x/2) / x) ` {-pi..pi})" using * unfolding bounded_real by blast show "integrable (lebesgue_on {0..d}) ?g" using d apply (intro absolutely_integrable_imp_integrable absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Dirichlet_kernel] absolutely_integrable_bounded_measurable_product_real bo fbm measurable continuous_imp_measurable_on_sets_lebesgue continuous_intros, auto) done show "(\x. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..d})" using d apply (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) done have "Dirichlet_kernel n x * (2 * sin(x/2)) / x = sin ((real n + 1/2) * x) / x" if "0 < x" "x \ d" for x using that d by (simp add: Dirichlet_kernel_def divide_simps sin_zero_pi_iff) then show "AE x in lebesgue_on {0..d}. ?g x = sin ((real n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x" using d by (force intro: AE_I' [where N="{0}"]) qed ultimately have "?f n = (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / (2 * sin(x/2))) - (LINT x|lebesgue_on {0..d}. sin ((n + 1/2) * x) * (f(t+x) + f(t-x) - 2*l) / x)" by (simp add: right_diff_distrib [of _ _ "1/_"] left_diff_distrib) also have "\ = ?\ n - ?\ n" using d by (simp add: measurable_restrict_mono [OF absolutely_integrable_imp_borel_measurable [OF fbm]] integral_Dirichlet_kernel_mul_expand) finally show "?f n = ?\ n - ?\ n" . qed qed ultimately show ?thesis by simp qed subsection\Dini's test for the convergence of a Fourier series\ proposition Fourier_Dini_test: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and int0d: "integrable (lebesgue_on {0..d}) (\x. \f(t+x) + f(t-x) - 2*l\ / x)" and "0 < d" shows "(\n. (\k\n. Fourier_coefficient f k * trigonometric_set k t)) \ l" proof - define \ where "\ \ \n x. sin((n + 1/2) * x) * ((f(t+x) + f(t-x) - 2*l) / x)" have "(\n. LINT x|lebesgue_on {0..pi}. \ n x) \ 0" unfolding lim_sequentially dist_real_def proof (intro allI impI) fix e :: real assume "e > 0" define \ where "\ \ \x. LINT x|lebesgue_on {0..x}. \f(t+x) + f(t-x) - 2*l\ / x" have [simp]: "\ 0 = 0" by (simp add: \_def integral_eq_zero_null_sets) have cont: "continuous_on {0..d} \" - unfolding \_def using indefinite_integral_continuous_1 int0d by blast + unfolding \_def using indefinite_integral_continuous_real int0d by blast with \d > 0\ have "\e>0. \da>0. \x'\{0..d}. \x' - 0\ < da \ \\ x' - \ 0\ < e" by (force simp: continuous_on_real_range dest: bspec [where x=0]) with \e > 0\ obtain k where "k > 0" and k: "\x'. \0 \ x'; x' < k; x' \ d\ \ \\ x'\ < e/2" by (drule_tac x="e/2" in spec) auto define \ where "\ \ min d (min (k/2) pi)" have e2: "\\ \\ < e/2" and "\ > 0" "\ \ pi" unfolding \_def using k \k > 0\ \d > 0\ by auto have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have 1: "\ n absolutely_integrable_on {0..\}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((real n + 1/2) * x)) \ borel_measurable (lebesgue_on {0..\})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((real n + 1/2) * x)) ` {0..\})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. (f(t+x) + f(t-x) - 2*l) / x) \ borel_measurable (lebesgue_on {0..\})" using \d > 0\ unfolding \_def by (intro measurable absolutely_integrable_imp_borel_measurable absolutely_integrable_on_subinterval [OF ftx] absolutely_integrable_on_subinterval [OF ftmx] absolutely_integrable_continuous_real continuous_intros, auto) moreover have "integrable (lebesgue_on {0..\}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" - proof (rule integrable_subinterval_real [of 0 d]) + proof (rule integrable_subinterval [of 0 d]) show "integrable (lebesgue_on {0..d}) (norm \ (\x. (f(t+x) + f(t-x) - 2*l) / x))" using int0d by (subst integrable_cong) (auto simp: o_def) show "{0..\} \ {0..d}" using \d > 0\ by (auto simp: \_def) qed ultimately show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {0..\}" by (auto simp: absolutely_integrable_measurable) qed auto have 2: "\ n absolutely_integrable_on {\..pi}" for n::nat unfolding \_def proof (rule absolutely_integrable_bounded_measurable_product_real) show "(\x. sin ((n + 1/2) * x)) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. sin ((n + 1/2) * x)) ` {\..pi})" by (simp add: bounded_iff) (use abs_sin_le_one in blast) have "(\x. inverse x * (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) have "(\x. 1/x) \ borel_measurable (lebesgue_on {\..pi})" by (auto simp: measurable_completion measurable_restrict_space1) then show "inverse \ borel_measurable (lebesgue_on {\..pi})" by (metis (no_types, lifting) inverse_eq_divide measurable_lebesgue_cong) have "\x\{\..pi}. inverse \x\ \ inverse \" using \0 < \\ by auto then show "bounded (inverse ` {\..pi})" by (auto simp: bounded_iff) show "(\x. f(t+x) + f(t-x) - 2*l) absolutely_integrable_on {\..pi}" proof (rule absolutely_integrable_on_subinterval) show "(\x. (f(t+x) + f(t-x) - 2*l)) absolutely_integrable_on {-pi..pi}" by (fast intro: ftx ftmx absolutely_integrable_on_const) show "{\..pi} \ {-pi..pi}" using \0 < \\ by auto qed qed auto then show "(\x. (f(t+x) + f(t-x) - 2*l) / x) absolutely_integrable_on {\..pi}" by (metis (no_types, lifting) divide_inverse mult.commute set_integrable_cong) qed auto have "(\x. f(t+x) - l) absolutely_integrable_on {-pi..pi}" using ftx by auto moreover have "bounded ((\x. 0) ` {x. \x\ < \})" using bounded_def by blast moreover have "bounded (inverse ` {x. \ \x\ < \})" using \\ > 0\ by (auto simp: divide_simps intro: boundedI [where B = "1/\"]) ultimately have "(\x. (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)) absolutely_integrable_on UNIV" apply (intro absolutely_integrable_bounded_measurable_product_real measurable set_integral_diff) apply (auto simp: lebesgue_on_UNIV_eq measurable_completion simp flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) done moreover have "(if x \ {-pi..pi} then if \x\ < \ then 0 else (f(t+x) - l) / x else 0) = (if \x\ < \ then 0 else inverse x) * (if x \ {-pi..pi} then f(t+x) - l else 0)" for x by (auto simp: divide_simps) ultimately have *: "(\x. if \x\ < \ then 0 else (f(t+x) - l) / x) absolutely_integrable_on {-pi..pi}" by (simp add: flip: absolutely_integrable_restrict_UNIV [where S = "{-pi..pi}"]) have "(\n. LINT x|lebesgue_on {0..pi}. sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x)) \ 0" using Riemann_lebesgue_sin_half [OF *] * by (simp add: integral_reflect_and_add absolutely_integrable_imp_integrable absolutely_integrable_sin_product cong: if_cong) moreover have "sin ((n + 1/2) * x) * (if \x\ < \ then 0 else (f(t+x) - l) / x) + sin ((n + 1/2) * -x) * (if \x\ < \ then 0 else (f(t-x) - l) / -x) = (if \ \x\ < \ then \ n x else 0)" for x n by simp (auto simp: divide_simps algebra_simps \_def) ultimately have "(\n. LINT x|lebesgue_on {0..pi}. (if \ \x\ < \ then \ n x else 0)) \ 0" by simp moreover have "(if 0 \ x \ x \ pi then if \ \x\ < \ then \ n x else 0 else 0) = (if \ \ x \ x \ pi then \ n x else 0)" for x n using \\ > 0\ by auto ultimately have \: "(\n. LINT x|lebesgue_on {\..pi}. \ n x) \ 0" - by (simp flip: integral_restrict_UNIV) + by (simp flip: Lebesgue_Measure.integral_restrict_UNIV) then obtain N::nat where N: "\n. n\N \ \LINT x|lebesgue_on {\..pi}. \ n x\ < e/2" unfolding lim_sequentially dist_real_def by (metis (full_types) \0 < e\ diff_zero half_gt_zero_iff) have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e" if "n \ N" for n::nat proof - have int: "integrable (lebesgue_on {0..pi}) (\ (real n))" by (intro integrable_combine [of concl: 0 pi] absolutely_integrable_imp_integrable) (use \\ > 0\ \\ \ pi\ 1 2 in auto) then have "integral\<^sup>L (lebesgue_on {0..pi}) (\ n) = integral\<^sup>L (lebesgue_on {0..\}) (\ n) + integral\<^sup>L (lebesgue_on {\..pi}) (\ n)" by (rule integral_combine) (use \0 < \\ \\ \ pi\ in auto) moreover have "\integral\<^sup>L (lebesgue_on {0..\}) (\ n)\ \ LINT x|lebesgue_on {0..\}. \f(t + x) + f(t - x) - 2 * l\ / x" proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on {0..\}) (\ (real n))" - by (meson integrable_subinterval_real \\ \ pi\ int atLeastatMost_subset_iff order_refl) + by (meson integrable_subinterval \\ \ pi\ int atLeastatMost_subset_iff order_refl) have "{0..\} \ {0..d}" by (auto simp: \_def) then show "integrable (lebesgue_on {0..\}) (\x. \f(t + x) + f(t - x) - 2 * l\ / x)" - by (rule integrable_subinterval_real [OF int0d]) + by (rule integrable_subinterval [OF int0d]) show "\\ (real n) x\ \ \f(t + x) + f(t - x) - 2 * l\ / x" if "x \ space (lebesgue_on {0..\})" for x using that apply (auto simp: \_def divide_simps abs_mult) by (simp add: mult.commute mult_left_le) qed ultimately have "\integral\<^sup>L (lebesgue_on {0..pi}) (\ n)\ < e/2 + e/2" using N [OF that] e2 unfolding \_def by linarith then show ?thesis by simp qed then show "\N. \n\N. \integral\<^sup>L (lebesgue_on {0..pi}) (\ (real n)) - 0\ < e" by force qed then show ?thesis unfolding \_def using Fourier_sum_limit_sine_part assms pi_gt_zero by blast qed subsection\Cesaro summability of Fourier series using Fejér kernel\ definition Fejer_kernel :: "nat \ real \ real" where "Fejer_kernel \ \n x. if n = 0 then 0 else (\r sin(x/2) = 0") case True have "(\rrx. Fejer_kernel n x * f x) absolutely_integrable_on {-pi..pi}" proof (rule absolutely_integrable_bounded_measurable_product_real) show "Fejer_kernel n \ borel_measurable (lebesgue_on {-pi..pi})" by (simp add: Fejer_kernel_continuous continuous_imp_measurable_on_sets_lebesgue) show "bounded (Fejer_kernel n ` {-pi..pi})" using Fejer_kernel_continuous compact_continuous_image compact_imp_bounded by blast qed (use assms in auto) lemma absolutely_integrable_mult_Fejer_kernel_reflected1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {-pi..pi}" using assms by (force intro: absolutely_integrable_mult_Fejer_kernel absolutely_integrable_periodic_offset) lemma absolutely_integrable_mult_Fejer_kernel_reflected2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {-pi..pi}" proof - have "(\x. f(t - x)) absolutely_integrable_on {-pi..pi}" using assms apply (subst absolutely_integrable_reflect_real [symmetric]) apply (simp add: absolutely_integrable_periodic_offset) done then show ?thesis by (rule absolutely_integrable_mult_Fejer_kernel) qed lemma absolutely_integrable_mult_Fejer_kernel_reflected3: shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel by blast lemma absolutely_integrable_mult_Fejer_kernel_reflected_part1: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t + x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected1]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part2: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * f(t - x)) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part3: assumes "d \ pi" shows "(\x. Fejer_kernel n x * c) absolutely_integrable_on {0..d}" by (rule absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel_reflected2]) (auto simp: assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part4: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * (f(t + x) + f(t - x))) absolutely_integrable_on {0..d}" unfolding distrib_left by (intro set_integral_add absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms) lemma absolutely_integrable_mult_Fejer_kernel_reflected_part5: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "d \ pi" shows "(\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - c)) absolutely_integrable_on {0..d}" unfolding distrib_left right_diff_distrib by (intro set_integral_add set_integral_diff absolutely_integrable_on_const absolutely_integrable_mult_Fejer_kernel_reflected_part1 absolutely_integrable_mult_Fejer_kernel_reflected_part2 assms, auto) lemma Fourier_sum_offset_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and "n > 0" shows "(\rk\2*r. Fourier_coefficient f k * trigonometric_set k t) / n - l = (LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi" proof - have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = (\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - l)" by (simp add: sum_subtractf) also have "\ = (\r = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" proof - have "integrable (lebesgue_on {0..pi}) (\x. Dirichlet_kernel i x * (f(t + x) + f(t - x) - 2 * l))" for i using absolutely_integrable_mult_Dirichlet_kernel_reflected_part2(2) f periodic by (force simp: intro!: absolutely_integrable_imp_integrable) then show ?thesis using \n > 0\ apply (simp add: Fejer_kernel_def flip: sum_divide_distrib) apply (simp add: sum_distrib_right flip: Bochner_Integration.integral_sum [symmetric]) done qed finally have "(\rk\2 * r. Fourier_coefficient f k * trigonometric_set k t) - real n * l = real n * ((LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * (f(t + x) + f(t - x) - 2 * l)) / pi)" . with \n > 0\ show ?thesis by (auto simp: mult.commute divide_simps) qed lemma Fourier_sum_limit_Fejer_kernel_half: fixes n::nat assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) \ l \ ((\n. integral\<^sup>L (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * ((f(t + x) + f(t - x)) - 2*l))) \ 0)" (is "?lhs = ?rhs") proof - have "?lhs \ (\n. ((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n - l) \ 0" by (simp add: LIM_zero_iff) also have "\ \ (\n. ((((\rk\2*r. Fourier_coefficient f k * trigonometric_set k t)) / n) - l) * pi) \ 0" using real_tendsto_mult_right_iff [OF pi_neq_zero] by simp also have "\ \ ?rhs" apply (intro Lim_transform_eq [OF Lim_transform_eventually [of "\n. 0"]] eventually_sequentiallyI [of 1]) apply (simp_all add: Fourier_sum_offset_Fejer_kernel_half assms) done finally show ?thesis . qed lemma has_integral_Fejer_kernel: "has_bochner_integral (lebesgue_on {-pi..pi}) (Fejer_kernel n) (if n = 0 then 0 else pi)" proof - have "has_bochner_integral (lebesgue_on {-pi..pi}) (\x. (\rrx. (\rr 0" by (simp add: Fejer_kernel) theorem Fourier_Fejer_Cesaro_summable: assumes f: "f absolutely_integrable_on {-pi..pi}" and periodic: "\x. f(x + 2*pi) = f x" and fl: "(f \ l) (at t within atMost t)" and fr: "(f \ r) (at t within atLeast t)" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k t) / n) \ (l+r) / 2" proof - define h where "h \ \u. (f(t+u) + f(t-u)) - (l+r)" have "(\n. LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u) \ 0" proof - have h0: "(h \ 0) (at 0 within atLeast 0)" proof - have l0: "((\u. f(t-u) - l) \ 0) (at 0 within {0..})" using fl unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t-x" in bspec) apply (auto simp: dist_norm) done have r0: "((\u. f(t + u) - r) \ 0) (at 0 within {0..})" using fr unfolding Lim_within apply (elim all_forward imp_forward ex_forward conj_forward asm_rl, clarify) apply (drule_tac x="t+x" in bspec) apply (auto simp: dist_norm) done show ?thesis using tendsto_add [OF l0 r0] by (simp add: h_def algebra_simps) qed show ?thesis unfolding lim_sequentially dist_real_def diff_0_right proof clarify fix e::real assume "e > 0" then obtain x' where "0 < x'" "\x. \0 < x; x < x'\ \ \h x\ < e / (2 * pi)" using h0 unfolding Lim_within dist_real_def by (auto simp: dest: spec [where x="e/2/pi"]) then obtain \ where "0 < \" "\ < pi" and \: "\x. 0 < x \ x \ \ \ \h x\ < e/2/pi" apply (intro that [where \="min x' pi/2"], auto) using m2pi_less_pi by linarith have ftx: "(\x. f(t+x)) absolutely_integrable_on {-pi..pi}" using absolutely_integrable_periodic_offset assms by auto then have ftmx: "(\x. f(t-x)) absolutely_integrable_on {-pi..pi}" by (simp flip: absolutely_integrable_reflect_real [where f= "(\x. f(t+x))"]) have h_aint: "h absolutely_integrable_on {-pi..pi}" unfolding h_def by (intro absolutely_integrable_on_const set_integral_diff set_integral_add, auto simp: ftx ftmx) have "(\n. LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ 0" proof (rule Lim_null_comparison) define \ where "\ \ \n. (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * sin(x/2) ^ 2)) / n" show "\\<^sub>F n in sequentially. norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" proof (rule eventually_sequentiallyI) fix n::nat assume "n \ 1" have hint: "(\x. h x / (2 * sin(x/2) ^ 2)) absolutely_integrable_on {\..pi}" unfolding divide_inverse mult.commute [of "h _"] proof (rule absolutely_integrable_bounded_measurable_product_real) have cont: "continuous_on {\..pi} (\x. inverse (2 * (sin (x * inverse 2))\<^sup>2))" using \0 < \\ by (intro continuous_intros) (auto simp: sin_zero_pi_iff) show "(\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) \ borel_measurable (lebesgue_on {\..pi})" using \0 < \\ by (intro cont continuous_imp_measurable_on_sets_lebesgue) auto show "bounded ((\x. inverse (2 * (sin (x * inverse 2))\<^sup>2)) ` {\..pi})" using cont by (simp add: compact_continuous_image compact_imp_bounded) show "h absolutely_integrable_on {\..pi}" using \0 < \\ \\ < pi\ by (auto intro: absolutely_integrable_on_subinterval [OF h_aint]) qed auto then have *: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2))" by (simp add: absolutely_integrable_measurable o_def) define \ where "\ \ \x. (if n = 0 then 0 else if x = 0 then n/2 else (sin (real n / 2 * x))\<^sup>2 / (2 * real n * (sin (x/2))\<^sup>2)) * h x" have "\LINT x|lebesgue_on {\..pi}. \ x\ \ (LINT x|lebesgue_on {\..pi}. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" proof (rule integral_abs_bound_integral) show **: "integrable (lebesgue_on {\..pi}) (\x. \h x\ / (2 * (sin (x/2))\<^sup>2) / n)" using Bochner_Integration.integrable_mult_left [OF *, of "1/n"] by (simp add: field_simps) show \: "\\ x\ \ \h x\ / (2 * (sin (x/2))\<^sup>2) / real n" if "x \ space (lebesgue_on {\..pi})" for x using that \0 < \\ apply (simp add: \_def divide_simps mult_less_0_iff abs_mult) apply (auto simp: square_le_1 mult_left_le_one_le) done show "integrable (lebesgue_on {\..pi}) \" proof (rule measurable_bounded_by_integrable_imp_lebesgue_integrable [OF _ **]) let ?g = "\x. \h x\ / (2 * sin(x/2) ^ 2) / n" have ***: "integrable (lebesgue_on {\..pi}) (\x. (sin (n/2 * x))\<^sup>2 * (inverse (2 * (sin (x/2))\<^sup>2) * h x))" proof (rule absolutely_integrable_imp_integrable [OF absolutely_integrable_bounded_measurable_product_real]) show "(\x. (sin (real n / 2 * x))\<^sup>2) \ borel_measurable (lebesgue_on {\..pi})" by (intro continuous_imp_measurable_on_sets_lebesgue continuous_intros) auto show "bounded ((\x. (sin (real n / 2 * x))\<^sup>2) ` {\..pi})" by (force simp: square_le_1 intro: boundedI [where B=1]) show "(\x. inverse (2 * (sin (x/2))\<^sup>2) * h x) absolutely_integrable_on {\..pi}" using hint by (simp add: divide_simps) qed auto show "\ \ borel_measurable (lebesgue_on {\..pi})" apply (rule borel_measurable_integrable) apply (rule integrable_cong [where f = "\x. sin(n / 2 * x) ^ 2 / (2 * n * sin(x/2) ^ 2) * h x", OF refl, THEN iffD1]) using \0 < \\ ** apply (force simp: \_def divide_simps algebra_simps mult_less_0_iff abs_mult) using Bochner_Integration.integrable_mult_left [OF ***, of "1/n"] by (simp add: field_simps) show "norm (\ x) \ ?g x" if "x \ {\..pi}" for x using that \ by (simp add: \_def) qed auto qed then show "norm (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x) \ \ n" by (simp add: Fejer_kernel \_def \_def flip: Bochner_Integration.integral_divide [OF *]) qed show "\ \ 0" unfolding \_def divide_inverse by (simp add: tendsto_mult_right_zero lim_inverse_n) qed then obtain N where N: "\n. n \ N \ \LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" unfolding lim_sequentially by (metis half_gt_zero_iff norm_conv_dist real_norm_def \e > 0\) show "\N. \n\N. \(LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u)\ < e" proof (intro exI allI impI) fix n :: nat assume n: "n \ max 1 N" with N have 1: "\LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x\ < e/2" by simp have "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ integral\<^sup>L (lebesgue_on {0..pi}) (Fejer_kernel n)" using \\ < pi\ has_bochner_integral_iff has_integral_Fejer_kernel_half by (force intro!: integral_mono_lebesgue_on_AE) also have "\ \ pi" using has_integral_Fejer_kernel_half by (simp add: has_bochner_integral_iff) finally have int_le_pi: "integral\<^sup>L (lebesgue_on {0..\}) (Fejer_kernel n) \ pi" . have 2: "\LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x\ \ (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * e/2/pi)" proof - have eq: "integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * h x) = integral\<^sup>L (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" proof (rule integral_cong_AE) have \: "{u. u \ 0 \ P u} \ {0..\} = {0} \ Collect P \ {0..\}" "{u. u \ 0 \ P u} \ {0..\} = (Collect P \ {0..\}) - {0}" for P using \0 < \\ by auto have *: "- {0} \ A \ {0..\} = A \ {0..\} - {0}" for A by auto show "(\x. Fejer_kernel n x * h x) \ borel_measurable (lebesgue_on {0..\})" using \\ < pi\ by (intro absolutely_integrable_imp_borel_measurable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel], auto) then show "(\x. Fejer_kernel n x * (if x = 0 then 0 else h x)) \ borel_measurable (lebesgue_on {0..\})" apply (simp add: in_borel_measurable Ball_def vimage_def Collect_conj_eq Collect_imp_eq * flip: Collect_neg_eq) apply (elim all_forward imp_forward asm_rl) using \0 < \\ apply (auto simp: \ insert_sets_lebesgue_on cong: conj_cong) done have 0: "{0} \ null_sets (lebesgue_on {0..\})" using \0 < \\ by (simp add: null_sets_restrict_space) then show "AE x in lebesgue_on {0..\}. Fejer_kernel n x * h x = Fejer_kernel n x * (if x = 0 then 0 else h x)" by (auto intro: AE_I' [OF 0]) qed show ?thesis unfolding eq proof (rule integral_abs_bound_integral) have "(\x. if x = 0 then 0 else h x) absolutely_integrable_on {- pi..pi}" proof (rule absolutely_integrable_spike [OF h_aint]) show "negligible {0}" by auto qed auto with \0 < \\ \\ < pi\ show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * (if x = 0 then 0 else h x))" by (intro absolutely_integrable_imp_integrable h_aint absolutely_integrable_on_subinterval [OF absolutely_integrable_mult_Fejer_kernel]) auto show "integrable (lebesgue_on {0..\}) (\x. Fejer_kernel n x * e / 2 / pi)" by (simp add: absolutely_integrable_imp_integrable \\ < pi\ absolutely_integrable_mult_Fejer_kernel_reflected_part3 less_eq_real_def) show "\Fejer_kernel n x * (if x = 0 then 0 else h x)\ \ Fejer_kernel n x * e / 2 / pi" if "x \ space (lebesgue_on {0..\})" for x using that \ [of x] \e > 0\ by (auto simp: abs_mult eq simp flip: times_divide_eq_right intro: mult_left_mono) qed qed have 3: "\ \ e/2" using int_le_pi \0 < e\ by (simp add: divide_simps mult.commute [of e]) have "integrable (lebesgue_on {0..pi}) (\x. Fejer_kernel n x * h x)" unfolding h_def by (simp add: absolutely_integrable_imp_integrable absolutely_integrable_mult_Fejer_kernel_reflected_part5 assms) then have "LINT x|lebesgue_on {0..pi}. Fejer_kernel n x * h x = (LINT x|lebesgue_on {0..\}. Fejer_kernel n x * h x) + (LINT x|lebesgue_on {\..pi}. Fejer_kernel n x * h x)" by (rule integral_combine) (use \0 < \\ \\ < pi\ in auto) then show "\LINT u|lebesgue_on {0..pi}. Fejer_kernel n u * h u\ < e" using 1 2 3 by linarith qed qed qed then show ?thesis unfolding h_def by (simp add: Fourier_sum_limit_Fejer_kernel_half assms add_divide_distrib) qed corollary Fourier_Fejer_Cesaro_summable_simple: assumes f: "continuous_on UNIV f" and periodic: "\x. f(x + 2*pi) = f x" shows "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ f x" proof - have "(\n. (\mk\2*m. Fourier_coefficient f k * trigonometric_set k x) / n) \ (f x + f x) / 2" proof (rule Fourier_Fejer_Cesaro_summable) show "f absolutely_integrable_on {- pi..pi}" using absolutely_integrable_continuous_real continuous_on_subset f by blast show "(f \ f x) (at x within {..x})" "(f \ f x) (at x within {x..})" using Lim_at_imp_Lim_at_within continuous_on_def f by blast+ qed (auto simp: periodic Lim_at_imp_Lim_at_within continuous_on_def f) then show ?thesis by simp qed end diff --git a/thys/Fourier/Fourier_Aux1.thy b/thys/Fourier/Fourier_Aux1.thy deleted file mode 100644 --- a/thys/Fourier/Fourier_Aux1.thy +++ /dev/null @@ -1,577 +0,0 @@ -section\Material already in development version as of September 2019\ - -theory Fourier_Aux1 - imports "HOL-Analysis.Analysis" -begin - -lemma square_powr_half [simp]: - fixes x::real shows "x\<^sup>2 powr (1/2) = \x\" - by (simp add: powr_half_sqrt) - -lemma frac_1_eq: "frac (x+1) = frac x" - by (simp add: frac_def) - -context comm_monoid_set -begin - -lemma in_pairs: "F g {2*m..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {m..n}" -proof (induction n) - case 0 - show ?case - by (cases "m=0") auto -next - case (Suc n) - then show ?case - by (auto simp: assoc split: if_split_asm) -qed - -lemma in_pairs_0: "F g {..Suc(2*n)} = F (\i. g(2*i) \<^bold>* g(Suc(2*i))) {..n}" - using in_pairs [of _ 0 n] by (simp add: atLeast0AtMost) - -end - -context linordered_nonzero_semiring -begin - -lemma numeral_le_iff: "numeral m \ numeral n \ m \ n" -proof - - have "of_nat (numeral m) \ of_nat (numeral n) \ m \ n" - by (simp only: less_eq_num_def nat_of_num_numeral of_nat_le_iff) - then show ?thesis by simp -qed - -lemma one_le_numeral: "1 \ numeral n" - using numeral_le_iff [of num.One n] by (simp add: numeral_One) - -lemma numeral_le_one_iff: "numeral n \ 1 \ n \ num.One" - using numeral_le_iff [of n num.One] by (simp add: numeral_One) - -lemma numeral_less_iff: "numeral m < numeral n \ m < n" -proof - - have "of_nat (numeral m) < of_nat (numeral n) \ m < n" - unfolding less_num_def nat_of_num_numeral of_nat_less_iff .. - then show ?thesis by simp -qed - -lemma not_numeral_less_one: "\ numeral n < 1" - using numeral_less_iff [of n num.One] by (simp add: numeral_One) - -lemma one_less_numeral_iff: "1 < numeral n \ num.One < n" - using numeral_less_iff [of num.One n] by (simp add: numeral_One) - -lemma zero_le_numeral: "0 \ numeral n" - using dual_order.trans one_le_numeral zero_le_one by blast - -lemma zero_less_numeral: "0 < numeral n" - using less_linear not_numeral_less_one order.strict_trans zero_less_one by blast - -lemma not_numeral_le_zero: "\ numeral n \ 0" - by (simp add: not_le zero_less_numeral) - -lemma not_numeral_less_zero: "\ numeral n < 0" - by (simp add: not_less zero_le_numeral) - -lemmas le_numeral_extra = - zero_le_one not_one_le_zero - order_refl [of 0] order_refl [of 1] - -lemmas less_numeral_extra = - zero_less_one not_one_less_zero - less_irrefl [of 0] less_irrefl [of 1] - -lemmas le_numeral_simps [simp] = - numeral_le_iff - one_le_numeral - numeral_le_one_iff - zero_le_numeral - not_numeral_le_zero - -lemmas less_numeral_simps [simp] = - numeral_less_iff - one_less_numeral_iff - not_numeral_less_one - zero_less_numeral - not_numeral_less_zero - -end - -lemma powr_numeral [simp]: "0 \ x \ x powr (numeral n :: real) = x ^ (numeral n)" - by (metis less_le power_zero_numeral powr_0 of_nat_numeral powr_realpow) - -context linordered_semidom -begin - -lemma power_mono_iff [simp]: - shows "\a \ 0; b \ 0; n>0\ \ a ^ n \ b ^ n \ a \ b" - using power_mono [of a b] power_strict_mono [of b a] not_le by auto - -end - -lemma tendsto_sup[tendsto_intros]: - fixes X :: "'a \ 'b::ordered_euclidean_space" - assumes "(X \ x) net" "(Y \ y) net" - shows "((\i. sup (X i) (Y i)) \ sup x y) net" - unfolding sup_max eucl_sup by (intro assms tendsto_intros) - -lemma tendsto_inf[tendsto_intros]: - fixes X :: "'a \ 'b::ordered_euclidean_space" - assumes "(X \ x) net" "(Y \ y) net" - shows "((\i. inf (X i) (Y i)) \ inf x y) net" - unfolding inf_min eucl_inf by (intro assms tendsto_intros) - -lemma continuous_on_max [continuous_intros]: - fixes f g :: "'a::topological_space \ 'b::linorder_topology" - shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. max (f x) (g x))" - by (auto simp: continuous_on_def intro!: tendsto_max) - -lemma continuous_on_min [continuous_intros]: - fixes f g :: "'a::topological_space \ 'b::linorder_topology" - shows "continuous_on A f \ continuous_on A g \ continuous_on A (\x. min (f x) (g x))" - by (auto simp: continuous_on_def intro!: tendsto_min) - -lemma Ints_eq_abs_less1: - fixes x:: "'a :: linordered_idom" - shows "\x \ Ints; y \ Ints\ \ x = y \ abs (x-y) < 1" - using eq_iff_diff_eq_0 by (fastforce intro: Ints_nonzero_abs_less1) - -lemma indicator_times_eq_if: - fixes f :: "'a \ 'b::comm_ring_1" - shows "indicator S x * f x = (if x \ S then f x else 0)" "f x * indicator S x = (if x \ S then f x else 0)" - by auto - -lemma indicator_scaleR_eq_if: - fixes f :: "'a \ 'b::real_vector" - shows "indicator S x *\<^sub>R f x = (if x \ S then f x else 0)" - by simp - -lemma lebesgue_on_UNIV_eq: "lebesgue_on UNIV = lebesgue" -proof - - have "measure_of UNIV (sets lebesgue) (emeasure lebesgue) = lebesgue" - by (metis measure_of_of_measure space_borel space_completion space_lborel) - then show ?thesis - by (auto simp: restrict_space_def) -qed - -lemma mem_limsup_iff: "x \ limsup A \ (\\<^sub>F n in sequentially. x \ A n)" - by (simp add: Limsup_def) (metis (mono_tags) eventually_mono not_frequently) - -lemma integral_restrict_Int: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "S \ sets lebesgue" "T \ sets lebesgue" - shows "integral\<^sup>L (lebesgue_on T) (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on (S \ T)) f" -proof - - have "(\x. indicat_real T x *\<^sub>R (if x \ S then f x else 0)) = (\x. indicat_real (S \ T) x *\<^sub>R f x)" - by (force simp: indicator_def) - then show ?thesis - by (simp add: assms sets.Int Bochner_Integration.integral_restrict_space) -qed - -lemma space_lebesgue_on [simp]: "space (lebesgue_on S) = S" - by (simp add: space_restrict_space) - -corollary absolutely_integrable_on_const [simp]: - fixes c :: "'a::euclidean_space" - assumes "S \ lmeasurable" - shows "(\x. c) absolutely_integrable_on S" - by (metis (full_types) assms absolutely_integrable_integrable_bound integrable_on_const order_refl) - -lemma integral_restrict_UNIV: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "S \ sets lebesgue" - shows "integral\<^sup>L lebesgue (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on S) f" - using integral_restrict_Int [of S UNIV f] assms - by (simp add: lebesgue_on_UNIV_eq) - -lemma tendsto_eventually: "eventually (\x. f x = c) F \ filterlim f (nhds c) F" - by (simp add: eventually_mono eventually_nhds_x_imp_x filterlim_iff) - -lemma content_cbox_plus: - fixes x :: "'a::euclidean_space" - shows "content(cbox x (x + h *\<^sub>R One)) = (if h \ 0 then h ^ DIM('a) else 0)" - by (simp add: algebra_simps content_cbox_if box_eq_empty) - -lemma finite_measure_lebesgue_on: "S \ lmeasurable \ finite_measure (lebesgue_on S)" - by (auto simp: finite_measureI fmeasurable_def emeasure_restrict_space) - -lemma measurable_bounded_by_integrable_imp_absolutely_integrable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" - and "g integrable_on S" and "\x. x \ S \ norm(f x) \ (g x)" - shows "f absolutely_integrable_on S" - using assms absolutely_integrable_integrable_bound measurable_bounded_by_integrable_imp_integrable by blast - -lemma absolutely_integrable_measurable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "S \ sets lebesgue" - shows "f absolutely_integrable_on S \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (norm \ f)" - (is "?lhs = ?rhs") -proof - assume L: ?lhs - then have "f \ borel_measurable (lebesgue_on S)" - by (simp add: absolutely_integrable_on_def integrable_imp_measurable) - then show ?rhs - using assms set_integrable_norm [of lebesgue S f] L - by (simp add: integrable_restrict_space set_integrable_def) -next - assume ?rhs then show ?lhs - using assms integrable_on_lebesgue_on - by (metis absolutely_integrable_integrable_bound comp_def eq_iff measurable_bounded_by_integrable_imp_integrable) -qed - -lemma lim_const_over_n [tendsto_intros]: - fixes a :: "'a::real_normed_field" - shows "(\n. a / of_nat n) \ 0" - using tendsto_mult [OF tendsto_const [of a] lim_1_over_n] by simp - -lemma add_mono_ennreal: "x < ennreal y \ x' < ennreal y' \ x + x' < ennreal (y + y')" - by (metis (full_types) add_strict_mono ennreal_less_zero_iff ennreal_plus less_le not_less zero_le) - -lemma enn2real_sum:"(\i. i \ I \ f i < top) \ enn2real (sum f I) = sum (enn2real \ f) I" - by (induction I rule: infinite_finite_induct) (auto simp: enn2real_plus) - -lemma sets_lebesgue_outer_open: - fixes e::real - assumes S: "S \ sets lebesgue" and "e > 0" - obtains T where "open T" "S \ T" "(T - S) \ lmeasurable" "emeasure lebesgue (T - S) < ennreal e" -proof - - obtain S' where S': "S \ S'" "S' \ sets borel" - and null: "S' - S \ null_sets lebesgue" - and em: "emeasure lebesgue S = emeasure lborel S'" - using completion_upper[of S lborel] S by auto - then have f_S': "S' \ sets borel" - using S by (auto simp: fmeasurable_def) - with outer_regular_lborel[OF _ \0] - obtain U where U: "open U" "S' \ U" "emeasure lborel (U - S') < e" - by blast - show thesis - proof - show "open U" "S \ U" - using f_S' U S' by auto - have "(U - S) = (U - S') \ (S' - S)" - using S' U by auto - then have eq: "emeasure lebesgue (U - S) = emeasure lborel (U - S')" - using null by (simp add: U(1) emeasure_Un_null_set f_S' sets.Diff) - have "(U - S) \ sets lebesgue" - by (simp add: S U(1) sets.Diff) - then show "(U - S) \ lmeasurable" - unfolding fmeasurable_def using U(3) eq less_le_trans by fastforce - with eq U show "emeasure lebesgue (U - S) < e" - by (simp add: eq) - qed -qed - -lemma sets_lebesgue_inner_closed: - fixes e::real - assumes "S \ sets lebesgue" "e > 0" - obtains T where "closed T" "T \ S" "S-T \ lmeasurable" "emeasure lebesgue (S - T) < ennreal e" -proof - - have "-S \ sets lebesgue" - using assms by (simp add: Compl_in_sets_lebesgue) - then obtain T where "open T" "-S \ T" - and T: "(T - -S) \ lmeasurable" "emeasure lebesgue (T - -S) < e" - using sets_lebesgue_outer_open assms by blast - show thesis - proof - show "closed (-T)" - using \open T\ by blast - show "-T \ S" - using \- S \ T\ by auto - show "S - ( -T) \ lmeasurable" "emeasure lebesgue (S - (- T)) < e" - using T by (auto simp: Int_commute) - qed -qed - -corollary eventually_ae_filter_negligible: - "eventually P (ae_filter lebesgue) \ (\N. negligible N \ {x. \ P x} \ N)" - by (auto simp: completion.AE_iff_null_sets negligible_iff_null_sets null_sets_completion_subset) - - -lemma borel_measurable_affine: - fixes f :: "'a :: euclidean_space \ 'b :: euclidean_space" - assumes f: "f \ borel_measurable lebesgue" and "c \ 0" - shows "(\x. f(t + c *\<^sub>R x)) \ borel_measurable lebesgue" -proof - - { fix a b - have "{x. f x \ cbox a b} \ sets lebesgue" - using f cbox_borel lebesgue_measurable_vimage_borel by blast - then have "(\x. (x - t) /\<^sub>R c) ` {x. f x \ cbox a b} \ sets lebesgue" - proof (rule differentiable_image_in_sets_lebesgue) - show "(\x. (x - t) /\<^sub>R c) differentiable_on {x. f x \ cbox a b}" - unfolding differentiable_on_def differentiable_def - by (rule \c \ 0\ derivative_eq_intros strip exI | simp)+ - qed auto - moreover - have "{x. f(t + c *\<^sub>R x) \ cbox a b} = (\x. (x-t) /\<^sub>R c) ` {x. f x \ cbox a b}" - using \c \ 0\ by (auto simp: image_def) - ultimately have "{x. f(t + c *\<^sub>R x) \ cbox a b} \ sets lebesgue" - by (auto simp: borel_measurable_vimage_closed_interval) } - then show ?thesis - by (subst lebesgue_on_UNIV_eq [symmetric]; auto simp: borel_measurable_vimage_closed_interval) -qed - - -corollary lebesgue_real_affine: - "c \ 0 \ lebesgue = density (distr lebesgue lebesgue (\x. t + c * x)) (\_. ennreal (abs c))" - using lebesgue_affine_euclidean [where c= "\x::real. c"] by simp - -lemma nn_integral_real_affine_lebesgue: - fixes c :: real assumes f[measurable]: "f \ borel_measurable lebesgue" and c: "c \ 0" - shows "(\\<^sup>+x. f x \lebesgue) = ennreal\c\ * (\\<^sup>+x. f(t + c * x) \lebesgue)" -proof - - have "(\\<^sup>+x. f x \lebesgue) = (\\<^sup>+x. f x \density (distr lebesgue lebesgue (\x. t + c * x)) (\x. ennreal \c\))" - using lebesgue_real_affine c by auto - also have "\ = \\<^sup>+ x. ennreal \c\ * f x \distr lebesgue lebesgue (\x. t + c * x)" - by (subst nn_integral_density) auto - also have "\ = ennreal \c\ * integral\<^sup>N (distr lebesgue lebesgue (\x. t + c * x)) f" - using f measurable_distr_eq1 nn_integral_cmult by blast - also have "\ = \c\ * (\\<^sup>+x. f(t + c * x) \lebesgue)" - using lebesgue_affine_measurable[where c= "\x::real. c"] - by (subst nn_integral_distr) (force+) - finally show ?thesis . -qed - -lemma lebesgue_integrable_real_affine: - fixes f :: "real \ 'a :: euclidean_space" - assumes f: "integrable lebesgue f" and "c \ 0" - shows "integrable lebesgue (\x. f(t + c * x))" -proof - - have "(\x. norm (f x)) \ borel_measurable lebesgue" - by (simp add: borel_measurable_integrable f) - then show ?thesis - using assms borel_measurable_affine [of f c] - unfolding integrable_iff_bounded - by (subst (asm) nn_integral_real_affine_lebesgue[where c=c and t=t]) (auto simp: ennreal_mult_less_top) -qed - -lemma lebesgue_integrable_real_affine_iff: - fixes f :: "real \ 'a :: euclidean_space" - shows "c \ 0 \ integrable lebesgue (\x. f(t + c * x)) \ integrable lebesgue f" - using lebesgue_integrable_real_affine[of f c t] - lebesgue_integrable_real_affine[of "\x. f(t + c * x)" "1/c" "-t/c"] - by (auto simp: field_simps) - -lemma\<^marker>\tag important\ lebesgue_integral_real_affine: - fixes f :: "real \ 'a :: euclidean_space" and c :: real - assumes c: "c \ 0" shows "(\x. f x \ lebesgue) = \c\ *\<^sub>R (\x. f(t + c * x) \lebesgue)" -proof cases - have "(\x. t + c * x) \ lebesgue \\<^sub>M lebesgue" - using lebesgue_affine_measurable[where c= "\x::real. c"] \c \ 0\ by simp - moreover - assume "integrable lebesgue f" - ultimately show ?thesis - by (subst lebesgue_real_affine[OF c, of t]) (auto simp: integral_density integral_distr) -next - assume "\ integrable lebesgue f" with c show ?thesis - by (simp add: lebesgue_integrable_real_affine_iff not_integrable_integral_eq) -qed - -lemma has_bochner_integral_lebesgue_real_affine_iff: - fixes i :: "'a :: euclidean_space" - shows "c \ 0 \ - has_bochner_integral lebesgue f i \ - has_bochner_integral lebesgue (\x. f(t + c * x)) (i /\<^sub>R \c\)" - unfolding has_bochner_integral_iff lebesgue_integrable_real_affine_iff - by (simp_all add: lebesgue_integral_real_affine[symmetric] divideR_right cong: conj_cong) - -lemma absolutely_integrable_measurable_real: - fixes f :: "'a::euclidean_space \ real" - assumes "S \ sets lebesgue" - shows "f absolutely_integrable_on S \ - f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (\x. \f x\)" - by (simp add: absolutely_integrable_measurable assms o_def) - -lemma continuous_imp_integrable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "continuous_on (cbox a b) f" - shows "integrable (lebesgue_on (cbox a b)) f" -proof - - have "f absolutely_integrable_on cbox a b" - by (simp add: absolutely_integrable_continuous assms) - then show ?thesis - by (simp add: integrable_restrict_space set_integrable_def) -qed - -lemma continuous_imp_integrable_real: - fixes f :: "real \ 'b::euclidean_space" - assumes "continuous_on {a..b} f" - shows "integrable (lebesgue_on {a..b}) f" - by (metis assms continuous_imp_integrable interval_cbox) - -lemma integral_abs_bound_integral: - fixes f :: "'a::euclidean_space \ real" - assumes "integrable M f" "integrable M g" "\x. x \ space M \ \f x\ \ g x" - shows "\\x. f x \M\ \ (\x. g x \M)" -proof - - have "LINT x|M. \f x\ \ integral\<^sup>L M g" - by (simp add: assms integral_mono) - then show ?thesis - by (metis dual_order.trans integral_abs_bound) -qed - -lemma ennreal_tendsto_0_iff: "(\n. f n \ 0) \ ((\n. ennreal (f n)) \ 0) \ (f \ 0)" - by (metis (mono_tags) ennreal_0 eventuallyI order_refl tendsto_ennreal_iff) - -lemma borel_measurable_continuous_onI: "continuous_on UNIV f \ f \ borel_measurable borel" - by (drule borel_measurable_continuous_on_restrict) simp - -lemma has_integral_integral_lebesgue_on: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" - shows "(f has_integral (integral\<^sup>L (lebesgue_on S) f)) S" -proof - - let ?f = "\x. if x \ S then f x else 0" - have "integrable lebesgue (\x. indicat_real S x *\<^sub>R f x)" - using indicator_scaleR_eq_if [of S _ f] assms - by (metis (full_types) integrable_restrict_space sets.Int_space_eq2) - then have "integrable lebesgue ?f" - using indicator_scaleR_eq_if [of S _ f] assms by auto - then have "(?f has_integral (integral\<^sup>L lebesgue ?f)) UNIV" - by (rule has_integral_integral_lebesgue) - then have "(f has_integral (integral\<^sup>L lebesgue ?f)) S" - using has_integral_restrict_UNIV by blast - moreover - have "S \ space lebesgue \ sets lebesgue" - by (simp add: assms) - then have "(integral\<^sup>L lebesgue ?f) = (integral\<^sup>L (lebesgue_on S) f)" - by (simp add: integral_restrict_space indicator_scaleR_eq_if) - ultimately show ?thesis - by auto -qed - -lemma lebesgue_integral_eq_integral: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "integrable (lebesgue_on S) f" "S \ sets lebesgue" - shows "integral\<^sup>L (lebesgue_on S) f = integral S f" - by (metis has_integral_integral_lebesgue_on assms integral_unique) - -declare isCont_Ln' [continuous_intros] - -lemma absolutely_integrable_imp_integrable: - assumes "f absolutely_integrable_on S" "S \ sets lebesgue" - shows "integrable (lebesgue_on S) f" - by (meson assms integrable_restrict_space set_integrable_def sets.Int sets.top) - -lemma insert_null_sets_iff [simp]: "insert a N \ null_sets lebesgue \ N \ null_sets lebesgue" - (is "?lhs = ?rhs") -proof - assume ?lhs then show ?rhs - by (meson completion.complete2 subset_insertI) -next - assume ?rhs then show ?lhs - by (simp add: null_sets.insert_in_sets null_setsI) -qed - -lemma - fixes a::real - shows lmeasurable_interval [iff]: "{a..b} \ lmeasurable" "{a<.. lmeasurable" - apply (metis box_real(2) lmeasurable_cbox) - by (metis box_real(1) lmeasurable_box) - -lemma integrable_const_ivl [iff]: - fixes a::"'a::ordered_euclidean_space" - shows "integrable (lebesgue_on {a..b}) (\x. c)" - by (metis cbox_interval finite_measure.integrable_const finite_measure_lebesgue_on lmeasurable_cbox) - -lemma absolutely_integrable_reflect[simp]: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - shows "(\x. f(-x)) absolutely_integrable_on cbox (-b) (-a) \ f absolutely_integrable_on cbox a b" - apply (auto simp: absolutely_integrable_on_def dest: integrable_reflect [THEN iffD1]) - unfolding integrable_on_def by auto - -lemma absolutely_integrable_reflect_real[simp]: - fixes f :: "real \ 'b::euclidean_space" - shows "(\x. f(-x)) absolutely_integrable_on {-b .. -a} \ f absolutely_integrable_on {a..b::real}" - unfolding box_real[symmetric] by (rule absolutely_integrable_reflect) - -lemma absolutely_integrable_on_subcbox: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - shows "\f absolutely_integrable_on S; cbox a b \ S\ \ f absolutely_integrable_on cbox a b" - by (meson absolutely_integrable_on_def integrable_on_subcbox) - -lemma absolutely_integrable_on_subinterval: - fixes f :: "real \ 'b::euclidean_space" - shows "\f absolutely_integrable_on S; {a..b} \ S\ \ f absolutely_integrable_on {a..b}" - using absolutely_integrable_on_subcbox by fastforce - -lemma integral_restrict: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "S \ T" "S \ sets lebesgue" "T \ sets lebesgue" - shows "integral\<^sup>L (lebesgue_on T) (\x. if x \ S then f x else 0) = integral\<^sup>L (lebesgue_on S) f" - using integral_restrict_Int [of S T f] assms - by (simp add: Int_absorb2) - -lemma absolutely_integrable_continuous_real: - fixes f :: "real \ 'b::euclidean_space" - shows "continuous_on {a..b} f \ f absolutely_integrable_on {a..b}" - by (metis absolutely_integrable_continuous box_real(2)) - -lemma absolutely_integrable_imp_borel_measurable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f absolutely_integrable_on S" "S \ sets lebesgue" - shows "f \ borel_measurable (lebesgue_on S)" - using absolutely_integrable_measurable assms by blast - -lemma measurable_bounded_by_integrable_imp_lebesgue_integrable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "integrable (lebesgue_on S) g" - and normf: "\x. x \ S \ norm(f x) \ g x" and S: "S \ sets lebesgue" - shows "integrable (lebesgue_on S) f" -proof - - have "f absolutely_integrable_on S" - by (metis (no_types) S absolutely_integrable_integrable_bound f g integrable_on_lebesgue_on measurable_bounded_by_integrable_imp_integrable normf) - then show ?thesis - by (simp add: S integrable_restrict_space set_integrable_def) -qed - -lemma lebesgue_on_mono: - assumes major: "AE x in lebesgue_on S. P x" and minor: "\x.\P x; x \ S\ \ Q x" - shows "AE x in lebesgue_on S. Q x" -proof - - have "AE a in lebesgue_on S. P a \ Q a" - using minor space_restrict_space by fastforce - then show ?thesis - using major by auto -qed - -lemma integral_eq_zero_null_sets: - assumes "S \ null_sets lebesgue" - shows "integral\<^sup>L (lebesgue_on S) f = 0" -proof (rule integral_eq_zero_AE) - show "AE x in lebesgue_on S. f x = 0" - by (metis (no_types, lifting) assms AE_not_in lebesgue_on_mono null_setsD2 null_sets_restrict_space order_refl) -qed - -lemma sin_zero_pi_iff: - fixes x::real - assumes "\x\ < pi" - shows "sin x = 0 \ x = 0" -proof - show "x = 0" if "sin x = 0" - using that assms by (auto simp: sin_zero_iff) -qed auto - -lemma id_borel_measurable_lebesgue [iff]: "id \ borel_measurable lebesgue" - by (simp add: measurable_completion) - -lemma id_borel_measurable_lebesgue_on [iff]: "id \ borel_measurable (lebesgue_on S)" - by (simp add: measurable_completion measurable_restrict_space1) - -lemma root_abs_power: "n > 0 \ abs (root n (y ^n)) = abs y" - using root_sgn_power [of n] - by (metis abs_ge_zero power_abs real_root_abs real_root_power_cancel) - -lemma power_tendsto_0_iff [simp]: - fixes f :: "'a \ real" - assumes "n > 0" - shows "((\x. f x ^ n) \ 0) F \ (f \ 0) F" -proof - - have "((\x. \root n (f x ^ n)\) \ 0) F \ (f \ 0) F" - by (auto simp: assms root_abs_power tendsto_rabs_zero_iff) - then have "((\x. f x ^ n) \ 0) F \ (f \ 0) F" - by (metis tendsto_real_root abs_0 real_root_zero tendsto_rabs) - with assms show ?thesis - by (auto simp: tendsto_null_power) -qed - -end - diff --git a/thys/Fourier/Fourier_Aux2.thy b/thys/Fourier/Fourier_Aux2.thy --- a/thys/Fourier/Fourier_Aux2.thy +++ b/thys/Fourier/Fourier_Aux2.thy @@ -1,2165 +1,139 @@ -section\Lemmas destined for the development version\ +section\Lemmas possibly destined for future Isabelle releases\ theory Fourier_Aux2 - imports Fourier_Aux1 Ergodic_Theory.SG_Library_Complement + imports Ergodic_Theory.SG_Library_Complement begin -lemma has_integral_divide: - fixes c :: "_ :: real_normed_div_algebra" - shows "(f has_integral y) S \ ((\x. f x / c) has_integral (y / c)) S" - unfolding divide_inverse - by (simp add: has_integral_mult_left) - -lemma integral_sin [simp]: - fixes a::real - assumes "a \ b" shows "integral {a..b} sin = cos a - cos b" -proof - - have "(sin has_integral (- cos b - - cos a)) {a..b}" - proof (rule fundamental_theorem_of_calculus) - show "((\a. - cos a) has_vector_derivative sin x) (at x within {a..b})" for x - unfolding has_field_derivative_iff_has_vector_derivative [symmetric] - by (rule derivative_eq_intros | force)+ - qed (use assms in auto) - then show ?thesis - by (simp add: integral_unique) -qed - -lemma integral_cos [simp]: - fixes a::real - assumes "a \ b" shows "integral {a..b} cos = sin b - sin a" -proof - - have "(cos has_integral (sin b - sin a)) {a..b}" - proof (rule fundamental_theorem_of_calculus) - show "(sin has_vector_derivative cos x) (at x within {a..b})" for x - unfolding has_field_derivative_iff_has_vector_derivative [symmetric] - by (rule derivative_eq_intros | force)+ - qed (use assms in auto) - then show ?thesis - by (simp add: integral_unique) -qed - lemma has_integral_sin_nx: "((\x. sin(real_of_int n * x)) has_integral 0) {-pi..pi}" proof (cases "n = 0") case False have "((\x. sin (n * x)) has_integral (- cos (n * pi)/n - - cos (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. - cos (n * x) / n) has_vector_derivative sin (n * a)) (at a within {-pi..pi})" if "a \ {-pi..pi}" for a :: real using that False apply (simp only: has_vector_derivative_def) apply (rule derivative_eq_intros | force)+ done qed auto then show ?thesis by simp qed auto lemma integral_sin_nx: "integral {-pi..pi} (\x. sin(x * real_of_int n)) = 0" using has_integral_sin_nx [of n] by (force simp: mult.commute) lemma integral_sin_Z [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(x * n)) = 0" proof (subst lebesgue_integral_eq_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. sin (x * n))" by (intro continuous_imp_integrable_real continuous_intros) show "integral {-pi..pi} (\x. sin (x * n)) = 0" using assms Ints_cases integral_sin_nx by blast qed auto lemma integral_sin_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. sin(n * x)) = 0" by (metis assms integral_sin_Z mult_commute_abs) lemma has_integral_cos_nx: "((\x. cos(real_of_int n * x)) has_integral (if n = 0 then 2 * pi else 0)) {-pi..pi}" proof (cases "n = 0") case True then show ?thesis using has_integral_const_real [of "1::real" "-pi" pi] by auto next case False have "((\x. cos (n * x)) has_integral (sin (n * pi)/n - sin (n * - pi)/n)) {-pi..pi}" proof (rule fundamental_theorem_of_calculus) show "((\x. sin (n * x) / n) has_vector_derivative cos (n * x)) (at x within {-pi..pi})" if "x \ {-pi..pi}" for x :: real using that False apply (simp only: has_vector_derivative_def) apply (rule derivative_eq_intros | force)+ done qed auto with False show ?thesis by (simp add: mult.commute) qed lemma integral_cos_nx: "integral {-pi..pi} (\x. cos(x * real_of_int n)) = (if n = 0 then 2 * pi else 0)" using has_integral_cos_nx [of n] by (force simp: mult.commute) lemma integral_cos_Z [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(x * n)) = (if n = 0 then 2 * pi else 0)" proof (subst lebesgue_integral_eq_integral) show "integrable (lebesgue_on {-pi..pi}) (\x. cos (x * n))" by (intro continuous_imp_integrable_real continuous_intros) show "integral {-pi..pi} (\x. cos (x * n)) = (if n = 0 then 2 * pi else 0)" by (metis Ints_cases assms integral_cos_nx of_int_0_eq_iff) qed auto lemma integral_cos_Z' [simp]: assumes "n \ \" shows "integral\<^sup>L (lebesgue_on {-pi..pi}) (\x. cos(n * x)) = (if n = 0 then 2 * pi else 0)" by (metis assms integral_cos_Z mult_commute_abs) -lemma real_tendsto_mult_left_iff: - "c \ 0 \ tendsto(\x. c * f x) (c * l) F \ tendsto f l F" for c :: real - by (auto simp: tendsto_mult_left dest: tendsto_mult_left [where c = "1/c"]) - -lemma real_tendsto_mult_right_iff: - "c \ 0 \ tendsto(\x. f x * c) (l * c) F \ tendsto f l F" for c :: real - by (simp add: mult.commute real_tendsto_mult_left_iff) - lemma real_tendsto_zero_mult_right_iff [simp]: fixes c::real assumes "c \ 0" shows "(\n. a n * c)\ 0 \ a \ 0" by (metis assms mult_zero_left real_tendsto_mult_right_iff) lemma real_tendsto_zero_divide_iff [simp]: fixes c::real assumes "c \ 0" shows "(\n. a n / c)\ 0 \ a \ 0" using real_tendsto_zero_mult_right_iff [of "1/c" a] assms by (simp add: field_simps) -lemma measurable_restrict_mono: - assumes f: "f \ restrict_space M A \\<^sub>M N" and "B \ A" - shows "f \ restrict_space M B \\<^sub>M N" -by (rule measurable_compose[OF measurable_restrict_space3 f]) - (insert \B \ A\, auto) - -lemma borel_measurable_if_lebesgue_on: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "S \ sets lebesgue" "T \ sets lebesgue" "S \ T" - shows "(\x. if x \ S then f x else 0) \ borel_measurable (lebesgue_on T) \ f \ borel_measurable (lebesgue_on S)" - (is "?lhs = ?rhs") -proof - assume ?lhs then show ?rhs - using measurable_restrict_mono [OF _ \S \ T\] - by (subst measurable_lebesgue_cong [where g = "(\x. if x \ S then f x else 0)"]) auto -next - assume ?rhs then show ?lhs - by (simp add: \S \ sets lebesgue\ borel_measurable_if_I measurable_restrict_space1) -qed - - lemma insert_sets_lebesgue_on: assumes "A \ sets (lebesgue_on S)" "a \ S" "S \ sets lebesgue" shows "insert a A \ sets (lebesgue_on S)" by (metis assms borel_singleton insert_subset lborelD sets.Int_space_eq2 sets.empty_sets sets.insert_in_sets sets_completionI_sets sets_restrict_space_iff) -lemma integrable_lebesgue_on_empty [iff]: - fixes f :: "'a::euclidean_space \ 'b::{second_countable_topology,banach}" - shows "integrable (lebesgue_on {}) f" - by (simp add: integrable_restrict_space) - -lemma integral_lebesgue_on_empty [simp]: - fixes f :: "'a::euclidean_space \ 'b::{second_countable_topology,banach}" - shows "integral\<^sup>L (lebesgue_on {}) f = 0" - by (simp add: Bochner_Integration.integral_empty) - lemma odd_even_cases [case_names 0 odd even]: assumes "P 0" and odd: "\n. P(Suc (2 * n))" and even: "\n. P(2 * n + 2)" shows "P n" by (metis nat_induct2 One_nat_def Suc_1 add_Suc_right assms(1) dvdE even odd oddE) -lemma null_AE_impI: "\N \ null_sets (lebesgue_on S); S \ sets lebesgue\ \ AE x in lebesgue_on S. x \ N \ P x" - by (rule AE_I' [of N]) auto - -lemma has_bochner_integral_restrict_space: - fixes f :: "'a \ 'b::{banach, second_countable_topology}" - assumes \: "\ \ space M \ sets M" - shows "has_bochner_integral (restrict_space M \) f i - \ has_bochner_integral M (\x. indicator \ x *\<^sub>R f x) i" - by (simp add: integrable_restrict_space [OF assms] integral_restrict_space [OF assms] has_bochner_integral_iff) - -lemma integrable_restrict_UNIV: - fixes f :: "'a::euclidean_space \ 'b::{banach, second_countable_topology}" - assumes S: "S \ sets lebesgue" - shows "integrable lebesgue (\x. if x \ S then f x else 0) \ integrable (lebesgue_on S) f" - using has_bochner_integral_restrict_space [of S lebesgue f] assms - by (simp add: integrable.simps indicator_scaleR_eq_if) - -lemma integral_mono_lebesgue_on_AE: - fixes f::"_ \ real" - assumes f: "integrable (lebesgue_on T) f" - and gf: "AE x in (lebesgue_on S). g x \ f x" - and f0: "AE x in (lebesgue_on T). 0 \ f x" - and "S \ T" and S: "S \ sets lebesgue" and T: "T \ sets lebesgue" - shows "(\x. g x \(lebesgue_on S)) \ (\x. f x \(lebesgue_on T))" -proof - - have "(\x. g x \(lebesgue_on S)) = (\x. (if x \ S then g x else 0) \lebesgue)" - by (simp add: integral_restrict_UNIV S) - also have "\ \ (\x. (if x \ T then f x else 0) \lebesgue)" - proof (rule Bochner_Integration.integral_mono_AE') - show "integrable lebesgue (\x. if x \ T then f x else 0)" - by (simp add: integrable_restrict_UNIV T f) - show "AE x in lebesgue. (if x \ S then g x else 0) \ (if x \ T then f x else 0)" - using assms by (auto simp: AE_restrict_space_iff) - show "AE x in lebesgue. 0 \ (if x \ T then f x else 0)" - using f0 by (simp add: AE_restrict_space_iff T) - qed - also have "\ = (\x. f x \(lebesgue_on T))" - using integral_restrict_UNIV T by blast - finally show ?thesis . -qed - - - -lemma integrable_subinterval_real: - fixes f :: "real \ 'a::euclidean_space" - assumes "integrable (lebesgue_on {a..b}) f" - and "{c..d} \ {a..b}" - shows "integrable (lebesgue_on {c..d}) f" - by (metis (full_types) absolutely_integrable_on_def assms fmeasurableD integrable_restrict_space integrable_subinterval_real interval_cbox lmeasurable_cbox set_integrable_def sets.Int_space_eq2) - - -lemma indefinite_integral_continuous_1: - fixes f :: "real \ 'a::euclidean_space" - assumes "integrable (lebesgue_on {a..b}) f" - shows "continuous_on {a..b} (\x. integral\<^sup>L (lebesgue_on {a..x}) f)" -proof - - have "f integrable_on {a..b}" - by (simp add: assms integrable_on_lebesgue_on) - then have "continuous_on {a..b} (\x. integral {a..x} f)" - using indefinite_integral_continuous_1 by blast - moreover have "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" if "a \ x" "x \ b" for x - proof - - have "{a..x} \ {a..b}" - using that by auto - then have "integrable (lebesgue_on {a..x}) f" - using integrable_subinterval_real assms by blast - then show "integral\<^sup>L (lebesgue_on {a..x}) f = integral {a..x} f" - by (smt integrable_restrict_UNIV integral_restrict_UNIV Henstock_Kurzweil_Integration.integral_restrict_UNIV atLeastAtMost_borel integral_lebesgue sets_completionI_sets sets_lborel) - qed - ultimately show ?thesis - by (metis (no_types, lifting) atLeastAtMost_iff continuous_on_cong) -qed - - -lemma AE_null_sets_lebesgue: - assumes "AE x in lebesgue. P x" - obtains N where "N \ null_sets lebesgue" "\x. x \ N \ P x" - using assms unfolding completion.AE_iff_null_sets by auto - - lemma measure_lebesgue_on_ivl [simp]: "\{a..b} \ S; S \ sets lebesgue\ \ measure (lebesgue_on S) {a..b} = content {a..b::real}" by (simp add: measure_restrict_space) - -subsection\Linear functions are (uniformly) continuous on any set\ - -lemma linear_lim_0: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "linear f" - shows "(f \ 0) (at 0)" -proof - - obtain B where "B > 0" "\x. norm (f x) \ B * norm x" - using assms linear_bounded_pos by blast - then have "\e x. \0 < e; x \ 0; norm x < e / B\ \ norm (f x) < e" - by (metis le_less_trans mult.commute pos_less_divide_eq) - then show ?thesis - unfolding Lim_at dist_norm - by (metis \0 < B\ diff_zero divide_pos_pos zero_less_norm_iff) -qed - -lemma linear_continuous_at: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "linear f" - shows "continuous (at a) f" - using assms has_derivative_continuous linear_imp_has_derivative by blast - - -lemma linear_continuous_within: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "linear f" - shows "continuous (at x within S) f" - by (simp add: assms differentiable_imp_continuous_within linear_imp_differentiable) - -lemma linear_continuous_on: - fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" - assumes "linear f" - shows "continuous_on S f" - by (simp add: continuous_at_imp_continuous_on linear_continuous_within assms) - -lemma Lim_linear: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and h :: "'b \ 'c::real_normed_vector" - assumes "(f \ l) F" "linear h" - shows "((\x. h(f x)) \ h l) F" -proof - - obtain B where B: "B > 0" "\x. norm (h x) \ B * norm x" - using linear_bounded_pos [OF \linear h\] by blast - show ?thesis - unfolding tendsto_iff - proof (intro allI impI) - show "\\<^sub>F x in F. dist (h (f x)) (h l) < e" if "e > 0" for e - proof - - have "\\<^sub>F x in F. dist (f x) l < e/B" - by (simp add: \0 < B\ assms(1) tendstoD that) - then show ?thesis - unfolding dist_norm - proof (rule eventually_mono) - show "norm (h (f x) - h l) < e" if "norm (f x - l) < e / B" for x - using that B - apply (simp add: divide_simps) - by (metis \linear h\ le_less_trans linear_diff mult.commute) - qed - qed - qed -qed - -lemma linear_continuous_compose: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" - assumes "continuous F f" "linear g" - shows "continuous F (\x. g(f x))" - using assms unfolding continuous_def by (rule Lim_linear) - -lemma linear_continuous_on_compose: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" and g :: "'b \ 'c::real_normed_vector" - assumes "continuous_on S f" "linear g" - shows "continuous_on S (\x. g(f x))" - using assms by (simp add: continuous_on_eq_continuous_within linear_continuous_compose) - - -subsection\Also bilinear functions, in composition form\ - - -lemma bilinear_continuous_compose: - fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" - assumes "continuous F f" "continuous F g" "bilinear h" - shows "continuous F (\x. h (f x) (g x))" - using assms Lim_bilinear bilinear_conv_bounded_bilinear - unfolding continuous_def - by blast - -lemma bilinear_continuous_on_compose: - fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::real_normed_vector" - and f :: "'d::t2_space \ 'a" - assumes "continuous_on S f" "continuous_on S g" "bilinear h" - shows "continuous_on S (\x. h (f x) (g x))" - using assms - by (simp add: continuous_on_eq_continuous_within bilinear_continuous_compose) - - - - - -subsection\Austin's Lemma\ - -lemma Austin_Lemma: - fixes \ :: "'a::euclidean_space set set" - assumes "finite \" and \: "\D. D \ \ \ \k a b. D = cbox a b \ (\i \ Basis. b\i - a\i = k)" - obtains \ where "\ \ \" "pairwise disjnt \" - "measure lebesgue (\\) \ measure lebesgue (\\) / 3 ^ (DIM('a))" - using assms -proof (induction "card \" arbitrary: \ thesis rule: less_induct) - case less - show ?case - proof (cases "\ = {}") - case True - then show thesis - using less by auto - next - case False - then have "Max (Sigma_Algebra.measure lebesgue ` \) \ Sigma_Algebra.measure lebesgue ` \" - using Max_in finite_imageI \finite \\ by blast - then obtain D where "D \ \" and "measure lebesgue D = Max (measure lebesgue ` \)" - by auto - then have D: "\C. C \ \ \ measure lebesgue C \ measure lebesgue D" - by (simp add: \finite \\) - let ?\ = "{C. C \ \ - {D} \ disjnt C D}" - obtain \' where \'sub: "\' \ ?\" and \'dis: "pairwise disjnt \'" - and \'m: "measure lebesgue (\\') \ measure lebesgue (\?\) / 3 ^ (DIM('a))" - proof (rule less.hyps) - have *: "?\ \ \" - using \D \ \\ by auto - then show "card ?\ < card \" "finite ?\" - by (auto simp: \finite \\ psubset_card_mono) - show "\k a b. D = cbox a b \ (\i\Basis. b \ i - a \ i = k)" if "D \ ?\" for D - using less.prems(3) that by auto - qed - then have [simp]: "\\' - D = \\'" - by (auto simp: disjnt_iff) - show ?thesis - proof (rule less.prems) - show "insert D \' \ \" - using \'sub \D \ \\ by blast - show "disjoint (insert D \')" - using \'dis \'sub by (fastforce simp add: pairwise_def disjnt_sym) - obtain a3 b3 where m3: "content (cbox a3 b3) = 3 ^ DIM('a) * measure lebesgue D" - and sub3: "\C. \C \ \; \ disjnt C D\ \ C \ cbox a3 b3" - proof - - obtain k a b where ab: "D = cbox a b" and k: "\i. i \ Basis \ b\i - a\i = k" - using less.prems \D \ \\ by blast - then have eqk: "\i. i \ Basis \ a \ i \ b \ i \ k \ 0" - by force - show thesis - proof - let ?a = "(a + b) /\<^sub>R 2 - (3/2) *\<^sub>R (b - a)" - let ?b = "(a + b) /\<^sub>R 2 + (3/2) *\<^sub>R (b - a)" - have eq: "(\i\Basis. b \ i * 3 - a \ i * 3) = (\i\Basis. b \ i - a \ i) * 3 ^ DIM('a)" - by (simp add: comm_monoid_mult_class.prod.distrib flip: left_diff_distrib inner_diff_left) - show "content (cbox ?a ?b) = 3 ^ DIM('a) * measure lebesgue D" - by (simp add: content_cbox_if box_eq_empty algebra_simps eq ab k) - show "C \ cbox ?a ?b" if "C \ \" and CD: "\ disjnt C D" for C - proof - - obtain k' a' b' where ab': "C = cbox a' b'" and k': "\i. i \ Basis \ b'\i - a'\i = k'" - using less.prems \C \ \\ by blast - then have eqk': "\i. i \ Basis \ a' \ i \ b' \ i \ k' \ 0" - by force - show ?thesis - proof (clarsimp simp add: disjoint_interval disjnt_def ab ab' not_less subset_box algebra_simps) - show "a \ i * 2 \ a' \ i + b \ i \ a \ i + b' \ i \ b \ i * 2" - if * [rule_format]: "\j\Basis. a' \ j \ b' \ j" and "i \ Basis" for i - proof - - have "a' \ i \ b' \ i \ a \ i \ b \ i \ a \ i \ b' \ i \ a' \ i \ b \ i" - using \i \ Basis\ CD by (simp_all add: disjoint_interval disjnt_def ab ab' not_less) - then show ?thesis - using D [OF \C \ \\] \i \ Basis\ - apply (simp add: ab ab' k k' eqk eqk' content_cbox_cases) - using k k' by fastforce - qed - qed - qed - qed - qed - have \lm: "\D. D \ \ \ D \ lmeasurable" - using less.prems(3) by blast - have "measure lebesgue (\\) \ measure lebesgue (cbox a3 b3 \ (\\ - cbox a3 b3))" - proof (rule measure_mono_fmeasurable) - show "\\ \ sets lebesgue" - using \lm \finite \\ by blast - show "cbox a3 b3 \ (\\ - cbox a3 b3) \ lmeasurable" - by (simp add: \lm fmeasurable.Un fmeasurable.finite_Union less.prems(2) subset_eq) - qed auto - also have "\ = content (cbox a3 b3) + measure lebesgue (\\ - cbox a3 b3)" - by (simp add: \lm fmeasurable.finite_Union less.prems(2) measure_Un2 subsetI) - also have "\ \ (measure lebesgue D + measure lebesgue (\\')) * 3 ^ DIM('a)" - proof - - have "(\\ - cbox a3 b3) \ \?\" - using sub3 by fastforce - then have "measure lebesgue (\\ - cbox a3 b3) \ measure lebesgue (\?\)" - proof (rule measure_mono_fmeasurable) - show "\ \ - cbox a3 b3 \ sets lebesgue" - by (simp add: \lm fmeasurableD less.prems(2) sets.Diff sets.finite_Union subsetI) - show "\ {C \ \ - {D}. disjnt C D} \ lmeasurable" - using \lm less.prems(2) by auto - qed - then have "measure lebesgue (\\ - cbox a3 b3) / 3 ^ DIM('a) \ measure lebesgue (\ \')" - using \'m by (simp add: divide_simps) - then show ?thesis - by (simp add: m3 field_simps) - qed - also have "\ \ measure lebesgue (\(insert D \')) * 3 ^ DIM('a)" - proof (simp add: \lm \D \ \\) - show "measure lebesgue D + measure lebesgue (\\') \ measure lebesgue (D \ \ \')" - proof (subst measure_Un2) - show "\ \' \ lmeasurable" - by (meson \lm \insert D \' \ \\ fmeasurable.finite_Union less.prems(2) finite_subset subset_eq subset_insertI) - show "measure lebesgue D + measure lebesgue (\ \') \ measure lebesgue D + measure lebesgue (\ \' - D)" - using \insert D \' \ \\ infinite_super less.prems(2) by force - qed (simp add: \lm \D \ \\) - qed - finally show "measure lebesgue (\\) / 3 ^ DIM('a) \ measure lebesgue (\(insert D \'))" - by (simp add: divide_simps) - qed - qed -qed - - -subsection\A differentiability-like property of the indefinite integral. \ - -proposition integrable_ccontinuous_explicit: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "\a b::'a. f integrable_on cbox a b" - obtains N where - "negligible N" - "\x e. \x \ N; 0 < e\ \ - \d>0. \h. 0 < h \ h < d \ - norm(integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e" -proof - - define BOX where "BOX \ \h. \x::'a. cbox x (x + h *\<^sub>R One)" - define BOX2 where "BOX2 \ \h. \x::'a. cbox (x - h *\<^sub>R One) (x + h *\<^sub>R One)" - define i where "i \ \h x. integral (BOX h x) f /\<^sub>R h ^ DIM('a)" - define \ where "\ \ \x r. \d>0. \h. 0 < h \ h < d \ r \ norm(i h x - f x)" - let ?N = "{x. \e>0. \ x e}" - have "\N. negligible N \ (\x e. x \ N \ 0 < e \ \ \ x e)" - proof (rule exI ; intro conjI allI impI) - let ?M = "\n. {x. \ x (inverse(real n + 1))}" - have "negligible ({x. \ x \} \ cbox a b)" - if "\ > 0" for a b \ - proof (cases "negligible(cbox a b)") - case True - then show ?thesis - by (simp add: negligible_Int) - next - case False - then have "box a b \ {}" - by (simp add: negligible_interval) - then have ab: "\i. i \ Basis \ a\i < b\i" - by (simp add: box_ne_empty) - show ?thesis - unfolding negligible_outer_le - proof (intro allI impI) - fix e::real - let ?ee = "(e * \) / 2 / 6 ^ (DIM('a))" - assume "e > 0" - then have gt0: "?ee > 0" - using \\ > 0\ by auto - have f': "f integrable_on cbox (a - One) (b + One)" - using assms by blast - obtain \ where "gauge \" - and \: "\p. \p tagged_partial_division_of (cbox (a - One) (b + One)); \ fine p\ - \ (\(x, k)\p. norm (content k *\<^sub>R f x - integral k f)) < ?ee" - using Henstock_lemma [OF f' gt0] that by auto - let ?E = "{x. x \ cbox a b \ \ x \}" - have "\h>0. BOX h x \ \ x \ - BOX h x \ cbox (a - One) (b + One) \ \ \ norm (i h x - f x)" - if "x \ cbox a b" "\ x \" for x - proof - - obtain d where "d > 0" and d: "ball x d \ \ x" - using gaugeD [OF \gauge \\, of x] openE by blast - then obtain h where "0 < h" "h < 1" and hless: "h < d / real DIM('a)" - and mule: "\ \ norm (i h x - f x)" - using \\ x \\ [unfolded \_def, rule_format, of "min 1 (d / DIM('a))"] - by auto - show ?thesis - proof (intro exI conjI) - show "0 < h" "\ \ norm (i h x - f x)" by fact+ - have "BOX h x \ ball x d" - proof (clarsimp simp: BOX_def mem_box dist_norm algebra_simps) - fix y - assume "\i\Basis. x \ i \ y \ i \ y \ i \ h + x \ i" - then have lt: "\(x - y) \ i\ < d / real DIM('a)" if "i \ Basis" for i - using hless that by (force simp: inner_diff_left) - have "norm (x - y) \ (\i\Basis. \(x - y) \ i\)" - using norm_le_l1 by blast - also have "\ < d" - using sum_bounded_above_strict [of Basis "\i. \(x - y) \ i\" "d / DIM('a)", OF lt] - by auto - finally show "norm (x - y) < d" . - qed - with d show "BOX h x \ \ x" - by blast - show "BOX h x \ cbox (a - One) (b + One)" - using that \h < 1\ - by (force simp: BOX_def mem_box algebra_simps intro: subset_box_imp) - qed - qed - then obtain \ where h0: "\x. x \ ?E \ \ x > 0" - and BOX_\: "\x. x \ ?E \ BOX (\ x) x \ \ x" - and "\x. x \ ?E \ BOX (\ x) x \ cbox (a - One) (b + One) \ \ \ norm (i (\ x) x - f x)" - by simp metis - then have BOX_cbox: "\x. x \ ?E \ BOX (\ x) x \ cbox (a - One) (b + One)" - and \_le: "\x. x \ ?E \ \ \ norm (i (\ x) x - f x)" - by blast+ - define \' where "\' \ \x. if x \ cbox a b \ \ x \ then ball x (\ x) else \ x" - have "gauge \'" - using \gauge \\ by (auto simp: h0 gauge_def \'_def) - obtain \ where "countable \" - and \: "\\ \ cbox a b" - "\K. K \ \ \ interior K \ {} \ (\c d. K = cbox c d)" - and Dcovered: "\K. K \ \ \ \x. x \ cbox a b \ \ x \ \ x \ K \ K \ \' x" - and subUD: "?E \ \\" - by (rule covering_lemma [of ?E a b \']) (simp_all add: Bex_def \box a b \ {}\ \gauge \'\) - then have "\ \ sets lebesgue" - by fastforce - show "\T. {x. \ x \} \ cbox a b \ T \ T \ lmeasurable \ measure lebesgue T \ e" - proof (intro exI conjI) - show "{x. \ x \} \ cbox a b \ \\" - apply auto - using subUD by auto - have mUE: "measure lebesgue (\ \) \ measure lebesgue (cbox a b)" - if "\ \ \" "finite \" for \ - proof (rule measure_mono_fmeasurable) - show "\ \ \ cbox a b" - using \(1) that(1) by blast - show "\ \ \ sets lebesgue" - by (metis \(2) fmeasurable.finite_Union fmeasurableD lmeasurable_cbox subset_eq that) - qed auto - then show "\\ \ lmeasurable" - by (metis \(2) \countable \\ fmeasurable_Union_bound lmeasurable_cbox) - then have leab: "measure lebesgue (\\) \ measure lebesgue (cbox a b)" - by (meson \(1) fmeasurableD lmeasurable_cbox measure_mono_fmeasurable) - obtain \ where "\ \ \" "finite \" - and \: "measure lebesgue (\\) \ 2 * measure lebesgue (\\)" - proof (cases "measure lebesgue (\\) = 0") - case True - then show ?thesis - by (force intro: that [where \ = "{}"]) - next - case False - obtain \ where "\ \ \" "finite \" - and \: "measure lebesgue (\\)/2 < measure lebesgue (\\)" - proof (rule measure_countable_Union_approachable [of \ "measure lebesgue (\\) / 2" "content (cbox a b)"]) - show "countable \" - by fact - show "0 < measure lebesgue (\ \) / 2" - using False by (simp add: zero_less_measure_iff) - show Dlm: "D \ lmeasurable" if "D \ \" for D - using \(2) that by blast - show "measure lebesgue (\ \) \ content (cbox a b)" - if "\ \ \" "finite \" for \ - proof - - have "measure lebesgue (\ \) \ measure lebesgue (\\)" - proof (rule measure_mono_fmeasurable) - show "\ \ \ \ \" - by (simp add: Sup_subset_mono \\ \ \\) - show "\ \ \ sets lebesgue" - by (meson Dlm fmeasurableD sets.finite_Union subset_eq that) - show "\ \ \ lmeasurable" - by fact - qed - also have "\ \ measure lebesgue (cbox a b)" - proof (rule measure_mono_fmeasurable) - show "\ \ \ sets lebesgue" - by (simp add: \\ \ \ lmeasurable\ fmeasurableD) - qed (auto simp:\(1)) - finally show ?thesis - by simp - qed - qed auto - then show ?thesis - using that by auto - qed - obtain tag where tag_in_E: "\D. D \ \ \ tag D \ ?E" - and tag_in_self: "\D. D \ \ \ tag D \ D" - and tag_sub: "\D. D \ \ \ D \ \' (tag D)" - using Dcovered by simp metis - then have sub_ball_tag: "\D. D \ \ \ D \ ball (tag D) (\ (tag D))" - by (simp add: \'_def) - define \ where "\ \ \D. BOX (\(tag D)) (tag D)" - define \2 where "\2 \ \D. BOX2 (\(tag D)) (tag D)" - obtain \ where "\ \ \2 ` \" "pairwise disjnt \" - "measure lebesgue (\\) \ measure lebesgue (\(\2`\)) / 3 ^ (DIM('a))" - proof (rule Austin_Lemma) - show "finite (\2`\)" - using \finite \\ by blast - have "\k a b. \2 D = cbox a b \ (\i\Basis. b \ i - a \ i = k)" if "D \ \" for D - apply (rule_tac x="2 * \(tag D)" in exI) - apply (rule_tac x="tag D - \(tag D) *\<^sub>R One" in exI) - apply (rule_tac x="tag D + \(tag D) *\<^sub>R One" in exI) - using that - apply (auto simp: \2_def BOX2_def algebra_simps) - done - then show "\D. D \ \2 ` \ \ \k a b. D = cbox a b \ (\i\Basis. b \ i - a \ i = k)" - by blast - qed auto - then obtain \ where "\ \ \" and disj: "pairwise disjnt (\2 ` \)" - and "measure lebesgue (\(\2 ` \)) \ measure lebesgue (\(\2`\)) / 3 ^ (DIM('a))" - unfolding \2_def subset_image_iff - by (meson empty_subsetI equals0D pairwise_imageI) - moreover - have "measure lebesgue (\(\2 ` \)) * 3 ^ DIM('a) \ e/2" - proof - - have "finite \" - using \finite \\ \\ \ \\ infinite_super by blast - have BOX2_m: "\x. x \ tag ` \ \ BOX2 (\ x) x \ lmeasurable" - by (auto simp: BOX2_def) - have BOX_m: "\x. x \ tag ` \ \ BOX (\ x) x \ lmeasurable" - by (auto simp: BOX_def) - have BOX_sub: "BOX (\ x) x \ BOX2 (\ x) x" for x - by (auto simp: BOX_def BOX2_def subset_box algebra_simps) - have DISJ2: "BOX2 (\ (tag X)) (tag X) \ BOX2 (\ (tag Y)) (tag Y) = {}" - if "X \ \" "Y \ \" "tag X \ tag Y" for X Y - proof - - obtain i where i: "i \ Basis" "tag X \ i \ tag Y \ i" - using \tag X \ tag Y\ by (auto simp: euclidean_eq_iff [of "tag X"]) - have XY: "X \ \" "Y \ \" - using \\ \ \\ \\ \ \\ that by auto - then have "0 \ \ (tag X)" "0 \ \ (tag Y)" - by (meson h0 le_cases not_le tag_in_E)+ - with XY i have "BOX2 (\ (tag X)) (tag X) \ BOX2 (\ (tag Y)) (tag Y)" - unfolding eq_iff - by (fastforce simp add: BOX2_def subset_box algebra_simps) - then show ?thesis - using disj that by (auto simp: pairwise_def disjnt_def \2_def) - qed - then have BOX2_disj: "pairwise (\x y. negligible (BOX2 (\ x) x \ BOX2 (\ y) y)) (tag ` \)" - by (simp add: pairwise_imageI) - then have BOX_disj: "pairwise (\x y. negligible (BOX (\ x) x \ BOX (\ y) y)) (tag ` \)" - proof (rule pairwise_mono) - show "negligible (BOX (\ x) x \ BOX (\ y) y)" - if "negligible (BOX2 (\ x) x \ BOX2 (\ y) y)" for x y - by (metis (no_types, hide_lams) that Int_mono negligible_subset BOX_sub) - qed auto - - have eq: "\box. (\D. box (\ (tag D)) (tag D)) ` \ = (\t. box (\ t) t) ` tag ` \" - by (simp add: image_comp) - have "measure lebesgue (BOX2 (\ t) t) * 3 ^ DIM('a) - = measure lebesgue (BOX (\ t) t) * (2*3) ^ DIM('a)" - if "t \ tag ` \" for t - proof - - have "content (cbox (t - \ t *\<^sub>R One) (t + \ t *\<^sub>R One)) - = content (cbox t (t + \ t *\<^sub>R One)) * 2 ^ DIM('a)" - using that by (simp add: algebra_simps content_cbox_if box_eq_empty) - then show ?thesis - by (simp add: BOX2_def BOX_def flip: power_mult_distrib) - qed - then have "measure lebesgue (\(\2 ` \)) * 3 ^ DIM('a) = measure lebesgue (\(\ ` \)) * 6 ^ DIM('a)" - unfolding \_def \2_def eq - by (simp add: measure_negligible_finite_Union_image - \finite \\ BOX2_m BOX_m BOX2_disj BOX_disj sum_distrib_right - del: UN_simps) - also have "\ \ e/2" - proof - - have "\ * measure lebesgue (\D\\. \ D) \ \ * (\D \ \`\. measure lebesgue D)" - using \\ > 0\ \finite \\ by (force simp: BOX_m \_def fmeasurableD intro: measure_Union_le) - also have "\ = (\D \ \`\. measure lebesgue D * \)" - by (metis mult.commute sum_distrib_right) - also have "\ \ (\(x, K) \ (\D. (tag D, \ D)) ` \. norm (content K *\<^sub>R f x - integral K f))" - proof (rule sum_le_included; clarify?) - fix D - assume "D \ \" - then have "\ (tag D) > 0" - using \\ \ \\ \\ \ \\ h0 tag_in_E by auto - then have m_\: "measure lebesgue (\ D) > 0" - by (simp add: \_def BOX_def algebra_simps) - have "\ \ norm (i (\(tag D)) (tag D) - f(tag D))" - using \_le \D \ \\ \\ \ \\ \\ \ \\ tag_in_E by auto - also have "\ = norm ((content (\ D) *\<^sub>R f(tag D) - integral (\ D) f) /\<^sub>R measure lebesgue (\ D))" - using m_\ - unfolding i_def \_def BOX_def - by (simp add: algebra_simps content_cbox_plus norm_minus_commute) - finally have "measure lebesgue (\ D) * \ \ norm (content (\ D) *\<^sub>R f(tag D) - integral (\ D) f)" - using m_\ by (simp add: field_simps) - then show "\y\(\D. (tag D, \ D)) ` \. - snd y = \ D \ measure lebesgue (\ D) * \ \ (case y of (x, k) \ norm (content k *\<^sub>R f x - integral k f))" - using \D \ \\ by auto - qed (use \finite \\ in auto) - also have "\ < ?ee" - proof (rule \) - show "(\D. (tag D, \ D)) ` \ tagged_partial_division_of cbox (a - One) (b + One)" - unfolding tagged_partial_division_of_def - proof (intro conjI allI impI ; clarify ?) - show "tag D \ \ D" - if "D \ \" for D - using that \\ \ \\ \\ \ \\ h0 tag_in_E - by (auto simp: \_def BOX_def mem_box algebra_simps eucl_less_le_not_le in_mono) - show "y \ cbox (a - One) (b + One)" if "D \ \" "y \ \ D" for D y - using that BOX_cbox \_def \\ \ \\ \\ \ \\ tag_in_E by blast - show "tag D = tag E \ \ D = \ E" - if "D \ \" "E \ \" and ne: "interior (\ D) \ interior (\ E) \ {}" for D E - proof - - have "BOX2 (\ (tag D)) (tag D) \ BOX2 (\ (tag E)) (tag E) = {} \ tag E = tag D" - using DISJ2 \D \ \\ \E \ \\ by force - then have "BOX (\ (tag D)) (tag D) \ BOX (\ (tag E)) (tag E) = {} \ tag E = tag D" - using BOX_sub by blast - then show "tag D = tag E \ \ D = \ E" - by (metis \_def interior_Int interior_empty ne) - qed - qed (use \finite \\ \_def BOX_def in auto) - show "\ fine (\D. (tag D, \ D)) ` \" - unfolding fine_def \_def using BOX_\ \\ \ \\ \\ \ \\ tag_in_E by blast - qed - finally show ?thesis - using \\ > 0\ by (auto simp: divide_simps) - qed - finally show ?thesis . - qed - moreover - have "measure lebesgue (\\) \ measure lebesgue (\(\2`\))" - proof (rule measure_mono_fmeasurable) - have "D \ ball (tag D) (\(tag D))" if "D \ \" for D - using \\ \ \\ sub_ball_tag that by blast - moreover have "ball (tag D) (\(tag D)) \ BOX2 (\ (tag D)) (tag D)" if "D \ \" for D - proof (clarsimp simp: \2_def BOX2_def mem_box algebra_simps dist_norm) - fix x and i::'a - assume "norm (tag D - x) < \ (tag D)" and "i \ Basis" - then have "\tag D \ i - x \ i\ \ \ (tag D)" - by (metis eucl_less_le_not_le inner_commute inner_diff_right norm_bound_Basis_le) - then show "tag D \ i \ x \ i + \ (tag D) \ x \ i \ \ (tag D) + tag D \ i" - by (simp add: abs_diff_le_iff) - qed - ultimately show "\\ \ \(\2`\)" - by (force simp: \2_def) - show "\\ \ sets lebesgue" - using \finite \\ \\ \ sets lebesgue\ \\ \ \\ by blast - show "\(\2`\) \ lmeasurable" - unfolding \2_def BOX2_def using \finite \\ by blast - qed - ultimately - have "measure lebesgue (\\) \ e/2" - by (auto simp: divide_simps) - then show "measure lebesgue (\\) \ e" - using \ by linarith - qed - qed - qed - then have "\j. negligible {x. \ x (inverse(real j + 1))}" - using negligible_on_intervals - by (metis (full_types) inverse_positive_iff_positive le_add_same_cancel1 linorder_not_le nat_le_real_less not_add_less1 of_nat_0) - then have "negligible ?M" - by auto - moreover have "?N \ ?M" - proof (clarsimp simp: dist_norm) - fix y e - assume "0 < e" - and ye [rule_format]: "\ y e" - then obtain k where k: "0 < k" "inverse (real k + 1) < e" - by (metis One_nat_def add.commute less_add_same_cancel2 less_imp_inverse_less less_trans neq0_conv of_nat_1 of_nat_Suc reals_Archimedean zero_less_one) - with ye show "\n. \ y (inverse (real n + 1))" - apply (rule_tac x=k in exI) - unfolding \_def - by (force intro: less_le_trans) - qed - ultimately show "negligible ?N" - by (blast intro: negligible_subset) - show "\ \ x e" if "x \ ?N \ 0 < e" for x e - using that by blast - qed - with that show ?thesis - unfolding i_def BOX_def \_def by (fastforce simp add: not_le) -qed - - - -subsection\HOL Light measurability\ - -lemma measurable_on_UNIV: - "(\x. if x \ S then f x else 0) measurable_on UNIV \ f measurable_on S" - by (auto simp: measurable_on_def) - -lemma measurable_on_spike_set: - assumes f: "f measurable_on S" and neg: "negligible ((S - T) \ (T - S))" - shows "f measurable_on T" -proof - - obtain N and F - where N: "negligible N" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ N \ (\n. F n x) \ (if x \ S then f x else 0)" - using f by (auto simp: measurable_on_def) - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "continuous_on UNIV (\x. F n x)" for n - by (intro conF continuous_intros) - show "negligible (N \ (S - T) \ (T - S))" - by (metis (full_types) N neg negligible_Un_eq) - show "(\n. F n x) \ (if x \ T then f x else 0)" - if "x \ (N \ (S - T) \ (T - S))" for x - using that tendsF [of x] by auto - qed -qed - -text\ Various common equivalent forms of function measurability. \ - -lemma measurable_on_0 [simp]: "(\x. 0) measurable_on S" - unfolding measurable_on_def -proof (intro exI conjI allI impI) - show "(\n. 0) \ (if x \ S then 0::'b else 0)" for x - by force -qed auto - -lemma measurable_on_scaleR_const: - assumes f: "f measurable_on S" - shows "(\x. c *\<^sub>R f x) measurable_on S" -proof - - obtain NF and F - where NF: "negligible NF" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ S then f x else 0)" - using f by (auto simp: measurable_on_def) - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "continuous_on UNIV (\x. c *\<^sub>R F n x)" for n - by (intro conF continuous_intros) - show "(\n. c *\<^sub>R F n x) \ (if x \ S then c *\<^sub>R f x else 0)" - if "x \ NF" for x - using tendsto_scaleR [OF tendsto_const tendsF, of x] that by auto - qed (auto simp: NF) -qed - - -lemma measurable_on_cmul: - fixes c :: real - assumes "f measurable_on S" - shows "(\x. c * f x) measurable_on S" - using measurable_on_scaleR_const [OF assms] by simp - -lemma measurable_on_cdivide: - fixes c :: real - assumes "f measurable_on S" - shows "(\x. f x / c) measurable_on S" -proof (cases "c=0") - case False - then show ?thesis - using measurable_on_cmul [of f S "1/c"] - by (simp add: assms) -qed auto - - -lemma measurable_on_minus: - "f measurable_on S \ (\x. -(f x)) measurable_on S" - using measurable_on_scaleR_const [of f S "-1"] by auto - - -lemma continuous_imp_measurable_on: - "continuous_on UNIV f \ f measurable_on UNIV" - unfolding measurable_on_def - apply (rule_tac x="{}" in exI) - apply (rule_tac x="\n. f" in exI, auto) - done - -proposition integrable_subintervals_imp_measurable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "\a b. f integrable_on cbox a b" - shows "f measurable_on UNIV" -proof - - define BOX where "BOX \ \h. \x::'a. cbox x (x + h *\<^sub>R One)" - define i where "i \ \h x. integral (BOX h x) f /\<^sub>R h ^ DIM('a)" - obtain N where "negligible N" - and k: "\x e. \x \ N; 0 < e\ - \ \d>0. \h. 0 < h \ h < d \ - norm (integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e" - using integrable_ccontinuous_explicit assms by blast - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "continuous_on UNIV ((\n x. i (inverse(Suc n)) x) n)" for n - proof (clarsimp simp: continuous_on_iff) - show "\d>0. \x'. dist x' x < d \ dist (i (inverse (1 + real n)) x') (i (inverse (1 + real n)) x) < e" - if "0 < e" - for x e - proof - - let ?e = "e / (1 + real n) ^ DIM('a)" - have "?e > 0" - using \e > 0\ by auto - moreover have "x \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" - by (simp add: mem_box inner_diff_left inner_left_distrib) - moreover have "x + One /\<^sub>R real (Suc n) \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" - by (auto simp: mem_box inner_diff_left inner_left_distrib field_simps) - ultimately obtain \ where "\ > 0" - and \: "\c' d'. \c' \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One); - d' \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One); - norm(c' - x) \ \; norm(d' - (x + One /\<^sub>R Suc n)) \ \\ - \ norm(integral(cbox c' d') f - integral(cbox x (x + One /\<^sub>R Suc n)) f) < ?e" - by (blast intro: indefinite_integral_continuous [of f _ _ x] assms) - show ?thesis - proof (intro exI impI conjI allI) - show "min \ 1 > 0" - using \\ > 0\ by auto - show "dist (i (inverse (1 + real n)) y) (i (inverse (1 + real n)) x) < e" - if "dist y x < min \ 1" for y - proof - - have no: "norm (y - x) < 1" - using that by (auto simp: dist_norm) - have le1: "inverse (1 + real n) \ 1" - by (auto simp: divide_simps) - have "norm (integral (cbox y (y + One /\<^sub>R real (Suc n))) f - - integral (cbox x (x + One /\<^sub>R real (Suc n))) f) - < e / (1 + real n) ^ DIM('a)" - proof (rule \) - show "y \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" - using no by (auto simp: mem_box algebra_simps dest: Basis_le_norm [of _ "y-x"]) - show "y + One /\<^sub>R real (Suc n) \ cbox (x - 2 *\<^sub>R One) (x + 2 *\<^sub>R One)" - proof (simp add: dist_norm mem_box algebra_simps, intro ballI conjI) - fix i::'a - assume "i \ Basis" - then have 1: "\y \ i - x \ i\ < 1" - by (metis inner_commute inner_diff_right no norm_bound_Basis_lt) - moreover have "\ < (2 + inverse (1 + real n))" "1 \ 2 - inverse (1 + real n)" - by (auto simp: field_simps) - ultimately show "x \ i \ y \ i + (2 + inverse (1 + real n))" - "y \ i + inverse (1 + real n) \ x \ i + 2" - by linarith+ - qed - show "norm (y - x) \ \" "norm (y + One /\<^sub>R real (Suc n) - (x + One /\<^sub>R real (Suc n))) \ \" - using that by (auto simp: dist_norm) - qed - then show ?thesis - using that by (simp add: dist_norm i_def BOX_def field_simps flip: scaleR_diff_right) - qed - qed - qed - qed - show "negligible N" - by (simp add: \negligible N\) - show "(\n. i (inverse (Suc n)) x) \ (if x \ UNIV then f x else 0)" - if "x \ N" for x - unfolding lim_sequentially - proof clarsimp - show "\no. \n\no. dist (i (inverse (1 + real n)) x) (f x) < e" - if "0 < e" for e - proof - - obtain d where "d > 0" - and d: "\h. \0 < h; h < d\ \ - norm (integral (cbox x (x + h *\<^sub>R One)) f /\<^sub>R h ^ DIM('a) - f x) < e" - using k [of x e] \x \ N\ \0 < e\ by blast - then obtain M where M: "M \ 0" "0 < inverse (real M)" "inverse (real M) < d" - using real_arch_invD by auto - show ?thesis - proof (intro exI allI impI) - show "dist (i (inverse (1 + real n)) x) (f x) < e" - if "M \ n" for n - proof - - have *: "0 < inverse (1 + real n)" "inverse (1 + real n) \ inverse M" - using that \M \ 0\ by auto - show ?thesis - using that M - apply (simp add: i_def BOX_def dist_norm) - apply (blast intro: le_less_trans * d) - done - qed - qed - qed - qed - qed -qed - -subsection\Composing continuous and measurable functions; a few variants\ - -lemma measurable_on_compose_continuous: - assumes f: "f measurable_on UNIV" and g: "continuous_on UNIV g" - shows "(g \ f) measurable_on UNIV" -proof - - obtain N and F - where "negligible N" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ N \ (\n. F n x) \ f x" - using f by (auto simp: measurable_on_def) - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "negligible N" - by fact - show "continuous_on UNIV (g \ (F n))" for n - using conF continuous_on_compose continuous_on_subset g by blast - show "(\n. (g \ F n) x) \ (if x \ UNIV then (g \ f) x else 0)" - if "x \ N" for x :: 'a - using that g tendsF by (auto simp: continuous_on_def intro: tendsto_compose) - qed -qed - -lemma measurable_on_compose_continuous_0: - assumes f: "f measurable_on S" and g: "continuous_on UNIV g" and "g 0 = 0" - shows "(g \ f) measurable_on S" -proof - - have f': "(\x. if x \ S then f x else 0) measurable_on UNIV" - using f measurable_on_UNIV by blast - show ?thesis - using measurable_on_compose_continuous [OF f' g] - by (simp add: measurable_on_UNIV o_def if_distrib \g 0 = 0\ cong: if_cong) -qed - - -lemma measurable_on_compose_continuous_box: - assumes fm: "f measurable_on UNIV" and fab: "\x. f x \ box a b" - and contg: "continuous_on (box a b) g" - shows "(g \ f) measurable_on UNIV" -proof - - have "\\. (\n. continuous_on UNIV (\ n)) \ (\x. x \ N \ (\n. \ n x) \ g (f x))" - if "negligible N" - and conth [rule_format]: "\n. continuous_on UNIV (\x. h n x)" - and tends [rule_format]: "\x. x \ N \ (\n. h n x) \ f x" - for N and h :: "nat \ 'a \ 'b" - proof - - define \ where "\ \ \n x. (\i\Basis. (max (a\i + (b\i - a\i) / real (n+2)) - (min ((h n x)\i) - (b\i - (b\i - a\i) / real (n+2)))) *\<^sub>R i)" - have aibi: "\i. i \ Basis \ a \ i < b \ i" - using box_ne_empty(2) fab by auto - then have *: "\i n. i \ Basis \ a \ i + real n * (a \ i) < b \ i + real n * (b \ i)" - by (meson add_mono_thms_linordered_field(3) less_eq_real_def mult_left_mono of_nat_0_le_iff) - show ?thesis - proof (intro exI conjI allI impI) - show "continuous_on UNIV (g \ (\ n))" for n :: nat - unfolding \_def - apply (intro continuous_on_compose2 [OF contg] continuous_intros conth) - apply (auto simp: aibi * mem_box less_max_iff_disj min_less_iff_disj algebra_simps divide_simps) - done - show "(\n. (g \ \ n) x) \ g (f x)" - if "x \ N" for x - unfolding o_def - proof (rule isCont_tendsto_compose [where g=g]) - show "isCont g (f x)" - using contg fab continuous_on_eq_continuous_at by blast - have "(\n. \ n x) \ (\i\Basis. max (a \ i) (min (f x \ i) (b \ i)) *\<^sub>R i)" - unfolding \_def - proof (intro tendsto_intros \x \ N\ tends) - fix i::'b - assume "i \ Basis" - have a: "(\n. a \ i + (b \ i - a \ i) / real n) \ a\i + 0" - by (intro tendsto_add lim_const_over_n tendsto_const) - show "(\n. a \ i + (b \ i - a \ i) / real (n + 2)) \ a \ i" - using LIMSEQ_ignore_initial_segment [where k=2, OF a] by simp - have b: "(\n. b\i - (b \ i - a \ i) / (real n)) \ b\i - 0" - by (intro tendsto_diff lim_const_over_n tendsto_const) - show "(\n. b \ i - (b \ i - a \ i) / real (n + 2)) \ b \ i" - using LIMSEQ_ignore_initial_segment [where k=2, OF b] by simp - qed - also have "(\i\Basis. max (a \ i) (min (f x \ i) (b \ i)) *\<^sub>R i) = (\i\Basis. (f x \ i) *\<^sub>R i)" - apply (rule sum.cong) - using fab - apply auto - apply (intro order_antisym) - apply (auto simp: mem_box) - using less_imp_le apply blast - by (metis (full_types) linear max_less_iff_conj min.bounded_iff not_le) - also have "\ = f x" - using euclidean_representation by blast - finally show "(\n. \ n x) \ f x" . - qed - qed - qed - then show ?thesis - using fm by (auto simp: measurable_on_def) -qed - -lemma measurable_on_Pair: - assumes f: "f measurable_on S" and g: "g measurable_on S" - shows "(\x. (f x, g x)) measurable_on S" -proof - - obtain NF and F - where NF: "negligible NF" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ S then f x else 0)" - using f by (auto simp: measurable_on_def) - obtain NG and G - where NG: "negligible NG" - and conG: "\n. continuous_on UNIV (G n)" - and tendsG: "\x. x \ NG \ (\n. G n x) \ (if x \ S then g x else 0)" - using g by (auto simp: measurable_on_def) - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "negligible (NF \ NG)" - by (simp add: NF NG) - show "continuous_on UNIV (\x. (F n x, G n x))" for n - using conF conG continuous_on_Pair by blast - show "(\n. (F n x, G n x)) \ (if x \ S then (f x, g x) else 0)" - if "x \ NF \ NG" for x - using tendsto_Pair [OF tendsF tendsG, of x x] that unfolding zero_prod_def - by (simp add: split: if_split_asm) - qed -qed - -lemma measurable_on_combine: - assumes f: "f measurable_on S" and g: "g measurable_on S" - and h: "continuous_on UNIV (\x. h (fst x) (snd x))" and "h 0 0 = 0" - shows "(\x. h (f x) (g x)) measurable_on S" -proof - - have *: "(\x. h (f x) (g x)) = (\x. h (fst x) (snd x)) \ (\x. (f x, g x))" - by auto - show ?thesis - unfolding * by (auto simp: measurable_on_compose_continuous_0 measurable_on_Pair assms) -qed - -lemma measurable_on_add: - assumes f: "f measurable_on S" and g: "g measurable_on S" - shows "(\x. f x + g x) measurable_on S" - by (intro continuous_intros measurable_on_combine [OF assms]) auto - -lemma measurable_on_diff: - assumes f: "f measurable_on S" and g: "g measurable_on S" - shows "(\x. f x - g x) measurable_on S" - by (intro continuous_intros measurable_on_combine [OF assms]) auto - -lemma measurable_on_scaleR: - assumes f: "f measurable_on S" and g: "g measurable_on S" - shows "(\x. f x *\<^sub>R g x) measurable_on S" - by (intro continuous_intros measurable_on_combine [OF assms]) auto - -lemma measurable_on_sum: - assumes "finite I" "\i. i \ I \ f i measurable_on S" - shows "(\x. sum (\i. f i x) I) measurable_on S" - using assms by (induction I) (auto simp: measurable_on_add) - -lemma measurable_on_spike: - assumes f: "f measurable_on T" and "negligible S" and gf: "\x. x \ T - S \ g x = f x" - shows "g measurable_on T" -proof - - obtain NF and F - where NF: "negligible NF" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ T then f x else 0)" - using f by (auto simp: measurable_on_def) - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "negligible (NF \ S)" - by (simp add: NF \negligible S\) - show "\x. x \ NF \ S \ (\n. F n x) \ (if x \ T then g x else 0)" - by (metis (full_types) Diff_iff Un_iff gf tendsF) - qed (auto simp: conF) -qed - - -lemma measurable_on_preimage_lemma0: - fixes f :: "'a::euclidean_space \ real" - assumes "m \ \" and f: "m / 2^n \ (f x)" "(f x) < (m+1) / 2^n" and m: "\m\ \ 2^(2 * n)" - shows "(\k\{k \ \. \k\ \ 2^(2 * n)}. - (k / 2^n) * indicator {y. k / 2^n \ f y \ f y < (k+1) / 2^n} x) - = (m / 2^n)" (is "?lhs = ?rhs") -proof - - have "?lhs = (\k\{m}. (k / 2^n) * indicator {y. k / 2^n \ f y \ f y < (k+1) / 2^n} x)" - proof (intro sum.mono_neutral_right ballI) - show "finite {k::real. k \ \ \ \k\ \ 2^(2 * n)}" - using finite_abs_int_segment by blast - show "(i / 2^n) * indicat_real {y. i / 2^n \ f y \ f y < (i+1) / 2^n} x = 0" - if "i \ {N \ \. \N\ \ 2^(2 * n)} - {m}" for i - using f m \m \ \\ that Ints_eq_abs_less1 [of i m] - by (auto simp: indicator_def divide_simps) - qed (auto simp: assms) - also have "\ = ?rhs" - using assms by (auto simp: indicator_def) - finally show ?thesis . -qed - -(*see HOL Light's lebesgue_measurable BUT OUR lmeasurable IS NOT THE SAME. It's more like "sets lebesgue" - `lebesgue_measurable s <=> (indicator s) measurable_on (:real^N)`;; -*) - -proposition indicator_measurable_on: - assumes "S \ sets lebesgue" - shows "indicat_real S measurable_on UNIV" -proof - - { fix n::nat - let ?\ = "(1::real) / (2 * 2^n)" - have \: "?\ > 0" - by auto - obtain T where "closed T" "T \ S" "S-T \ lmeasurable" and ST: "emeasure lebesgue (S - T) < ?\" - by (meson \ assms sets_lebesgue_inner_closed) - obtain U where "open U" "S \ U" "(U - S) \ lmeasurable" and US: "emeasure lebesgue (U - S) < ?\" - by (meson \ assms sets_lebesgue_outer_open) - have eq: "-T \ U = (S-T) \ (U - S)" - using \T \ S\ \S \ U\ by auto - have "emeasure lebesgue ((S-T) \ (U - S)) \ emeasure lebesgue (S - T) + emeasure lebesgue (U - S)" - using \S - T \ lmeasurable\ \U - S \ lmeasurable\ emeasure_subadditive by blast - also have "\ < ?\ + ?\" - using ST US add_mono_ennreal by metis - finally have le: "emeasure lebesgue (-T \ U) < ennreal (1 / 2^n)" - by (simp add: eq) - have 1: "continuous_on (T \ -U) (indicat_real S)" - unfolding indicator_def - proof (rule continuous_on_cases [OF \closed T\]) - show "closed (- U)" - using \open U\ by blast - show "continuous_on T (\x. 1::real)" "continuous_on (- U) (\x. 0::real)" - by (auto simp: continuous_on) - show "\x. x \ T \ x \ S \ x \ - U \ x \ S \ (1::real) = 0" - using \T \ S\ \S \ U\ by auto - qed - have 2: "closedin (top_of_set UNIV) (T \ -U)" - using \closed T\ \open U\ by auto - obtain g where "continuous_on UNIV g" "\x. x \ T \ -U \ g x = indicat_real S x" "\x. norm(g x) \ 1" - by (rule Tietze [OF 1 2, of 1]) auto - with le have "\g E. continuous_on UNIV g \ (\x \ -E. g x = indicat_real S x) \ - (\x. norm(g x) \ 1) \ E \ sets lebesgue \ emeasure lebesgue E < ennreal (1 / 2^n)" - apply (rule_tac x=g in exI) - apply (rule_tac x="-T \ U" in exI) - using \S - T \ lmeasurable\ \U - S \ lmeasurable\ eq by auto - } - then obtain g E where cont: "\n. continuous_on UNIV (g n)" - and geq: "\n x. x \ - E n \ g n x = indicat_real S x" - and ng1: "\n x. norm(g n x) \ 1" - and Eset: "\n. E n \ sets lebesgue" - and Em: "\n. emeasure lebesgue (E n) < ennreal (1 / 2^n)" - by metis - have null: "limsup E \ null_sets lebesgue" - proof (rule borel_cantelli_limsup1 [OF Eset]) - show "emeasure lebesgue (E n) < \" for n - by (metis Em infinity_ennreal_def order.asym top.not_eq_extremum) - show "summable (\n. measure lebesgue (E n))" - proof (rule summable_comparison_test' [OF summable_geometric, of "1/2" 0]) - show "norm (measure lebesgue (E n)) \ (1/2) ^ n" for n - using Em [of n] by (simp add: measure_def enn2real_leI power_one_over) - qed auto - qed - have tends: "(\n. g n x) \ indicat_real S x" if "x \ limsup E" for x - proof - - have "\\<^sub>F n in sequentially. x \ - E n" - using that by (simp add: mem_limsup_iff not_frequently) - then show ?thesis - unfolding tendsto_iff dist_real_def - by (simp add: eventually_mono geq) - qed - show ?thesis - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "negligible (limsup E)" - using negligible_iff_null_sets null by blast - show "continuous_on UNIV (g n)" for n - using cont by blast - qed (use tends in auto) -qed - -lemma measurable_on_restrict: - assumes f: "f measurable_on UNIV" and S: "S \ sets lebesgue" - shows "(\x. if x \ S then f x else 0) measurable_on UNIV" -proof - - have "indicat_real S measurable_on UNIV" - by (simp add: S indicator_measurable_on) - then show ?thesis - using measurable_on_scaleR [OF _ f, of "indicat_real S"] - by (simp add: indicator_scaleR_eq_if) -qed - -lemma measurable_on_const_UNIV: "(\x. k) measurable_on UNIV" - by (simp add: continuous_imp_measurable_on) - -lemma measurable_on_const [simp]: "S \ sets lebesgue \ (\x. k) measurable_on S" - using measurable_on_UNIV measurable_on_const_UNIV measurable_on_restrict by blast - -lemma simple_function_indicator_representation_real: - fixes f ::"'a \ real" - assumes f: "simple_function M f" and x: "x \ space M" and nn: "\x. f x \ 0" - shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" -proof - - have f': "simple_function M (ennreal \ f)" - by (simp add: f) - have *: "f x = - enn2real - (\y\ ennreal ` f ` space M. - y * indicator ((ennreal \ f) -` {y} \ space M) x)" - using arg_cong [OF simple_function_indicator_representation [OF f' x], of enn2real, simplified nn o_def] nn - unfolding o_def image_comp - by (metis enn2real_ennreal) - have "enn2real (\y\ennreal ` f ` space M. if ennreal (f x) = y \ x \ space M then y else 0) - = sum (enn2real \ (\y. if ennreal (f x) = y \ x \ space M then y else 0)) - (ennreal ` f ` space M)" - by (rule enn2real_sum) auto - also have "\ = sum (enn2real \ (\y. if ennreal (f x) = y \ x \ space M then y else 0) \ ennreal) - (f ` space M)" - by (rule sum.reindex) (use nn in \auto simp: inj_on_def intro: sum.cong\) - also have "\ = (\y\f ` space M. if f x = y \ x \ space M then y else 0)" - using nn - by (auto simp: inj_on_def intro: sum.cong) - finally show ?thesis - by (subst *) (simp add: enn2real_sum indicator_def if_distrib cong: if_cong) -qed - -lemma\<^marker>\tag important\ simple_function_induct_real - [consumes 1, case_names cong set mult add, induct set: simple_function]: - fixes u :: "'a \ real" - assumes u: "simple_function M u" - assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" - assumes set: "\A. A \ sets M \ P (indicator A)" - assumes mult: "\u c. P u \ P (\x. c * u x)" - assumes add: "\u v. P u \ P v \ P (\x. u x + v x)" - and nn: "\x. u x \ 0" - shows "P u" -proof (rule cong) - from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" - proof eventually_elim - fix x assume x: "x \ space M" - from simple_function_indicator_representation_real[OF u x] nn - show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" - by metis - qed -next - from u have "finite (u ` space M)" - unfolding simple_function_def by auto - then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" - proof induct - case empty - then show ?case - using set[of "{}"] by (simp add: indicator_def[abs_def]) - next - case (insert a F) - have eq: "\ {y. u x = y \ (y = a \ y \ F) \ x \ space M} - = (if u x = a \ x \ space M then a else 0) + \ {y. u x = y \ y \ F \ x \ space M}" for x - proof (cases "x \ space M") - case True - have *: "{y. u x = y \ (y = a \ y \ F)} = {y. u x = a \ y = a} \ {y. u x = y \ y \ F}" - by auto - show ?thesis - using insert by (simp add: * True) - qed auto - have a: "P (\x. a * indicator (u -` {a} \ space M) x)" - proof (intro mult set) - show "u -` {a} \ space M \ sets M" - using u by auto - qed - show ?case - using nn insert a - by (simp add: eq indicator_times_eq_if [where f = "\x. a"] add) - qed -next - show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" - apply (subst simple_function_cong) - apply (rule simple_function_indicator_representation_real[symmetric]) - apply (auto intro: u nn) - done -qed fact - -proposition simple_function_measurable_on_UNIV: - fixes f :: "'a::euclidean_space \ real" - assumes f: "simple_function lebesgue f" and nn: "\x. f x \ 0" - shows "f measurable_on UNIV" - using f -proof (induction f) - case (cong f g) - then obtain N where "negligible N" "{x. g x \ f x} \ N" - by (auto simp: eventually_ae_filter_negligible eq_commute) - then show ?case - by (blast intro: measurable_on_spike cong) -next - case (set S) - then show ?case - by (simp add: indicator_measurable_on) -next - case (mult u c) - then show ?case - by (simp add: measurable_on_cmul) - case (add u v) - then show ?case - by (simp add: measurable_on_add) -qed (auto simp: nn) - -lemma simple_function_lebesgue_if: - fixes f :: "'a::euclidean_space \ real" - assumes f: "simple_function lebesgue f" and S: "S \ sets lebesgue" - shows "simple_function lebesgue (\x. if x \ S then f x else 0)" -proof - - have ffin: "finite (range f)" and fsets: "\x. f -` {f x} \ sets lebesgue" - using f by (auto simp: simple_function_def) - have "finite (f ` S)" - by (meson finite_subset subset_image_iff ffin top_greatest) - moreover have "finite ((\x. 0::real) ` T)" for T :: "'a set" - by (auto simp: image_def) - moreover have if_sets: "(\x. if x \ S then f x else 0) -` {f a} \ sets lebesgue" for a - proof - - have *: "(\x. if x \ S then f x else 0) -` {f a} - = (if f a = 0 then -S \ f -` {f a} else (f -` {f a}) \ S)" - by (auto simp: split: if_split_asm) - show ?thesis - unfolding * by (metis Compl_in_sets_lebesgue S sets.Int sets.Un fsets) - qed - moreover have "(\x. if x \ S then f x else 0) -` {0} \ sets lebesgue" - proof (cases "0 \ range f") - case True - then show ?thesis - by (metis (no_types, lifting) if_sets rangeE) - next - case False - then have "(\x. if x \ S then f x else 0) -` {0} = -S" - by auto - then show ?thesis - by (simp add: Compl_in_sets_lebesgue S) - qed - ultimately show ?thesis - by (auto simp: simple_function_def) -qed - -corollary simple_function_measurable_on: - fixes f :: "'a::euclidean_space \ real" - assumes f: "simple_function lebesgue f" and nn: "\x. f x \ 0" and S: "S \ sets lebesgue" - shows "f measurable_on S" - by (simp add: measurable_on_UNIV [symmetric, of f] S f simple_function_lebesgue_if nn simple_function_measurable_on_UNIV) - -lemma - fixes f :: "'a::euclidean_space \ 'b::ordered_euclidean_space" - assumes f: "f measurable_on S" and g: "g measurable_on S" - shows measurable_on_sup: "(\x. f x \ g x) measurable_on S" - and measurable_on_inf: "(\x. f x \ g x) measurable_on S" -proof - - obtain NF and F - where NF: "negligible NF" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ NF \ (\n. F n x) \ (if x \ S then f x else 0)" - using f by (auto simp: measurable_on_def) - obtain NG and G - where NG: "negligible NG" - and conG: "\n. continuous_on UNIV (G n)" - and tendsG: "\x. x \ NG \ (\n. G n x) \ (if x \ S then g x else 0)" - using g by (auto simp: measurable_on_def) - show "(\x. f x \ g x) measurable_on S" - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "continuous_on UNIV (\x. F n x \ G n x)" for n - unfolding sup_max eucl_sup by (intro conF conG continuous_intros) - show "(\n. F n x \ G n x) \ (if x \ S then f x \ g x else 0)" - if "x \ NF \ NG" for x - using tendsto_sup [OF tendsF tendsG, of x x] that by auto - qed (simp add: NF NG) - show "(\x. f x \ g x) measurable_on S" - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "continuous_on UNIV (\x. F n x \ G n x)" for n - unfolding inf_min eucl_inf by (intro conF conG continuous_intros) - show "(\n. F n x \ G n x) \ (if x \ S then f x \ g x else 0)" - if "x \ NF \ NG" for x - using tendsto_inf [OF tendsF tendsG, of x x] that by auto - qed (simp add: NF NG) -qed - -proposition measurable_on_componentwise_UNIV: - "f measurable_on UNIV \ (\i\Basis. (\x. (f x \ i) *\<^sub>R i) measurable_on UNIV)" - (is "?lhs = ?rhs") -proof - assume L: ?lhs - show ?rhs - proof - fix i::'b - assume "i \ Basis" - have cont: "continuous_on UNIV (\x. (x \ i) *\<^sub>R i)" - by (intro continuous_intros) - show "(\x. (f x \ i) *\<^sub>R i) measurable_on UNIV" - using measurable_on_compose_continuous [OF L cont] - by (simp add: o_def) - qed -next - assume ?rhs - then have "\N g. negligible N \ - (\n. continuous_on UNIV (g n)) \ - (\x. x \ N \ (\n. g n x) \ (f x \ i) *\<^sub>R i)" - if "i \ Basis" for i - by (simp add: measurable_on_def that) - then obtain N g where N: "\i. i \ Basis \ negligible (N i)" - and cont: "\i n. i \ Basis \ continuous_on UNIV (g i n)" - and tends: "\i x. \i \ Basis; x \ N i\ \ (\n. g i n x) \ (f x \ i) *\<^sub>R i" - by metis - show ?lhs - unfolding measurable_on_def - proof (intro exI conjI allI impI) - show "negligible (\i \ Basis. N i)" - using N eucl.finite_Basis by blast - show "continuous_on UNIV (\x. (\i\Basis. g i n x))" for n - by (intro continuous_intros cont) - next - fix x - assume "x \ (\i \ Basis. N i)" - then have "\i. i \ Basis \ x \ N i" - by auto - then have "(\n. (\i\Basis. g i n x)) \ (\i\Basis. (f x \ i) *\<^sub>R i)" - by (intro tends tendsto_intros) - then show "(\n. (\i\Basis. g i n x)) \ (if x \ UNIV then f x else 0)" - by (simp add: euclidean_representation) - qed -qed - -corollary measurable_on_componentwise: - "f measurable_on S \ (\i\Basis. (\x. (f x \ i) *\<^sub>R i) measurable_on S)" - apply (subst measurable_on_UNIV [symmetric]) - apply (subst measurable_on_componentwise_UNIV) - apply (simp add: measurable_on_UNIV if_distrib [of "\x. inner x _"] if_distrib [of "\x. scaleR x _"] cong: if_cong) - done - - -lemma\<^marker>\tag important\ borel_measurable_implies_simple_function_sequence_real: - fixes u :: "'a \ real" - assumes u[measurable]: "u \ borel_measurable M" and nn: "\x. u x \ 0" - shows "\f. incseq f \ (\i. simple_function M (f i)) \ (\x. bdd_above (range (\i. f i x))) \ - (\i x. 0 \ f i x) \ u = (SUP i. f i)" -proof - - define f where [abs_def]: - "f i x = real_of_int (floor ((min i (u x)) * 2^i)) / 2^i" for i x - - have [simp]: "0 \ f i x" for i x - by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg nn) - - have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x - by simp - - have "real_of_int \real i * 2 ^ i\ = real_of_int \i * 2 ^ i\" for i - by (intro arg_cong[where f=real_of_int]) simp - then have [simp]: "real_of_int \real i * 2 ^ i\ = i * 2 ^ i" for i - unfolding floor_of_nat by simp - - have bdd: "bdd_above (range (\i. f i x))" for x - by (rule bdd_aboveI [where M = "u x"]) (auto simp: f_def field_simps min_def) - - have "incseq f" - proof (intro monoI le_funI) - fix m n :: nat and x assume "m \ n" - moreover - { fix d :: nat - have "\2^d::real\ * \2^m * (min (of_nat m) (u x))\ \ \2^d * (2^m * (min (of_nat m) (u x)))\" - by (rule le_mult_floor) (auto simp: nn) - also have "\ \ \2^d * (2^m * (min (of_nat d + of_nat m) (u x)))\" - by (intro floor_mono mult_mono min.mono) - (auto simp: nn min_less_iff_disj of_nat_less_top) - finally have "f m x \ f(m + d) x" - unfolding f_def - by (auto simp: field_simps power_add * simp del: of_int_mult) } - ultimately show "f m x \ f n x" - by (auto simp: le_iff_add) - qed - then have inc_f: "incseq (\i. f i x)" for x - by (auto simp: incseq_def le_fun_def) - moreover - have "simple_function M (f i)" for i - proof (rule simple_function_borel_measurable) - have "\(min (of_nat i) (u x)) * 2 ^ i\ \ \int i * 2 ^ i\" for x - by (auto split: split_min intro!: floor_mono) - then have "f i ` space M \ (\n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}" - unfolding floor_of_int by (auto simp: f_def nn intro!: imageI) - then show "finite (f i ` space M)" - by (rule finite_subset) auto - show "f i \ borel_measurable M" - unfolding f_def enn2real_def by measurable - qed - moreover - { fix x - have "(SUP i. (f i x)) = u x" - proof - - obtain n where "u x \ of_nat n" using real_arch_simple by auto - then have min_eq_r: "\\<^sub>F i in sequentially. min (real i) (u x) = u x" - by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min) - have "(\i. real_of_int \min (real i) (u x) * 2^i\ / 2^i) \ u x" - proof (rule tendsto_sandwich) - show "(\n. u x - (1/2)^n) \ u x" - by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero) - show "\\<^sub>F n in sequentially. real_of_int \min (real n) (u x) * 2 ^ n\ / 2 ^ n \ u x" - using min_eq_r by eventually_elim (auto simp: field_simps) - have *: "u x * (2 ^ n * 2 ^ n) \ 2^n + 2^n * real_of_int \u x * 2 ^ n\" for n - using real_of_int_floor_ge_diff_one[of "u x * 2^n", THEN mult_left_mono, of "2^n"] - by (auto simp: field_simps) - show "\\<^sub>F n in sequentially. u x - (1/2)^n \ real_of_int \min (real n) (u x) * 2 ^ n\ / 2 ^ n" - using min_eq_r by eventually_elim (insert *, auto simp: field_simps) - qed auto - then have "(\i. (f i x)) \ u x" - by (simp add: f_def) - from LIMSEQ_unique LIMSEQ_incseq_SUP [OF bdd inc_f] this - show ?thesis - by blast - qed } - ultimately show ?thesis - by (intro exI [of _ "\i x. f i x"]) (auto simp: \incseq f\ bdd image_comp) -qed - - -lemma homeomorphic_open_interval_UNIV: - fixes a b:: real - assumes "a < b" - shows "{a<.. {}" - shows "box a b homeomorphic (UNIV::'a set)" -proof - - have "{a \ i <.. i} homeomorphic (UNIV::real set)" if "i \ Basis" for i - using assms box_ne_empty that by (blast intro: homeomorphic_open_interval_UNIV) - then have "\f g. (\x. a \ i < x \ x < b \ i \ g (f x) = x) \ - (\y. a \ i < g y \ g y < b \ i \ f(g y) = y) \ - continuous_on {a \ i<.. i} f \ - continuous_on (UNIV::real set) g" - if "i \ Basis" for i - using that by (auto simp: homeomorphic_minimal mem_box Ball_def) - then obtain f g where gf: "\i x. \i \ Basis; a \ i < x; x < b \ i\ \ g i (f i x) = x" - and fg: "\i y. i \ Basis \ a \ i < g i y \ g i y < b \ i \ f i (g i y) = y" - and contf: "\i. i \ Basis \ continuous_on {a \ i<.. i} (f i)" - and contg: "\i. i \ Basis \ continuous_on (UNIV::real set) (g i)" - by metis - define F where "F \ \x. \i\Basis. (f i (x \ i)) *\<^sub>R i" - define G where "G \ \x. \i\Basis. (g i (x \ i)) *\<^sub>R i" - show ?thesis - unfolding homeomorphic_minimal - proof (intro exI conjI ballI) - show "G y \ box a b" for y - using fg by (simp add: G_def mem_box) - show "G (F x) = x" if "x \ box a b" for x - using that by (simp add: F_def G_def gf mem_box euclidean_representation) - show "F (G y) = y" for y - by (simp add: F_def G_def fg mem_box euclidean_representation) - show "continuous_on (box a b) F" - unfolding F_def - proof (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_inner]) - show "(\x. x \ i) ` box a b \ {a \ i<.. i}" if "i \ Basis" for i - using that by (auto simp: mem_box) - qed - show "continuous_on UNIV G" - unfolding G_def - by (intro continuous_intros continuous_on_compose2 [OF contg continuous_on_inner]) auto - qed auto -qed - - - -lemma diff_null_sets_lebesgue: "\N \ null_sets (lebesgue_on S); X-N \ sets (lebesgue_on S); N \ X\ - \ X \ sets (lebesgue_on S)" - by (metis Int_Diff_Un inf.commute inf.orderE null_setsD2 sets.Un) - -lemma borel_measurable_diff_null: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes N: "N \ null_sets (lebesgue_on S)" and S: "S \ sets lebesgue" - shows "f \ borel_measurable (lebesgue_on (S-N)) \ f \ borel_measurable (lebesgue_on S)" - unfolding in_borel_measurable lebesgue_on_UNIV_eq space_lebesgue_on sets_restrict_UNIV -proof (intro ball_cong iffI) - show "f -` T \ S \ sets (lebesgue_on S)" - if "f -` T \ (S-N) \ sets (lebesgue_on (S-N))" for T - using that assms - by (smt Diff_Int_distrib completion.complete2 diff_null_sets_lebesgue inf.idem inf_le2 inf_mono lebesgue_on_UNIV_eq null_setsD2 null_sets_restrict_space sets.Diff sets_restrict_space_iff space_lebesgue_on space_restrict_space) - show "f -` T \ (S-N) \ sets (lebesgue_on (S-N))" - if "f -` T \ S \ sets (lebesgue_on S)" for T - using image_eqI inf.commute inf_top_right sets_restrict_space that - by (smt Int_Diff S sets.Int_space_eq2 sets_restrict_space_iff space_lebesgue_on) -qed auto - -lemma lebesgue_measurable_diff_null: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "N \ null_sets lebesgue" - shows "f \ borel_measurable (lebesgue_on (-N)) \ f \ borel_measurable lebesgue" - by (simp add: Compl_eq_Diff_UNIV assms borel_measurable_diff_null lebesgue_on_UNIV_eq) - - - -proposition measurable_on_imp_borel_measurable_lebesgue_UNIV: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f measurable_on UNIV" - shows "f \ borel_measurable lebesgue" -proof - - obtain N and F - where NF: "negligible N" - and conF: "\n. continuous_on UNIV (F n)" - and tendsF: "\x. x \ N \ (\n. F n x) \ f x" - using assms by (auto simp: measurable_on_def) - obtain N where "N \ null_sets lebesgue" "f \ borel_measurable (lebesgue_on (-N))" - proof - show "f \ borel_measurable (lebesgue_on (- N))" - proof (rule borel_measurable_LIMSEQ_metric) - show "F i \ borel_measurable (lebesgue_on (- N))" for i - by (meson Compl_in_sets_lebesgue NF conF continuous_imp_measurable_on_sets_lebesgue continuous_on_subset negligible_imp_sets subset_UNIV) - show "(\i. F i x) \ f x" if "x \ space (lebesgue_on (- N))" for x - using that - by (simp add: tendsF) - qed - show "N \ null_sets lebesgue" - using NF negligible_iff_null_sets by blast - qed - then show ?thesis - using lebesgue_measurable_diff_null by blast -qed - -corollary measurable_on_imp_borel_measurable_lebesgue: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f measurable_on S" and S: "S \ sets lebesgue" - shows "f \ borel_measurable (lebesgue_on S)" -proof - - have "(\x. if x \ S then f x else 0) measurable_on UNIV" - using assms(1) measurable_on_UNIV by blast - then show ?thesis - by (simp add: borel_measurable_if_D measurable_on_imp_borel_measurable_lebesgue_UNIV) -qed - - -proposition measurable_on_limit: - fixes f :: "nat \ 'a::euclidean_space \ 'b::euclidean_space" - assumes f: "\n. f n measurable_on S" and N: "negligible N" - and lim: "\x. x \ S - N \ (\n. f n x) \ g x" - shows "g measurable_on S" -proof - - have "box (0::'b) One homeomorphic (UNIV::'b set)" - by (simp add: homeomorphic_box_UNIV) - then obtain h h':: "'b\'b" where hh': "\x. x \ box 0 One \ h (h' x) = x" - and h'im: "h' ` box 0 One = UNIV" - and conth: "continuous_on UNIV h" - and conth': "continuous_on (box 0 One) h'" - and h'h: "\y. h' (h y) = y" - and rangeh: "range h = box 0 One" - by (auto simp: homeomorphic_def homeomorphism_def) - have "norm y \ DIM('b)" if y: "y \ box 0 One" for y::'b - proof - - have y01: "0 < y \ i" "y \ i < 1" if "i \ Basis" for i - using that y by (auto simp: mem_box) - have "norm y \ (\i\Basis. \y \ i\)" - using norm_le_l1 by blast - also have "\ \ (\i::'b\Basis. 1)" - proof (rule sum_mono) - show "\y \ i\ \ 1" if "i \ Basis" for i - using y01 that by fastforce - qed - also have "\ \ DIM('b)" - by auto - finally show ?thesis . - qed - then have norm_le: "norm(h y) \ DIM('b)" for y - by (metis UNIV_I image_eqI rangeh) - have "(h' \ (h \ (\x. if x \ S then g x else 0))) measurable_on UNIV" - proof (rule measurable_on_compose_continuous_box) - let ?\ = "h \ (\x. if x \ S then g x else 0)" - let ?f = "\n. h \ (\x. if x \ S then f n x else 0)" - show "?\ measurable_on UNIV" - proof (rule integrable_subintervals_imp_measurable) - show "?\ integrable_on cbox a b" for a b - proof (rule integrable_spike_set) - show "?\ integrable_on (cbox a b - N)" - proof (rule dominated_convergence_integrable) - show const: "(\x. DIM('b)) integrable_on cbox a b - N" - by (simp add: N has_integral_iff integrable_const integrable_negligible integrable_setdiff negligible_diff) - show "norm ((h \ (\x. if x \ S then g x else 0)) x) \ DIM('b)" if "x \ cbox a b - N" for x - using that norm_le by (simp add: o_def) - show "(\k. ?f k x) \ ?\ x" if "x \ cbox a b - N" for x - using that lim [of x] conth - by (auto simp: continuous_on_def intro: tendsto_compose) - show "(?f n) absolutely_integrable_on cbox a b - N" for n - proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable) - show "?f n \ borel_measurable (lebesgue_on (cbox a b - N))" - proof (rule measurable_on_imp_borel_measurable_lebesgue [OF measurable_on_spike_set]) - show "?f n measurable_on cbox a b" - unfolding measurable_on_UNIV [symmetric, of _ "cbox a b"] - proof (rule measurable_on_restrict) - have f': "(\x. if x \ S then f n x else 0) measurable_on UNIV" - by (simp add: f measurable_on_UNIV) - show "?f n measurable_on UNIV" - using measurable_on_compose_continuous [OF f' conth] by auto - qed auto - show "negligible (cbox a b \ (cbox a b - N))" - by (auto intro: negligible_subset [OF N]) - show "cbox a b - N \ sets lebesgue" - by (simp add: N negligible_imp_sets sets.Diff) - qed - show "cbox a b - N \ sets lebesgue" - by (simp add: N negligible_imp_sets sets.Diff) - show "norm (?f n x) \ DIM('b)" - if "x \ cbox a b - N" for x - using that local.norm_le by simp - qed (auto simp: const) - qed - show "negligible {x \ cbox a b - N - cbox a b. ?\ x \ 0}" - by (auto simp: empty_imp_negligible) - have "{x \ cbox a b - (cbox a b - N). ?\ x \ 0} \ N" - by auto - then show "negligible {x \ cbox a b - (cbox a b - N). ?\ x \ 0}" - using N negligible_subset by blast - qed - qed - show "?\ x \ box 0 One" for x - using rangeh by auto - show "continuous_on (box 0 One) h'" - by (rule conth') - qed - then show ?thesis - by (simp add: o_def h'h measurable_on_UNIV) -qed - - -lemma measurable_on_if_simple_function_limit: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - shows "\\n. g n measurable_on UNIV; \n. finite (range (g n)); \x. (\n. g n x) \ f x\ - \ f measurable_on UNIV" - by (force intro: measurable_on_limit [where N="{}"]) - - -lemma lebesgue_measurable_imp_measurable_on_nnreal_UNIV: - fixes u :: "'a::euclidean_space \ real" - assumes u: "u \ borel_measurable lebesgue" and nn: "\x. u x \ 0" - shows "u measurable_on UNIV" -proof - - obtain f where "incseq f" and f: "\i. simple_function lebesgue (f i)" - and bdd: "\x. bdd_above (range (\i. f i x))" - and nnf: "\i x. 0 \ f i x" and *: "u = (SUP i. f i)" - using borel_measurable_implies_simple_function_sequence_real nn u by metis - show ?thesis - unfolding * - proof (rule measurable_on_if_simple_function_limit [of concl: "\ range f"]) - show "(f i) measurable_on UNIV" for i - by (simp add: f nnf simple_function_measurable_on_UNIV) - show "finite (range (f i))" for i - by (metis f simple_function_def space_borel space_completion space_lborel) - show "(\i. f i x) \ (\ range f) x" for x - proof - - have "incseq (\i. f i x)" - using \incseq f\ apply (auto simp: incseq_def) - by (simp add: le_funD) - then show ?thesis - by (metis SUP_apply bdd LIMSEQ_incseq_SUP) - qed - qed -qed - -lemma lebesgue_measurable_imp_measurable_on_nnreal: - fixes u :: "'a::euclidean_space \ real" - assumes "u \ borel_measurable lebesgue" "\x. u x \ 0""S \ sets lebesgue" - shows "u measurable_on S" - unfolding measurable_on_UNIV [symmetric, of u] - using assms - by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal_UNIV) - -lemma lebesgue_measurable_imp_measurable_on_real: - fixes u :: "'a::euclidean_space \ real" - assumes u: "u \ borel_measurable lebesgue" and S: "S \ sets lebesgue" - shows "u measurable_on S" -proof - - let ?f = "\x. \u x\ + u x" - let ?g = "\x. \u x\ - u x" - have "?f measurable_on S" "?g measurable_on S" - using S u by (auto intro: lebesgue_measurable_imp_measurable_on_nnreal) - then have "(\x. (?f x - ?g x) / 2) measurable_on S" - using measurable_on_cdivide measurable_on_diff by blast - then show ?thesis - by auto -qed - - -proposition lebesgue_measurable_imp_measurable_on: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes f: "f \ borel_measurable lebesgue" and S: "S \ sets lebesgue" - shows "f measurable_on S" - unfolding measurable_on_componentwise [of f] -proof - fix i::'b - assume "i \ Basis" - have "(\x. (f x \ i)) \ borel_measurable lebesgue" - using \i \ Basis\ borel_measurable_euclidean_space f by blast - then have "(\x. (f x \ i)) measurable_on S" - using S lebesgue_measurable_imp_measurable_on_real by blast - then show "(\x. (f x \ i) *\<^sub>R i) measurable_on S" - by (intro measurable_on_scaleR measurable_on_const S) -qed - -proposition measurable_on_iff_borel_measurable: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "S \ sets lebesgue" - shows "f measurable_on S \ f \ borel_measurable (lebesgue_on S)" (is "?lhs = ?rhs") -proof - show "f \ borel_measurable (lebesgue_on S)" - if "f measurable_on S" - using that by (simp add: assms measurable_on_imp_borel_measurable_lebesgue) -next - assume "f \ borel_measurable (lebesgue_on S)" - then have "(\a. if a \ S then f a else 0) measurable_on UNIV" - by (simp add: assms borel_measurable_if_I lebesgue_measurable_imp_measurable_on) - then show "f measurable_on S" - using measurable_on_UNIV by blast -qed - -lemma measurable_on_bilinear: - fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::euclidean_space" - assumes h: "bilinear h" and f: "f measurable_on S" and g: "g measurable_on S" - shows "(\x. h (f x) (g x)) measurable_on S" -proof (rule measurable_on_combine [where h = h]) - show "continuous_on UNIV (\x. h (fst x) (snd x))" - by (metis (mono_tags, lifting) bilinear_continuous_on_compose continuous_on_cong continuous_on_fst continuous_on_id continuous_on_snd h) - show "h 0 0 = 0" - by (simp add: bilinear_lzero h) -qed (auto intro: assms) - -lemma borel_measurable_bilinear: - fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::euclidean_space" - assumes "bilinear h" "f \ borel_measurable (lebesgue_on S)" "g \ borel_measurable (lebesgue_on S)" - and S: "S \ sets lebesgue" - shows "(\x. h (f x) (g x)) \ borel_measurable (lebesgue_on S)" - using assms measurable_on_bilinear [of h f S g] - by (simp flip: measurable_on_iff_borel_measurable) - - -lemma absolutely_integrable_bounded_measurable_product: - fixes h :: "'a::euclidean_space \ 'b::euclidean_space \ 'c::euclidean_space" - assumes "bilinear h" and f: "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" - and bou: "bounded (f ` S)" and g: "g absolutely_integrable_on S" - shows "(\x. h (f x) (g x)) absolutely_integrable_on S" -proof - - obtain B where "B > 0" and B: "\x y. norm (h x y) \ B * norm x * norm y" - using bilinear_bounded_pos \bilinear h\ by blast - obtain C where "C > 0" and C: "\x. x \ S \ norm (f x) \ C" - using bounded_pos by (metis bou imageI) - show ?thesis - proof (rule measurable_bounded_by_integrable_imp_absolutely_integrable [OF _ \S \ sets lebesgue\]) - show "norm (h (f x) (g x)) \ B * C * norm(g x)" if "x \ S" for x - by (meson less_le mult_left_mono mult_right_mono norm_ge_zero order_trans that \B > 0\ B C) - show "(\x. h (f x) (g x)) \ borel_measurable (lebesgue_on S)" - using \bilinear h\ f g absolutely_integrable_measurable borel_measurable_bilinear by blast - show "(\x. B * C * norm(g x)) integrable_on S" - using \0 < B\ \0 < C\ absolutely_integrable_on_def g by auto - qed -qed - -lemma absolutely_integrable_bounded_measurable_product_real: - fixes f :: "real \ real" - assumes "f \ borel_measurable (lebesgue_on S)" "S \ sets lebesgue" - and "bounded (f ` S)" and "g absolutely_integrable_on S" - shows "(\x. f x * g x) absolutely_integrable_on S" - using absolutely_integrable_bounded_measurable_product bilinear_times assms by blast - - - - -lemma set_integrable_mult_right_iff [simp]: - fixes a :: "'a::{real_normed_field, second_countable_topology}" - assumes "a \ 0" - shows "set_integrable M A (\t. a * f t) \ set_integrable M A f" -proof - assume "set_integrable M A (\t. a * f t)" - then have "set_integrable M A (\t. 1/a * (a * f t))" - using set_integrable_mult_right by blast - then show "set_integrable M A f" - using assms by auto -qed auto - -lemma set_integrable_mult_left_iff [simp]: - fixes a :: "'a::{real_normed_field, second_countable_topology}" - assumes "a \ 0" - shows "set_integrable M A (\t. f t * a) \ set_integrable M A f" - using assms - by (subst set_integrable_mult_right_iff [symmetric]) (auto simp: mult.commute) - - -lemma set_integrable_mult_divide_iff [simp]: - fixes a :: "'a::{real_normed_field, second_countable_topology}" - assumes "a \ 0" - shows "set_integrable M A (\t. f t / a) \ set_integrable M A f" - by (simp add: divide_inverse assms) - - -lemma borel_measurable_AE: - fixes f :: "'a::euclidean_space \ 'b::euclidean_space" - assumes "f \ borel_measurable lebesgue" and ae: "AE x in lebesgue. f x = g x" - shows "g \ borel_measurable lebesgue" -proof - - obtain N where N: "N \ null_sets lebesgue" "\x. x \ N \ f x = g x" - using AE_null_sets_lebesgue [OF ae] by blast - have "f measurable_on UNIV" - by (simp add: assms lebesgue_measurable_imp_measurable_on) - then have "g measurable_on UNIV" - by (metis Diff_iff N measurable_on_spike negligible_iff_null_sets) - then show ?thesis - using measurable_on_imp_borel_measurable_lebesgue_UNIV by blast -qed - - -lemma has_bochner_integral_combine: - fixes f :: "real \ 'a::euclidean_space" - assumes "a \ c" "c \ b" - and ac: "has_bochner_integral (lebesgue_on {a..c}) f i" - and cb: "has_bochner_integral (lebesgue_on {c..b}) f j" - shows "has_bochner_integral (lebesgue_on {a..b}) f(i + j)" -proof - - have i: "has_bochner_integral lebesgue (\x. indicator {a..c} x *\<^sub>R f x) i" - and j: "has_bochner_integral lebesgue (\x. indicator {c..b} x *\<^sub>R f x) j" - using assms by (auto simp: has_bochner_integral_restrict_space) - have AE: "AE x in lebesgue. indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" - proof (rule AE_I') - have eq: "indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" if "x \ c" for x - using assms that by (auto simp: indicator_def) - then show "{x \ space lebesgue. indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x \ indicat_real {a..b} x *\<^sub>R f x} \ {c}" - by auto - qed auto - have "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) (i + j)" - proof (rule has_bochner_integralI_AE [OF has_bochner_integral_add [OF i j] _ AE]) - have eq: "indicat_real {a..c} x *\<^sub>R f x + indicat_real {c..b} x *\<^sub>R f x = indicat_real {a..b} x *\<^sub>R f x" if "x \ c" for x - using assms that by (auto simp: indicator_def) - show "(\x. indicat_real {a..b} x *\<^sub>R f x) \ borel_measurable lebesgue" - proof (rule borel_measurable_AE [OF borel_measurable_add AE]) - show "(\x. indicator {a..c} x *\<^sub>R f x) \ borel_measurable lebesgue" - "(\x. indicator {c..b} x *\<^sub>R f x) \ borel_measurable lebesgue" - using i j by auto - qed - qed - then show ?thesis - by (simp add: has_bochner_integral_restrict_space) -qed - -lemma integrable_combine: - fixes f :: "real \ 'a::euclidean_space" - assumes "integrable (lebesgue_on {a..c}) f" "integrable (lebesgue_on {c..b}) f" - and "a \ c" "c \ b" - shows "integrable (lebesgue_on {a..b}) f" - using assms has_bochner_integral_combine has_bochner_integral_iff by blast - -lemma integral_combine: - fixes f :: "real \ 'a::euclidean_space" - assumes f: "integrable (lebesgue_on {a..b}) f" and "a \ c" "c \ b" - shows "integral\<^sup>L (lebesgue_on {a..b}) f = integral\<^sup>L (lebesgue_on {a..c}) f + integral\<^sup>L (lebesgue_on {c..b}) f" -proof - - have i: "has_bochner_integral (lebesgue_on {a..c}) f(integral\<^sup>L (lebesgue_on {a..c}) f)" - using integrable_subinterval_real \c \ b\ f has_bochner_integral_iff by fastforce - have j: "has_bochner_integral (lebesgue_on {c..b}) f(integral\<^sup>L (lebesgue_on {c..b}) f)" - using integrable_subinterval_real \a \ c\ f has_bochner_integral_iff by fastforce - show ?thesis - by (meson \a \ c\ \c \ b\ has_bochner_integral_combine has_bochner_integral_iff i j) -qed - -lemma has_bochner_integral_reflect_real_lemma[intro]: - fixes f :: "real \ 'a::euclidean_space" - assumes "has_bochner_integral (lebesgue_on {a..b}) f i" - shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i" -proof - - have eq: "indicat_real {a..b} (- x) *\<^sub>R f(- x) = indicat_real {- b..- a} x *\<^sub>R f(- x)" for x - by (auto simp: indicator_def) - have i: "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) i" - using assms by (auto simp: has_bochner_integral_restrict_space) - then have "has_bochner_integral lebesgue (\x. indicator {-b..-a} x *\<^sub>R f(-x)) i" - using has_bochner_integral_lebesgue_real_affine_iff [of "-1" "(\x. indicator {a..b} x *\<^sub>R f x)" i 0] - by (auto simp: eq) - then show ?thesis - using assms by blast -qed - -lemma has_bochner_integral_reflect_real[simp]: - fixes f :: "real \ 'a::euclidean_space" - shows "has_bochner_integral (lebesgue_on {-b..-a}) (\x. f(-x)) i \ has_bochner_integral (lebesgue_on {a..b}) f i" - by (auto simp: dest: has_bochner_integral_reflect_real_lemma) - lemma has_bochner_integral_null [intro]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "N \ null_sets lebesgue" shows "has_bochner_integral (lebesgue_on N) f 0" unfolding has_bochner_integral_iff proof show "integrable (lebesgue_on N) f" proof (subst integrable_restrict_space) show "N \ space lebesgue \ sets lebesgue" using assms by force show "integrable lebesgue (\x. indicat_real N x *\<^sub>R f x)" proof (rule integrable_cong_AE_imp) show "integrable lebesgue (\x. 0)" by simp show *: "AE x in lebesgue. 0 = indicat_real N x *\<^sub>R f x" using assms by (simp add: indicator_def completion.null_sets_iff_AE eventually_mono) show "(\x. indicat_real N x *\<^sub>R f x) \ borel_measurable lebesgue" by (auto intro: borel_measurable_AE [OF _ *]) qed qed show "integral\<^sup>L (lebesgue_on N) f = 0" proof (rule integral_eq_zero_AE) show "AE x in lebesgue_on N. f x = 0" by (rule AE_I' [where N=N]) (auto simp: assms null_setsD2 null_sets_restrict_space) qed qed - lemma has_bochner_integral_null_eq[simp]: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "N \ null_sets lebesgue" shows "has_bochner_integral (lebesgue_on N) f i \ i = 0" using assms has_bochner_integral_eq by blast -lemma integrable_reflect_real[simp]: - fixes f :: "real \ 'a::euclidean_space" - shows "integrable (lebesgue_on {-b..-a}) (\x. f(-x)) \ integrable (lebesgue_on {a..b}) f" - by (metis has_bochner_integral_iff has_bochner_integral_reflect_real) - -lemma integral_reflect_real[simp]: - fixes f :: "real \ 'a::euclidean_space" - shows "integral\<^sup>L (lebesgue_on {-b .. -a}) (\x. f(-x)) = integral\<^sup>L (lebesgue_on {a..b::real}) f" - using has_bochner_integral_reflect_real [of b a f] - by (metis has_bochner_integral_iff not_integrable_integral_eq) - -lemma monoseq_convergent: - fixes X :: "nat \ real" - assumes X: "monoseq X" and B: "\i. \X i\ \ B" - obtains L where "X \ L" - using X unfolding monoseq_iff -proof - assume "incseq X" - show thesis - using abs_le_D1 [OF B] incseq_convergent [OF \incseq X\] that by meson -next - assume "decseq X" - show thesis - using decseq_convergent [OF \decseq X\] that - by (metis B abs_le_iff add.inverse_inverse neg_le_iff_le) -qed - end diff --git a/thys/Fourier/Lspace.thy b/thys/Fourier/Lspace.thy --- a/thys/Fourier/Lspace.thy +++ b/thys/Fourier/Lspace.thy @@ -1,150 +1,150 @@ section\Lspace as it is in HOL Light\ text\Mainly a repackaging of existing material from Lp\ theory Lspace - imports Lp.Lp Fourier_Aux1 + imports Lp.Lp begin abbreviation lspace :: "'a measure \ ennreal \ ('a \ real) set" where "lspace M p \ space\<^sub>N (\ p M)" lemma lspace_1: assumes "S \ sets lebesgue" shows "f \ lspace (lebesgue_on S) 1 \ f absolutely_integrable_on S" using assms by (simp add: L1_space integrable_restrict_space set_integrable_def) lemma lspace_ennreal_iff: assumes "p > 0" shows "lspace (lebesgue_on S) (ennreal p) = {f \ borel_measurable (lebesgue_on S). integrable (lebesgue_on S) (\x. (norm(f x) powr p))}" using assms by (fastforce simp: Lp_measurable Lp_D intro: Lp_I) lemma lspace_iff: assumes "\ > p" "p > 0" shows "lspace (lebesgue_on S) p = {f \ borel_measurable (lebesgue_on S). integrable (lebesgue_on S) (\x. (norm(f x) powr (enn2real p)))}" proof - obtain q::real where "p = enn2real q" using Lp_cases assms by auto then show ?thesis using assms lspace_ennreal_iff by auto qed lemma lspace_iff': assumes p: "\ > p" "p > 0" and S: "S \ sets lebesgue" shows "lspace (lebesgue_on S) p = {f \ borel_measurable (lebesgue_on S). (\x. (norm(f x) powr (enn2real p))) integrable_on S}" (is "?lhs = ?rhs") proof show "?lhs \ ?rhs" using assms integrable_on_lebesgue_on by (auto simp: lspace_iff) next show "?rhs \ ?lhs" proof (clarsimp simp: lspace_iff [OF p]) show "integrable (lebesgue_on S) (\xa. \f xa\ powr enn2real p)" if "f \ borel_measurable (lebesgue_on S)" and "(\x. \f x\ powr enn2real p) integrable_on S" for f proof - have "(\x. \f x\ powr enn2real p) absolutely_integrable_on S" by (simp add: absolutely_integrable_on_iff_nonneg that(2)) then show ?thesis using L1_space S lspace_1 by blast qed qed qed lemma lspace_mono: assumes "f \ lspace (lebesgue_on S) q" and S: "S \ lmeasurable" and "p > 0" "p \ q" "q < \" shows "f \ lspace (lebesgue_on S) p" proof - have "p < \" using assms by (simp add: top.not_eq_extremum) with assms show ?thesis proof (clarsimp simp add: lspace_iff') show "(\x. \f x\ powr enn2real p) integrable_on S" if "f \ borel_measurable (lebesgue_on S)" and "(\x. \f x\ powr enn2real q) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) show "(\x. \f x\ powr enn2real p) \ borel_measurable (lebesgue_on S)" using measurable_abs_powr that(1) by blast let ?g = "\x. max 1 (norm(f x) powr enn2real q)" have "?g absolutely_integrable_on S" proof (rule absolutely_integrable_max_1) show "(\x. norm (f x) powr enn2real q) absolutely_integrable_on S" by (simp add: nonnegative_absolutely_integrable_1 that(2)) qed (simp add: S) then show "?g integrable_on S" using absolutely_integrable_abs_iff by blast show "norm (\f x\ powr enn2real p) \ ?g x" if "x \ S" for x proof - have "\f x\ powr enn2real p \ \f x\ powr enn2real q" if "1 \ \f x\" using assms enn2real_mono powr_mono that by auto then show ?thesis using powr_le1 by (fastforce simp add: le_max_iff_disj) qed show "S \ sets lebesgue" by (simp add: S fmeasurableD) qed qed qed lemma lspace_inclusion: assumes "S \ lmeasurable" and "p > 0" "p \ q" "q < \" shows "lspace (lebesgue_on S) q \ lspace (lebesgue_on S) p" using assms lspace_mono by auto lemma lspace_const: fixes p::real assumes "p > 0""S \ lmeasurable" shows "(\x. c) \ lspace (lebesgue_on S) p" by (simp add: Lp_space assms finite_measure.integrable_const finite_measure_lebesgue_on) lemma lspace_max: fixes p::real assumes "f \ lspace (lebesgue_on S) p" "g \ lspace (lebesgue_on S) p" "p > 0" shows "(\x. max (f x) (g x)) \ lspace (lebesgue_on S) p" proof - have "integrable (lebesgue_on S) (\x. \max (f x) (g x)\ powr p)" if f: "f \ borel_measurable (lebesgue_on S)" "integrable (lebesgue_on S) (\x. \f x\ powr p)" and g: "g \ borel_measurable (lebesgue_on S)" "integrable (lebesgue_on S) (\x. \g x\ powr p)" proof - have "integrable (lebesgue_on S) (\x. \\f x\ powr p\)" "integrable (lebesgue_on S) (\x. \\g x\ powr p\)" using integrable_abs that by blast+ then have "integrable (lebesgue_on S) (\x. max (\\f x\ powr p\) (\\g x\ powr p\))" using integrable_max by blast then show ?thesis proof (rule Bochner_Integration.integrable_bound) show "(\x. \max (f x) (g x)\ powr p) \ borel_measurable (lebesgue_on S)" using borel_measurable_max measurable_abs_powr that by blast qed auto qed then show ?thesis using assms by (auto simp: Lp_space borel_measurable_max) qed lemma lspace_min: fixes p::real assumes "f \ lspace (lebesgue_on S) p" "g \ lspace (lebesgue_on S) p" "p > 0" shows "(\x. min (f x) (g x)) \ lspace (lebesgue_on S) p" proof - have "integrable (lebesgue_on S) (\x. \min (f x) (g x)\ powr p)" if f: "f \ borel_measurable (lebesgue_on S)" "integrable (lebesgue_on S) (\x. \f x\ powr p)" and g: "g \ borel_measurable (lebesgue_on S)" "integrable (lebesgue_on S) (\x. \g x\ powr p)" proof - have "integrable (lebesgue_on S) (\x. \\f x\ powr p\)" "integrable (lebesgue_on S) (\x. \\g x\ powr p\)" using integrable_abs that by blast+ then have "integrable (lebesgue_on S) (\x. max (\\f x\ powr p\) (\\g x\ powr p\))" using integrable_max by blast then show ?thesis proof (rule Bochner_Integration.integrable_bound) show "(\x. \min (f x) (g x)\ powr p) \ borel_measurable (lebesgue_on S)" using borel_measurable_min measurable_abs_powr that by blast qed auto qed then show ?thesis using assms by (auto simp: Lp_space borel_measurable_min) qed lemma Lp_space_numeral: assumes "numeral n > (0::int)" shows "space\<^sub>N (\ (numeral n) M) = {f \ borel_measurable M. integrable M (\x. \f x\ ^ numeral n)}" using Lp_space [of "numeral n" M] assms by simp end diff --git a/thys/Fourier/Periodic.thy b/thys/Fourier/Periodic.thy --- a/thys/Fourier/Periodic.thy +++ b/thys/Fourier/Periodic.thy @@ -1,187 +1,221 @@ section\Shifting the origin for integration of periodic functions\ theory Periodic - imports "HOL-Analysis.Analysis" Fourier_Aux2 + imports "HOL-Analysis.Analysis" begin +lemma has_bochner_integral_null [intro]: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "N \ null_sets lebesgue" + shows "has_bochner_integral (lebesgue_on N) f 0" + unfolding has_bochner_integral_iff +proof + show "integrable (lebesgue_on N) f" + proof (subst integrable_restrict_space) + show "N \ space lebesgue \ sets lebesgue" + using assms by force + show "integrable lebesgue (\x. indicat_real N x *\<^sub>R f x)" + proof (rule integrable_cong_AE_imp) + show "integrable lebesgue (\x. 0)" + by simp + show *: "AE x in lebesgue. 0 = indicat_real N x *\<^sub>R f x" + using assms + by (simp add: indicator_def completion.null_sets_iff_AE eventually_mono) + show "(\x. indicat_real N x *\<^sub>R f x) \ borel_measurable lebesgue" + by (auto intro: borel_measurable_AE [OF _ *]) + qed + qed + show "integral\<^sup>L (lebesgue_on N) f = 0" + proof (rule integral_eq_zero_AE) + show "AE x in lebesgue_on N. f x = 0" + by (rule AE_I' [where N=N]) (auto simp: assms null_setsD2 null_sets_restrict_space) + qed +qed + +lemma has_bochner_integral_null_eq[simp]: + fixes f :: "'a::euclidean_space \ 'b::euclidean_space" + assumes "N \ null_sets lebesgue" + shows "has_bochner_integral (lebesgue_on N) f i \ i = 0" + using assms has_bochner_integral_eq by blast + lemma periodic_integer_multiple: "(\x. f(x + a) = f x) \ (\x. \n \ \. f(x + n * a) = f x)" (is "?lhs = ?rhs") proof assume L [rule_format]: ?lhs have nat: "f(x + of_nat n * a) = f x" for x n proof (induction n) case (Suc n) with L [of "x + of_nat n * a"] show ?case by (simp add: algebra_simps) qed auto have "f(x + of_int z * a) = f x" for x z proof (cases "z \ 0") case True then show ?thesis by (metis nat of_nat_nat) next case False then show ?thesis using nat [of "x + of_int z * a" "nat (-z)"] by auto qed then show ?rhs by (auto simp: Ints_def) qed (use Ints_1 in fastforce) lemma has_integral_offset: fixes f :: "real \ 'a::euclidean_space" assumes "has_bochner_integral (lebesgue_on {a..b}) f i" shows "has_bochner_integral (lebesgue_on {a-c..b-c}) (\x. f(x + c)) i" proof - have eq: "indicat_real {a..b} (x + c) = indicat_real {a-c..b-c} x" for x by (auto simp: indicator_def) have "has_bochner_integral lebesgue (\x. indicator {a..b} x *\<^sub>R f x) i" using assms by (auto simp: has_bochner_integral_restrict_space) then have "has_bochner_integral lebesgue (\x. indicat_real {a-c..b-c} x *\<^sub>R f(x + c)) i" using has_bochner_integral_lebesgue_real_affine_iff [of 1 "(\x. indicator {a..b} x *\<^sub>R f x)" i c] by (simp add: add_ac eq) then show ?thesis using assms by (auto simp: has_bochner_integral_restrict_space) qed lemma has_integral_periodic_offset_lemma: fixes f :: "real \ 'a::euclidean_space" assumes periodic: "\x. f(x + (b-a)) = f x" and f: "has_bochner_integral (lebesgue_on {a..a+c}) f i" shows "has_bochner_integral (lebesgue_on {b..b+c}) f i" proof - have "f(x + a-b) = f x" for x using periodic [of "x + (a-b)"] by (simp add: algebra_simps) then show ?thesis using has_integral_offset [OF f, of "a-b"] by (auto simp: algebra_simps) qed lemma has_integral_periodic_offset_pos: fixes f :: "real \ real" assumes f: "has_bochner_integral (lebesgue_on {a..b}) f i" and periodic: "\x. f(x + (b-a)) = f x" and c: "c \ 0" "a + c \ b" shows "has_bochner_integral (lebesgue_on {a..b}) (\x. f(x + c)) i" proof - have "{a..a + c} \ {a..b}" by (simp add: assms(4)) then have 1: "has_bochner_integral (lebesgue_on {a..a+c}) f(integral\<^sup>L (lebesgue_on {a..a+c}) f)" - by (meson integrable_subinterval_real f has_bochner_integral_iff) + by (meson integrable_subinterval f has_bochner_integral_iff) then have 3: "has_bochner_integral (lebesgue_on {b..b+c}) f(integral\<^sup>L (lebesgue_on {a..a+c}) f)" using has_integral_periodic_offset_lemma periodic by blast have "{a + c..b} \ {a..b}" by (simp add: c) then have 2: "has_bochner_integral (lebesgue_on {a+c..b}) f(integral\<^sup>L (lebesgue_on {a+c..b}) f)" - by (meson integrable_subinterval_real f has_bochner_integral_integrable integrable.intros) + by (meson integrable_subinterval f has_bochner_integral_integrable integrable.intros) have "integral\<^sup>L (lebesgue_on {a + c..b}) f + integral\<^sup>L (lebesgue_on {a..a + c}) f = i" by (metis integral_combine add.commute c f has_bochner_integral_iff le_add_same_cancel1) then have "has_bochner_integral (lebesgue_on {a+c..b+c}) f i" using has_bochner_integral_combine [OF _ _ 2 3] 1 2 by (simp add: c) then show ?thesis by (metis add_diff_cancel_right' has_integral_offset) qed lemma has_integral_periodic_offset_weak: fixes f :: "real \ real" assumes f: "has_bochner_integral (lebesgue_on {a..b}) f i" and periodic: "\x. f(x + (b-a)) = f x" and c: "\c\ \ b-a" shows "has_bochner_integral (lebesgue_on {a..b}) (\x. f(x + c)) i" proof (cases "c \ 0") case True then show ?thesis using c f has_integral_periodic_offset_pos periodic by auto next case False have per': "f(- (x + (- a - - b))) = f(- x)" for x using periodic [of "(a-b)-x"] by simp have f': "has_bochner_integral (lebesgue_on {- b..- a}) (\x. f(- x)) i" using f by blast show ?thesis using has_integral_periodic_offset_pos [of "-b" "-a" "\x. f(-x)" i "-c", OF f' per'] c False by (simp flip: has_bochner_integral_reflect_real [of b a]) qed lemma has_integral_periodic_offset: fixes f :: "real \ real" assumes f: "has_bochner_integral (lebesgue_on {a..b}) f i" and periodic: "\x. f(x + (b-a)) = f x" shows "has_bochner_integral (lebesgue_on {a..b}) (\x. f(x + c)) i" proof - consider "b \ a" | "a < b" by linarith then show ?thesis proof cases case 1 then have "{a..b} \ null_sets lebesgue" using less_eq_real_def by auto with f show ?thesis by auto next case 2 define fba where "fba \ \x. f(x + (b-a) * frac(c / (b-a)))" have eq: "fba x = f(x + c)" if "x \ {a..b}" for x proof - have "f(x + n * (b-a)) = f x" if "n \ \" for n x using periodic periodic_integer_multiple that by blast then have "f((x + c) + - floor (c / (b - a)) * (b - a)) = f(x + c)" using Ints_of_int by blast moreover have "((x + c) + -floor (c / (b - a)) * (b - a)) = (x + (b - a) * frac (c / (b - a)))" using 2 by (simp add: field_simps frac_def) ultimately show ?thesis unfolding fba_def by metis qed have *: "has_bochner_integral (lebesgue_on {a..b}) fba i" unfolding fba_def proof (rule has_integral_periodic_offset_weak [OF f]) show "f(x + (b - a)) = f x" for x by fact have "\frac (c / (b - a))\ \ 1" using frac_unique_iff less_eq_real_def by auto then show "\(b - a) * frac (c / (b - a))\ \ b - a" using "2" by auto qed then show ?thesis proof (rule has_bochner_integralI_AE [OF _ _ AE_I2]) have "fba \ borel_measurable (lebesgue_on {a..b})" using "*" borel_measurable_has_bochner_integral by blast then show "(\x. f(x + c)) \ borel_measurable (lebesgue_on {a..b})" by (subst measurable_lebesgue_cong [OF eq, symmetric]) qed (auto simp: eq) qed qed lemma integrable_periodic_offset: fixes f :: "real \ real" assumes f: "integrable (lebesgue_on {a..b}) f" and periodic: "\x. f(x + (b-a)) = f x" shows "integrable (lebesgue_on {a..b}) (\x. f(x + c))" using f has_integral_periodic_offset periodic by (simp add: has_bochner_integral_iff) lemma absolutely_integrable_periodic_offset: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on {a..b}" and periodic: "\x. f(x + (b-a)) = f x" shows "(\x. f(x + c)) absolutely_integrable_on {a..b}" "(\x. f(c + x)) absolutely_integrable_on {a..b}" using assms integrable_periodic_offset [of a b "f"] by (auto simp: integrable_restrict_space set_integrable_def add.commute [of c]) lemma integral_periodic_offset: fixes f :: "real \ real" assumes periodic: "\x. f(x + (b-a)) = f x" shows "integral\<^sup>L (lebesgue_on {a..b}) (\x. f(x + c)) = integral\<^sup>L (lebesgue_on {a..b}) f" proof (cases "integrable (lebesgue_on {a..b}) f") case True then show ?thesis using has_integral_periodic_offset periodic by (simp add: has_bochner_integral_iff) next case False then have "\ integrable (lebesgue_on {a..b}) (\x. f(x + c))" using periodic[of "_+c"] by (auto simp: algebra_simps intro: dest: integrable_periodic_offset [where c = "-c"]) then have "integral\<^sup>L (lebesgue_on {a..b}) f = 0" "integral\<^sup>L (lebesgue_on {a..b}) (\x. f(x + c)) = 0" using False not_integrable_integral_eq by blast+ then show ?thesis by simp qed end diff --git a/thys/Fourier/Square_Integrable.thy b/thys/Fourier/Square_Integrable.thy --- a/thys/Fourier/Square_Integrable.thy +++ b/thys/Fourier/Square_Integrable.thy @@ -1,849 +1,824 @@ section\Square integrable functions over the reals\ theory Square_Integrable - imports Lspace Fourier_Aux2 + imports Lspace begin subsection\Basic definitions\ definition square_integrable:: "(real \ real) \ real set \ bool" (infixr "square'_integrable" 46) where "f square_integrable S \ S \ sets lebesgue \ f \ borel_measurable (lebesgue_on S) \ integrable (lebesgue_on S) (\x. f x ^ 2)" lemma square_integrable_imp_measurable: "f square_integrable S \ f \ borel_measurable (lebesgue_on S)" by (simp add: square_integrable_def) lemma square_integrable_imp_lebesgue: "f square_integrable S \ S \ sets lebesgue" by (simp add: square_integrable_def) lemma square_integrable_imp_lspace: assumes "f square_integrable S" shows "f \ lspace (lebesgue_on S) 2" proof - have "(\x. (f x)\<^sup>2) absolutely_integrable_on S" by (metis assms integrable_on_lebesgue_on nonnegative_absolutely_integrable_1 square_integrable_def zero_le_power2) moreover have "S \ sets lebesgue" using assms square_integrable_def by blast ultimately show ?thesis by (simp add: assms Lp_space_numeral integrable_restrict_space set_integrable_def square_integrable_imp_measurable) qed lemma square_integrable_iff_lspace: assumes "S \ sets lebesgue" shows "f square_integrable S \ f \ lspace (lebesgue_on S) 2" (is "?lhs = ?rhs") proof assume L: ?lhs then show ?rhs using square_integrable_imp_lspace by blast next assume ?rhs then show ?lhs using assms by (auto simp: Lp_space_numeral square_integrable_def integrable_on_lebesgue_on) qed lemma square_integrable_0 [simp]: "S \ sets lebesgue \ (\x. 0) square_integrable S" by (simp add: square_integrable_def power2_eq_square integrable_0) lemma square_integrable_neg_eq [simp]: "(\x. -(f x)) square_integrable S \ f square_integrable S" by (auto simp: square_integrable_def) lemma square_integrable_lmult [simp]: assumes "f square_integrable S" shows "(\x. c * f x) square_integrable S" proof (simp add: square_integrable_def, intro conjI) have f: "f \ borel_measurable (lebesgue_on S)" "integrable (lebesgue_on S) (\x. f x ^ 2)" using assms by (simp_all add: square_integrable_def) then show "(\x. c * f x) \ borel_measurable (lebesgue_on S)" using borel_measurable_scaleR [of "\x. c" "lebesgue_on S" f] by simp have "integrable (lebesgue_on S) (\x. c\<^sup>2 * (f x)\<^sup>2)" by (cases "c=0") (auto simp: f integrable_0) then show "integrable (lebesgue_on S) (\x. (c * f x)\<^sup>2)" by (simp add: power2_eq_square mult_ac) show "S \ sets lebesgue" using assms square_integrable_def by blast qed lemma square_integrable_rmult [simp]: "f square_integrable S \ (\x. f x * c) square_integrable S" using square_integrable_lmult [of f S c] by (simp add: mult.commute) lemma square_integrable_imp_absolutely_integrable_product: assumes f: "f square_integrable S" and g: "g square_integrable S" shows "(\x. f x * g x) absolutely_integrable_on S" proof - have fS: "integrable (lebesgue_on S) (\r. (f r)\<^sup>2)" "integrable (lebesgue_on S) (\r. (g r)\<^sup>2)" using assms square_integrable_def by blast+ have "integrable (lebesgue_on S) (\x. \f x * g x\)" proof (intro integrable_abs Holder_inequality [of 2 2]) show "f \ borel_measurable (lebesgue_on S)" "g \ borel_measurable (lebesgue_on S)" using f g square_integrable_def by blast+ show "integrable (lebesgue_on S) (\x. \f x\ powr 2)" "integrable (lebesgue_on S) (\x. \g x\ powr 2)" using nonnegative_absolutely_integrable_1 [of "(\x. (f x)\<^sup>2)"] nonnegative_absolutely_integrable_1 [of "(\x. (g x)\<^sup>2)"] by (simp_all add: fS integrable_restrict_space set_integrable_def) qed auto then show ?thesis using assms by (simp add: absolutely_integrable_measurable_real borel_measurable_times square_integrable_def) qed lemma square_integrable_imp_integrable_product: assumes "f square_integrable S" "g square_integrable S" shows "integrable (lebesgue_on S) (\x. f x * g x)" using absolutely_integrable_measurable assms integrable_abs_iff by (metis (full_types) absolutely_integrable_measurable_real square_integrable_def square_integrable_imp_absolutely_integrable_product) lemma square_integrable_add [simp]: assumes f: "f square_integrable S" and g: "g square_integrable S" shows "(\x. f x + g x) square_integrable S" unfolding square_integrable_def proof (intro conjI) show "S \ sets lebesgue" using assms square_integrable_def by blast show "(\x. f x + g x) \ borel_measurable (lebesgue_on S)" by (simp add: f g borel_measurable_add square_integrable_imp_measurable) show "integrable (lebesgue_on S) (\x. (f x + g x)\<^sup>2)" unfolding power2_eq_square distrib_right distrib_left proof (intro Bochner_Integration.integrable_add) show "integrable (lebesgue_on S) (\x. f x * f x)" "integrable (lebesgue_on S) (\x. g x * g x)" using f g square_integrable_imp_integrable_product by blast+ show "integrable (lebesgue_on S) (\x. f x * g x)" "integrable (lebesgue_on S) (\x. g x * f x)" using f g square_integrable_imp_integrable_product by blast+ qed qed lemma square_integrable_diff [simp]: "\f square_integrable S; g square_integrable S\ \ (\x. f x - g x) square_integrable S" using square_integrable_neg_eq square_integrable_add [of f S "\x. - (g x)"] by auto lemma square_integrable_abs [simp]: "f square_integrable S \ (\x. \f x\) square_integrable S" by (simp add: square_integrable_def borel_measurable_abs) lemma square_integrable_sum [simp]: assumes I: "finite I" "\i. i \ I \ f i square_integrable S" and S: "S \ sets lebesgue" shows "(\x. \i\I. f i x) square_integrable S" using I by induction (simp_all add: S) lemma continuous_imp_square_integrable [simp]: "continuous_on {a..b} f \ f square_integrable {a..b}" using continuous_imp_integrable [of a b "(\x. (f x)\<^sup>2)"] by (simp add: square_integrable_def continuous_on_power continuous_imp_measurable_on_sets_lebesgue) lemma square_integrable_imp_absolutely_integrable: assumes f: "f square_integrable S" and S: "S \ lmeasurable" shows "f absolutely_integrable_on S" proof - have "f \ lspace (lebesgue_on S) 2" using f S square_integrable_iff_lspace by blast then have "f \ lspace (lebesgue_on S) 1" by (rule lspace_mono) (use S in auto) then show ?thesis using S by (simp flip: lspace_1) qed lemma square_integrable_imp_integrable: assumes f: "f square_integrable S" and S: "S \ lmeasurable" shows "integrable (lebesgue_on S) f" by (meson S absolutely_integrable_measurable_real f fmeasurableD integrable_abs_iff square_integrable_imp_absolutely_integrable) subsection\ The norm and inner product in L2\ definition l2product :: "'a::euclidean_space set \ ('a \ real) \ ('a \ real) \ real" where "l2product S f g \ (\x. f x * g x \(lebesgue_on S))" definition l2norm :: "['a::euclidean_space set, 'a \ real] \ real" where "l2norm S f \ sqrt (l2product S f f)" definition lnorm :: "['a measure, real, 'a \ real] \ real" where "lnorm M p f \ (\x. \f x\ powr p \M) powr (1/p)" corollary Holder_inequality_lnorm: assumes "p > (0::real)" "q > 0" "1/p+1/q = 1" and "f \ borel_measurable M" "g \ borel_measurable M" "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr q)" shows "(\x. \f x * g x\ \M) \ lnorm M p f * lnorm M q g" "\\x. f x * g x \M \ \ lnorm M p f * lnorm M q g" by (simp_all add: Holder_inequality assms lnorm_def) lemma l2norm_lnorm: "l2norm S f = lnorm (lebesgue_on S) 2 f" proof - have "(LINT x|lebesgue_on S. (f x)\<^sup>2) \ 0" by simp then show ?thesis by (auto simp: lnorm_def l2norm_def l2product_def power2_eq_square powr_half_sqrt) qed lemma lnorm_nonneg: "lnorm M p f \ 0" by (simp add: lnorm_def) lemma lnorm_minus_commute: "lnorm M p (g - f) = lnorm M p (f - g)" by (simp add: lnorm_def abs_minus_commute) text\ Extending a continuous function in a periodic way\ -lemma isCont_id [continuous_intros,iff]: - fixes x :: "'a::t2_space" shows "isCont id x" - unfolding id_def by (intro continuous_intros) - -lemma continuous_floor: - fixes x::real - shows "x \ \ \ continuous (at x) (real_of_int \ floor)" - using floor_has_real_derivative [where f=id] - by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) - -lemma continuous_frac: - fixes x::real - assumes "x \ \" - shows "continuous (at x) frac" -proof - - have "isCont (\x. real_of_int \x\) x" - using continuous_floor [OF assms] by (simp add: o_def) - then have *: "continuous (at x) (\x. x - real_of_int \x\)" - by (intro continuous_intros) - moreover have "\\<^sub>F x in nhds x. frac x = x - real_of_int \x\" - by (simp add: frac_def) - ultimately show ?thesis - by (simp add: LIM_imp_LIM frac_def isCont_def) -qed - proposition continuous_on_compose_frac: fixes f:: "real \ real" assumes contf: "continuous_on {0..1} f" and f10: "f 1 = f 0" shows "continuous_on UNIV (f \ frac)" proof - have *: "isCont (f \ frac) x" if caf: "\x. \0 \ x; x \ 1\ \ continuous (at x within {0..1}) f" for x proof (cases "x \ \") case True then have [simp]: "frac x = 0" by simp show ?thesis proof (clarsimp simp add: continuous_at_eps_delta dist_real_def) have f0: "continuous (at 0 within {0..1}) f" and f1: "continuous (at 1 within {0..1}) f" by (auto intro: caf) show "\d>0. \x'. \x'-x\ < d \ \f(frac x') - f 0\ < e" if "0 < e" for e proof - obtain d0 where "d0 > 0" and d0: "\x'. \x'\{0..1}; \x'\ < d0\ \ \f x' - f 0\ < e" using \e > 0\ caf [of 0] dist_not_less_zero by (auto simp: continuous_within_eps_delta dist_real_def) obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; \x' - 1\ < d1\ \ \f x' - f 0\ < e" using \e > 0\ caf [of 1] dist_not_less_zero f10 by (auto simp: continuous_within_eps_delta dist_real_def) show ?thesis proof (intro exI conjI allI impI) show "0 < min 1 (min d0 d1)" by (auto simp: \d0 > 0\ \d1 > 0\) show "\f(frac x') - f 0\ < e" if "\x'-x\ < min 1 (min d0 d1)" for x' proof (cases "x \ x'") case True with \x \ \\ that have "frac x' = x' - x" by (simp add: frac_unique_iff) then show ?thesis using True d0 that by auto next case False then have [simp]: "frac x' = 1 - (x - x')" using that \x \ \\ by (simp add: not_le frac_unique_iff) show ?thesis using False d1 that by auto qed qed qed qed next case False show ?thesis proof (rule continuous_at_compose) show "isCont frac x" by (simp add: False continuous_frac) have "frac x \ {0<..<1}" by (simp add: False frac_lt_1) then show "isCont f(frac x)" by (metis at_within_Icc_at greaterThanLessThan_iff le_cases not_le that) qed qed then show ?thesis using contf by (simp add: o_def continuous_on_eq_continuous_within) qed proposition Tietze_periodic_interval: fixes f:: "real \ real" assumes contf: "continuous_on {a..b} f" and fab: "f a = f b" obtains g where "continuous_on UNIV g" "\x. x \ {a..b} \ g x = f x" "\x. g(x + (b-a)) = g x" proof (cases "a < b") case True let ?g = "f \ (\y. a + (b-a) * y) \ frac \ (\x. (x - a) / (b-a))" show ?thesis proof have "a + (b - a) * y \ b" if "a < b" "0 \ y" "y \ 1" for y using that affine_ineq by (force simp: field_simps) then have *: "continuous_on (range (\x. (x - a) / (b - a))) (f \ (\y. a + (b - a) * y) \ frac)" apply (intro continuous_on_subset [OF continuous_on_compose_frac] continuous_on_subset [OF contf] continuous_intros) using \a < b\ by (auto simp: fab) show "continuous_on UNIV ?g" by (intro * continuous_on_compose continuous_intros) (use True in auto) show "?g x = f x" if "x \ {a..b}" for x :: real proof (cases "x=b") case True then show ?thesis by (auto simp: frac_def intro: fab) next case False with \a < b\ that have "frac ((x - a) / (b - a)) = (x - a) / (b - a)" by (subst frac_eq) (auto simp: divide_simps) with \a < b\ show ?thesis by auto qed have "a + (b-a) * frac ((x + b - 2 * a) / (b-a)) = a + (b-a) * frac ((x - a) / (b-a))" for x using True frac_1_eq [of "(x - a) / (b-a)"] by (auto simp: divide_simps) then show "?g (x + (b-a)) = (?g x::real)" for x by force qed next case False show ?thesis proof show "f a = f x" if "x \ {a..b}" for x using that False order_trans by fastforce qed auto qed subsection\Lspace stuff\ lemma eNorm_triangle_eps: assumes "eNorm N (x' - x) < a" "defect N = 1" obtains e where "e > 0" "\y. eNorm N (y - x') < e \ eNorm N (y - x) < a" proof - let ?d = "a - Norm N (x' - x)" have nt: "eNorm N (x' - x) < \" using assms top.not_eq_extremum by fastforce with assms have d: "?d > 0" by (simp add: Norm_def diff_gr0_ennreal) have [simp]: "ennreal (1 - Norm N (x' - x)) = 1 - eNorm N (x' - x)" using that nt unfolding Norm_def by (metis enn2real_nonneg ennreal_1 ennreal_enn2real ennreal_minus) show ?thesis proof show "(0::ennreal) < ?d" using d ennreal_less_zero_iff by blast show "eNorm N (y - x) < a" if "eNorm N (y - x') < ?d" for y using that assms eNorm_triangular_ineq [of N "y - x'" "x' - x"] le_less_trans less_diff_eq_ennreal by (simp add: Norm_def nt) qed qed lemma topspace_topology\<^sub>N [simp]: assumes "defect N = 1" shows "topspace (topology\<^sub>N N) = UNIV" proof - have "x \ topspace (topology\<^sub>N N)" for x proof - have "\e>0. \y. eNorm N (y - x') < e \ eNorm N (y - x) < 1" if "eNorm N (x' - x) < 1" for x' using eNorm_triangle_eps by (metis assms that) then show ?thesis unfolding topspace_def by (rule_tac X="{y. eNorm N (y - x) < 1}" in UnionI) (auto intro: openin_topology\<^sub>N_I) qed then show ?thesis by auto qed lemma tendsto_ine\<^sub>N_iff_limitin: assumes "defect N = 1" shows "tendsto_ine\<^sub>N N u x = limitin (topology\<^sub>N N) u x sequentially" proof - have "\\<^sub>F x in sequentially. u x \ U" if 0: "(\n. eNorm N (u n - x)) \ 0" and U: "openin (topology\<^sub>N N) U" "x \ U" for U proof - obtain e where "e > 0" and e: "\y. eNorm N (y-x) < e \ y \ U" using openin_topology\<^sub>N_D U by metis then show ?thesis using eventually_mono order_tendstoD(2)[OF 0] by force qed moreover have "(\n. eNorm N (u n - x)) \ 0" if x: "x \ topspace (topology\<^sub>N N)" and *: "\U. \openin (topology\<^sub>N N) U; x \ U\ \ (\\<^sub>F x in sequentially. u x \ U)" proof (rule order_tendstoI) show "\\<^sub>F n in sequentially. eNorm N (u n - x) < a" if "a > 0" for a apply (rule * [OF openin_topology\<^sub>N_I, of "{v. eNorm N (v - x) < a}", simplified]) using assms eNorm_triangle_eps that apply blast+ done qed simp ultimately show ?thesis by (auto simp: tendsto_ine\<^sub>N_def limitin_def assms) qed corollary tendsto_ine\<^sub>N_iff_limitin_ge1: fixes p :: ennreal assumes "p \ 1" shows "tendsto_ine\<^sub>N (\ p M) u x = limitin (topology\<^sub>N (\ p M)) u x sequentially" proof (rule tendsto_ine\<^sub>N_iff_limitin) show "defect (\ p M) = 1" by (metis (full_types) L_infinity(2) L_zero(2) Lp(2) Lp_cases assms ennreal_ge_1) qed corollary tendsto_in\<^sub>N_iff_limitin: assumes "defect N = 1" "x \ space\<^sub>N N" "\n. u n \ space\<^sub>N N" shows "tendsto_in\<^sub>N N u x = limitin (topology\<^sub>N N) u x sequentially" using assms tendsto_ine\<^sub>N_iff_limitin tendsto_ine_in by blast corollary tendsto_in\<^sub>N_iff_limitin_ge1: fixes p :: ennreal assumes "p \ 1" "x \ lspace M p" "\n. u n \ lspace M p" shows "tendsto_in\<^sub>N (\ p M) u x = limitin (topology\<^sub>N (\ p M)) u x sequentially" proof (rule tendsto_in\<^sub>N_iff_limitin) show "defect (\ p M) = 1" by (metis (full_types) L_infinity(2) L_zero(2) Lp(2) Lp_cases \p \ 1\ ennreal_ge_1) qed (auto simp: assms) lemma l2product_sym: "l2product S f g = l2product S g f" by (simp add: l2product_def mult.commute) lemma l2product_pos_le: "f square_integrable S \ 0 \ l2product S f f" by (simp add: square_integrable_def l2product_def flip: power2_eq_square) lemma l2norm_pow_2: "f square_integrable S \ (l2norm S f) ^ 2 = l2product S f f" by (simp add: l2norm_def l2product_pos_le) lemma l2norm_pos_le: "f square_integrable S \ 0 \ l2norm S f" by (simp add: l2norm_def l2product_pos_le) lemma l2norm_le: "(l2norm S f \ l2norm S g \ l2product S f f \ l2product S g g)" by (simp add: l2norm_def) lemma l2norm_eq: "(l2norm S f = l2norm S g \ l2product S f f = l2product S g g)" by (simp add: l2norm_def) lemma Schwartz_inequality_strong: assumes "f square_integrable S" "g square_integrable S" shows "l2product S (\x. \f x\) (\x. \g x\) \ l2norm S f * l2norm S g" using Holder_inequality_lnorm [of 2 2 f "lebesgue_on S" g] assms by (simp add: square_integrable_def l2product_def abs_mult flip: l2norm_lnorm) lemma Schwartz_inequality_abs: assumes "f square_integrable S" "g square_integrable S" shows "\l2product S f g\ \ l2norm S f * l2norm S g" proof - have "\l2product S f g\ \ l2product S (\x. \f x\) (\x. \g x\)" unfolding l2product_def proof (rule integral_abs_bound_integral) show "integrable (lebesgue_on S) (\x. f x * g x)" "integrable (lebesgue_on S) (\x. \f x\ * \g x\)" by (simp_all add: assms square_integrable_imp_integrable_product) qed (simp add: abs_mult) also have "\ \ l2norm S f * l2norm S g" by (simp add: Schwartz_inequality_strong assms) finally show ?thesis . qed lemma Schwartz_inequality: assumes "f square_integrable S" "g square_integrable S" shows "l2product S f g \ l2norm S f * l2norm S g" using Schwartz_inequality_abs assms by fastforce lemma lnorm_triangle: assumes f: "f \ lspace M p" and g: "g \ lspace M p" and "p \ 1" shows "lnorm M p (\x. f x + g x) \ lnorm M p f + lnorm M p g" proof - have "p > 0" using assms by linarith then have "integrable M (\x. \f x\ powr p)" "integrable M (\x. \g x\ powr p)" by (simp_all add: Lp_D(2) assms) moreover have "f \ borel_measurable M" "g \ borel_measurable M" using Lp_measurable f g by blast+ ultimately show ?thesis unfolding lnorm_def using Minkowski_inequality(2) \p \ 1\ by blast qed lemma lnorm_triangle_fun: assumes f: "f \ lspace M p" and g: "g \ lspace M p" and "p \ 1" shows "lnorm M p (f + g) \ lnorm M p f + lnorm M p g" using lnorm_triangle [OF assms] by (simp add: plus_fun_def) lemma l2norm_triangle: assumes "f square_integrable S" "g square_integrable S" shows "l2norm S (\x. f x + g x) \ l2norm S f + l2norm S g" proof - have "f \ lspace (lebesgue_on S) 2" "g \ lspace (lebesgue_on S) 2" using assms by (simp_all add: square_integrable_imp_lspace) then show ?thesis using lnorm_triangle [of f 2 "lebesgue_on S"] by (simp add: l2norm_lnorm) qed lemma l2product_ladd: "\f square_integrable S; g square_integrable S; h square_integrable S\ \ l2product S (\x. f x + g x) h = l2product S f h + l2product S g h" by (simp add: l2product_def algebra_simps square_integrable_imp_integrable_product) lemma l2product_radd: "\f square_integrable S; g square_integrable S; h square_integrable S\ \ l2product S f(\x. g x + h x) = l2product S f g + l2product S f h" by (simp add: l2product_def algebra_simps square_integrable_imp_integrable_product) lemma l2product_ldiff: "\f square_integrable S; g square_integrable S; h square_integrable S\ \ l2product S (\x. f x - g x) h = l2product S f h - l2product S g h" by (simp add: l2product_def algebra_simps square_integrable_imp_integrable_product) lemma l2product_rdiff: "\f square_integrable S; g square_integrable S; h square_integrable S\ \ l2product S f(\x. g x - h x) = l2product S f g - l2product S f h" by (simp add: l2product_def algebra_simps square_integrable_imp_integrable_product) lemma l2product_lmult: "\f square_integrable S; g square_integrable S\ \ l2product S (\x. c * f x) g = c * l2product S f g" by (simp add: l2product_def algebra_simps) lemma l2product_rmult: "\f square_integrable S; g square_integrable S\ \ l2product S f(\x. c * g x) = c * l2product S f g" by (simp add: l2product_def algebra_simps) lemma l2product_lzero [simp]: "l2product S (\x. 0) f = 0" by (simp add: l2product_def) lemma l2product_rzero [simp]: "l2product S f(\x. 0) = 0" by (simp add: l2product_def) lemma l2product_lsum: assumes I: "finite I" "\i. i \ I \ (f i) square_integrable S" and S: "g square_integrable S" shows "l2product S (\x. \i\I. f i x) g = (\i\I. l2product S (f i) g)" using I proof induction case (insert i I) with S show ?case by (simp add: l2product_ladd square_integrable_imp_lebesgue) qed auto lemma l2product_rsum: assumes I: "finite I" "\i. i \ I \ (f i) square_integrable S" and S: "g square_integrable S" shows "l2product S g (\x. \i\I. f i x) = (\i\I. l2product S g (f i))" using l2product_lsum [OF assms] by (simp add: l2product_sym) lemma l2norm_lmult: "f square_integrable S \ l2norm S (\x. c * f x) = \c\ * l2norm S f" by (simp add: l2norm_def l2product_rmult l2product_sym real_sqrt_mult) lemma l2norm_rmult: "f square_integrable S \ l2norm S (\x. f x * c) = l2norm S f * \c\" using l2norm_lmult by (simp add: mult.commute) lemma l2norm_neg: "f square_integrable S \ l2norm S (\x. - f x) = l2norm S f" using l2norm_lmult [of f S "-1"] by simp lemma l2norm_diff: assumes "f square_integrable S" "g square_integrable S" shows "l2norm S (\x. f x - g x) = l2norm S (\x. g x - f x)" proof - have "(\x. f x - g x) square_integrable S" using assms square_integrable_diff by blast then show ?thesis using l2norm_neg [of "\x. f x - g x" S] by (simp add: algebra_simps) qed subsection\Completeness (Riesz-Fischer)\ lemma eNorm_eq_lnorm: "\f \ lspace M p; p > 0\ \ eNorm (\ (ennreal p) M) f = ennreal (lnorm M p f)" by (simp add: Lp_D(4) lnorm_def) lemma Norm_eq_lnorm: "\f \ lspace M p; p > 0\ \ Norm (\ (ennreal p) M) f = lnorm M p f" by (simp add: Lp_D(3) lnorm_def) lemma eNorm_ge1_triangular_ineq: assumes "p \ (1::real)" shows "eNorm (\ p M) (x + y) \ eNorm (\ p M) x + eNorm (\ p M) y" using eNorm_triangular_ineq [of "(\ p M)"] assms by (simp add: Lp(2)) text\A mere repackaging of the theorem @{thm Lp_complete}, but nearly as much work again.\ proposition l2_complete: assumes f: "\i::nat. f i square_integrable S" and cauchy: "\e. 0 < e \ \N. \m\N. \n\N. l2norm S (\x. f m x - f n x) < e" obtains g where "g square_integrable S" "((\n. l2norm S (\x. f n x - g x)) \ 0)" proof - have finite: "eNorm (\ 2 (lebesgue_on S)) (f n - f m) < \" for m n by (metis f infinity_ennreal_def spaceN_diff spaceN_iff square_integrable_imp_lspace) have *: "cauchy_ine\<^sub>N (\ 2 (lebesgue_on S)) f" proof (clarsimp simp: cauchy_ine\<^sub>N_def) show "\M. \n\M. \m\M. eNorm (\ 2 (lebesgue_on S)) (f n - f m) < e" if "e > 0" for e proof (cases e) case (real r) then have "r > 0" using that by auto with cauchy obtain N::nat where N: "\m n. \m \ N; n \ N\ \ l2norm S (\x. f n x - f m x) < r" by blast show ?thesis proof (intro exI allI impI) show "eNorm (\ 2 (lebesgue_on S)) (f n - f m) < e" if "N \ m" "N \ n" for m n proof - have fnm: "(f n - f m) \ borel_measurable (lebesgue_on S)" using f unfolding square_integrable_def by (blast intro: borel_measurable_diff') have "l2norm S (\x. f n x - f m x) = lnorm (lebesgue_on S) 2 (\x. f n x - f m x)" by (metis l2norm_lnorm) also have "\ = Norm (\ 2 (lebesgue_on S)) (f n - f m)" using Lp_Norm [OF _ fnm, of 2] by (simp add: lnorm_def) finally show ?thesis using N [OF that] real finite by (simp add: Norm_def) qed qed qed (simp add: finite) qed then obtain g where g: "tendsto_ine\<^sub>N (\ 2 (lebesgue_on S)) f g" using Lp_complete complete\<^sub>N_def by blast show ?thesis proof have fng_to_0: "(\n. eNorm (\ 2 (lebesgue_on S)) (\x. f n x - g x)) \ 0" using g Lp_D(4) [of 2 _ "lebesgue_on S"] by (simp add: tendsto_ine\<^sub>N_def minus_fun_def) then obtain M where "\n . n \ M \ eNorm (\ 2 (lebesgue_on S)) (\x. f n x - g x) < \" apply (simp add: lim_explicit) by (metis (full_types) open_lessThan diff_self eNorm_zero lessThan_iff local.finite) then have "eNorm (\ 2 (lebesgue_on S)) (\x. g x - f M x) < \" using eNorm_uminus [of _ "\x. g x - f _ x"] by (simp add: uminus_fun_def) moreover have "eNorm (\ 2 (lebesgue_on S)) (\x. f M x) < \" using f square_integrable_imp_lspace by (simp add: spaceN_iff) ultimately have "eNorm (\ 2 (lebesgue_on S)) g < \" using eNorm_ge1_triangular_ineq [of 2 "lebesgue_on S" "g - f M" "f M", simplified] not_le top.not_eq_extremum by (fastforce simp add: minus_fun_def) then have g_space: "g \ space\<^sub>N (\ 2 (lebesgue_on S))" by (simp add: spaceN_iff) show "g square_integrable S" unfolding square_integrable_def proof (intro conjI) show "g \ borel_measurable (lebesgue_on S)" using Lp_measurable g_space by blast show "S \ sets lebesgue" using f square_integrable_def by blast then show "integrable (lebesgue_on S) (\x. (g x)\<^sup>2)" using g_space square_integrable_def square_integrable_iff_lspace by blast qed then have "f n - g \ lspace (lebesgue_on S) 2" for n using f spaceN_diff square_integrable_imp_lspace by blast with fng_to_0 have "(\n. ennreal (lnorm (lebesgue_on S) 2 (\x. f n x - g x))) \ 0" by (simp add: minus_fun_def flip: eNorm_eq_lnorm) then have "(\n. lnorm (lebesgue_on S) 2 (\x. f n x - g x)) \ 0" by (simp add: ennreal_tendsto_0_iff lnorm_def) then show "(\n. l2norm S (\x. f n x - g x)) \ 0" using g by (simp add: l2norm_lnorm lnorm_def) qed qed subsection\Approximation of functions in Lp by bounded and continuous ones\ lemma lspace_bounded_measurable: fixes p::real assumes f: "f \ borel_measurable (lebesgue_on S)" and g: "g \ lspace (lebesgue_on S) p" and "p > 0" and le: " AE x in lebesgue_on S. norm (\f x\ powr p) \ norm (\g x\ powr p)" shows "f \ lspace (lebesgue_on S) p" using assms by (auto simp: lspace_ennreal_iff intro: Bochner_Integration.integrable_bound) lemma lspace_approximate_bounded: assumes f: "f \ lspace (lebesgue_on S) p" and S: "S \ lmeasurable" and "p > 0" "e > 0" obtains g where "g \ lspace (lebesgue_on S) p" "bounded (g ` S)" "lnorm (lebesgue_on S) p (f - g) < e" proof - have f_bm: "f \ borel_measurable (lebesgue_on S)" using Lp_measurable f by blast let ?f = "\n::nat. \x. max (- n) (min n (f x))" have "tendsto_in\<^sub>N (\ p (lebesgue_on S)) ?f f" proof (rule Lp_domination_limit) show "\n::nat. ?f n \ borel_measurable (lebesgue_on S)" by (intro f_bm borel_measurable_max borel_measurable_min borel_measurable_const) show "abs \ f \ lspace (lebesgue_on S) p" using Lp_Banach_lattice [OF f] by (simp add: o_def) have *: "\\<^sub>F n in sequentially. dist (?f n x) (f x) < e" if x: "x \ space (lebesgue_on S)" and "e > 0" for x e proof show "dist (?f n x) (f x) < e" if "nat \\f x\\ \ n" for n :: nat using that \0 < e\ by (simp add: dist_real_def max_def min_def abs_if split: if_split_asm) qed then show "AE x in lebesgue_on S. (\n::nat. max (- n) (min n (f x))) \ f x" by (blast intro: tendstoI) qed (auto simp: f_bm) moreover have lspace: "?f n \ lspace (lebesgue_on S) p" for n::nat by (intro f lspace_const lspace_min lspace_max \p > 0\ S) ultimately have "(\n. lnorm (lebesgue_on S) p (?f n - f)) \ 0" by (simp add: tendsto_in\<^sub>N_def Norm_eq_lnorm \p > 0\ f) with \e > 0\ obtain N where N: "\lnorm (lebesgue_on S) p (?f N - f)\ < e" by (auto simp: LIMSEQ_iff) show ?thesis proof have "\x\S. \max (- real N) (min (real N) (f x))\ \ N" by auto then show "bounded (?f N ` S::real set)" by (force simp: bounded_iff) show "lnorm (lebesgue_on S) p (f - ?f N) < e" using N by (simp add: lnorm_minus_commute) qed (auto simp: lspace) qed lemma borel_measurable_imp_continuous_limit: fixes h :: "'a::euclidean_space \ 'b::euclidean_space" assumes h: "h \ borel_measurable (lebesgue_on S)" and S: "S \ sets lebesgue" obtains g where "\n. continuous_on UNIV (g n)" "AE x in lebesgue_on S. (\n::nat. g n x) \ h x" proof - have "h measurable_on S" using S h measurable_on_iff_borel_measurable by blast then obtain N g where N: "N \ null_sets lebesgue" and g: "\n. continuous_on UNIV (g n)" and tends: "\x. x \ N \ (\n. g n x) \ (if x \ S then h x else 0)" by (auto simp: measurable_on_def negligible_iff_null_sets) moreover have "AE x in lebesgue_on S. (\n::nat. g n x) \ h x" proof (rule AE_I') show "N \ S \ null_sets (lebesgue_on S)" by (simp add: S N null_set_Int2 null_sets_restrict_space) show "{x \ space (lebesgue_on S). \ (\n. g n x) \ h x} \ N \ S" using tends by force qed ultimately show thesis using that by blast qed proposition lspace_approximate_continuous: assumes f: "f \ lspace (lebesgue_on S) p" and S: "S \ lmeasurable" and "1 \ p" "e > 0" obtains g where "continuous_on UNIV g" "g \ lspace (lebesgue_on S) p" "lnorm (lebesgue_on S) p (f - g) < e" proof - have "p > 0" using assms by simp obtain h where h: "h \ lspace (lebesgue_on S) p" and "bounded (h ` S)" and lesse2: "lnorm (lebesgue_on S) p (f - h) < e/2" by (rule lspace_approximate_bounded [of f p S "e/2"]) (use assms in auto) then obtain B where "B > 0" and B: "\x. x \ S \ \h x\ \ B" by (auto simp: bounded_pos) have bmh: "h \ borel_measurable (lebesgue_on S)" using h lspace_ennreal_iff [of p] \p \ 1\ by auto obtain g where contg: "\n. continuous_on UNIV (g n)" and gle: "\n x. x \ S \ \g n x\ \ B" and tends: "AE x in lebesgue_on S. (\n::nat. g n x) \ h x" proof - obtain \ where cont: "\n. continuous_on UNIV (\ n)" and tends: "AE x in lebesgue_on S. (\n::nat. \ n x) \ h x" using borel_measurable_imp_continuous_limit S bmh by blast let ?g = "\n::nat. \x. max (- B) (min B (\ n x))" show thesis proof show "continuous_on UNIV (?g n)" for n by (intro continuous_intros cont) show "\?g n x\ \ B" if "x \ S" for n x using that \B > 0\ by (auto simp: max_def min_def) have "(\n. max (- B) (min B (\ n x))) \ h x" if "(\n. \ n x) \ h x" "x \ S" for x using that \B > 0\ B [OF \x \ S\] unfolding LIMSEQ_def by (fastforce simp: min_def max_def dist_real_def) then show "AE x in lebesgue_on S. (\n. ?g n x) \ h x" using tends by auto qed qed have lspace_B: "(\x. B) \ lspace (lebesgue_on S) p" by (simp add: S \0 < p\ lspace_const) have lspace_g: "g n \ lspace (lebesgue_on S) p" for n proof (rule lspace_bounded_measurable) show "g n \ borel_measurable (lebesgue_on S)" by (simp add: borel_measurable_continuous_onI contg measurable_completion measurable_restrict_space1) show "AE x in lebesgue_on S. norm (\g n x\ powr p) \ norm (\B\ powr p)" using \B > 0\ gle S \0 < p\ powr_mono2 by auto qed (use \p > 0\ lspace_B in auto) have "tendsto_in\<^sub>N (\ p (lebesgue_on S)) g h" proof (rule Lp_domination_limit [OF bmh _ lspace_B tends]) show "\n::nat. g n \ borel_measurable (lebesgue_on S)" using Lp_measurable lspace_g by blast show "\n. AE x in lebesgue_on S. \g n x\ \ B" using S gle by auto qed then have 0: "(\n. Norm (\ p (lebesgue_on S)) (g n - h)) \ 0" by (simp add: tendsto_in\<^sub>N_def) have "\e. 0 < e \ \N. \n\N. lnorm (lebesgue_on S) p (g n - h) < e" using LIMSEQ_D [OF 0] \e > 0\ by (force simp: Norm_eq_lnorm \0 < p\ h lspace_g) then obtain N where N: "lnorm (lebesgue_on S) p (g N - h) < e/2" unfolding minus_fun_def by (meson \e>0\ half_gt_zero order_refl) show ?thesis proof show "continuous_on UNIV (g N)" by (simp add: contg) show "g N \ lspace (lebesgue_on S) (ennreal p)" by (simp add: lspace_g) have "lnorm (lebesgue_on S) p (f - h + - (g N - h)) \ lnorm (lebesgue_on S) p (f - h) + lnorm (lebesgue_on S) p (- (g N - h))" by (rule lnorm_triangle_fun) (auto simp: lspace_g h assms) also have "\ < e/2 + e/2" using lesse2 N by (simp add: lnorm_minus_commute) finally show "lnorm (lebesgue_on S) p (f - g N) < e" by simp qed qed proposition square_integrable_approximate_continuous: assumes f: "f square_integrable S" and S: "S \ lmeasurable" and "e > 0" obtains g where "continuous_on UNIV g" "g square_integrable S" "l2norm S (\x. f x - g x) < e" proof - have f2: "f \ lspace (lebesgue_on S) 2" by (simp add: f square_integrable_imp_lspace) then obtain g where contg: "continuous_on UNIV g" and g2: "g \ lspace (lebesgue_on S) 2" and less_e: "lnorm (lebesgue_on S) 2 (\x. f x - g x) < e" using lspace_approximate_continuous [of f 2 S e] S \0 < e\ by (auto simp: minus_fun_def) show thesis proof show "g square_integrable S" using g2 by (simp add: S fmeasurableD square_integrable_iff_lspace) show "l2norm S (\x. f x - g x) < e" using less_e by (simp add: l2norm_lnorm) qed (simp add: contg) qed lemma absolutely_integrable_approximate_continuous: fixes f :: "real \ real" assumes f: "f absolutely_integrable_on S" and S: "S \ lmeasurable" and "0 < e" obtains g where "continuous_on UNIV g" "g absolutely_integrable_on S" "integral\<^sup>L (lebesgue_on S) (\x. \f x - g x\) < e" proof - obtain g where "continuous_on UNIV g" "g \ lspace (lebesgue_on S) 1" and lnorm: "lnorm (lebesgue_on S) 1 (f - g) < e" proof (rule lspace_approximate_continuous) show "f \ lspace (lebesgue_on S) (ennreal 1)" by (simp add: S f fmeasurableD lspace_1) qed (auto simp: assms) show thesis proof show "continuous_on UNIV g" by fact show "g absolutely_integrable_on S" using S \g \ lspace (lebesgue_on S) 1\ lspace_1 by blast have *: "(\x. f x - g x) absolutely_integrable_on S" by (simp add: \g absolutely_integrable_on S\ f) moreover have "integrable (lebesgue_on S) (\x. \f x - g x\)" by (simp add: L1_D(2) S * fmeasurableD lspace_1) ultimately show "integral\<^sup>L (lebesgue_on S) (\x. \f x - g x\) < e" using lnorm S unfolding lnorm_def absolutely_integrable_on_def by simp qed qed end