diff -r b4552595b04e -r f6bb31879698 src/HOL/Analysis/Bochner_Integration.thy --- a/src/HOL/Analysis/Bochner_Integration.thy Tue Feb 16 17:12:02 2021 +0000 +++ b/src/HOL/Analysis/Bochner_Integration.thy Fri Feb 19 13:42:12 2021 +0100 @@ -1315,6 +1315,27 @@ using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f] unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto +lemma (in finite_measure) square_integrable_imp_integrable: + fixes f :: "'a \ 'b :: {second_countable_topology, banach, real_normed_div_algebra}" + assumes [measurable]: "f \ borel_measurable M" + assumes "integrable M (\x. f x ^ 2)" + shows "integrable M f" +proof - + have less_top: "emeasure M (space M) < top" + using finite_emeasure_space by (meson top.not_eq_extremum) + have "nn_integral M (\x. norm (f x)) ^ 2 \ + nn_integral M (\x. norm (f x) ^ 2) * emeasure M (space M)" + using Cauchy_Schwarz_nn_integral[of "\x. norm (f x)" M "\_. 1"] + by (simp add: ennreal_power) + also have "\ < \" + using assms(2) less_top + by (subst (asm) integrable_iff_bounded) (auto simp: norm_power ennreal_mult_less_top) + finally have "nn_integral M (\x. norm (f x)) < \" + by (simp add: power_less_top_ennreal) + thus ?thesis + by (simp add: integrable_iff_bounded) +qed + lemma integrable_bound: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ 'c::{banach, second_countable_topology}" @@ -2065,16 +2086,50 @@ have "{x \ space M. u x \ c} = {x \ space M. ennreal(1/c) * u x \ 1} \ (space M)" using \c>0\ by (auto simp: ennreal_mult'[symmetric]) - then have "emeasure M {x \ space M. u x \ c} = emeasure M ({x \ space M. ennreal(1/c) * u x \ 1} \ (space M))" + then have "emeasure M {x \ space M. u x \ c} = emeasure M ({x \ space M. ennreal(1/c) * u x \ 1})" by simp also have "... \ ennreal(1/c) * (\\<^sup>+ x. ennreal(u x) * indicator (space M) x \M)" by (rule nn_integral_Markov_inequality) (auto simp add: assms) also have "... \ ennreal(1/c) * (\x. u x \M)" - apply (rule mult_left_mono) using * \c>0\ by auto + by (rule mult_left_mono) (use * \c > 0\ in auto) finally show ?thesis using \0 by (simp add: ennreal_mult'[symmetric]) qed +theorem integral_Markov_inequality_measure: + assumes [measurable]: "integrable M u" and "A \ sets M" and "AE x in M. 0 \ u x" "0 < (c::real)" + shows "measure M {x\space M. u x \ c} \ (\x. u x \M) / c" +proof - + have le: "emeasure M {x\space M. u x \ c} \ ennreal ((1/c) * (\x. u x \M))" + by (rule integral_Markov_inequality) (use assms in auto) + also have "\ < top" + by auto + finally have "ennreal (measure M {x\space M. u x \ c}) = emeasure M {x\space M. u x \ c}" + by (intro emeasure_eq_ennreal_measure [symmetric]) auto + also note le + finally show ?thesis + by (subst (asm) ennreal_le_iff) + (auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms) +qed + +theorem%important (in finite_measure) second_moment_method: + assumes [measurable]: "f \ M \\<^sub>M borel" + assumes "integrable M (\x. f x ^ 2)" + defines "\ \ lebesgue_integral M f" + assumes "a > 0" + shows "measure M {x\space M. \f x\ \ a} \ lebesgue_integral M (\x. f x ^ 2) / a\<^sup>2" +proof - + have integrable: "integrable M f" + using assms by (blast dest: square_integrable_imp_integrable) + have "{x\space M. \f x\ \ a} = {x\space M. f x ^ 2 \ a\<^sup>2}" + using \a > 0\ by (simp flip: abs_le_square_iff) + hence "measure M {x\space M. \f x\ \ a} = measure M {x\space M. f x ^ 2 \ a\<^sup>2}" + by simp + also have "\ \ lebesgue_integral M (\x. f x ^ 2) / a\<^sup>2" + using assms by (intro integral_Markov_inequality_measure) auto + finally show ?thesis . +qed + lemma integral_ineq_eq_0_then_AE: fixes f::"_ \ real" assumes "AE x in M. f x \ g x" "integrable M f" "integrable M g"