diff -r e13e70f32407 -r e01015e49041 src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Tue Dec 29 22:41:22 2015 +0100 +++ b/src/HOL/Probability/Bochner_Integration.thy Tue Dec 29 23:04:53 2015 +0100 @@ -18,7 +18,7 @@ lemma borel_measurable_implies_sequence_metric: fixes f :: "'a \ 'b :: {metric_space, second_countable_topology}" assumes [measurable]: "f \ borel_measurable M" - shows "\F. (\i. simple_function M (F i)) \ (\x\space M. (\i. F i x) ----> f x) \ + shows "\F. (\i. simple_function M (F i)) \ (\x\space M. (\i. F i x) \ f x) \ (\i. \x\space M. dist (F i x) z \ 2 * dist (f x) z)" proof - obtain D :: "'b set" where "countable D" and D: "\X. open X \ X \ {} \ \d\D. d \ X" @@ -101,7 +101,7 @@ note * = this fix x assume "x \ space M" - show "(\i. F i x) ----> f x" + show "(\i. F i x) \ f x" proof cases assume "f x = z" then have "\i n. x \ A i n" @@ -173,7 +173,7 @@ assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" - assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ (\x. x \ space M \ (\i. U i x) ----> u x) \ P u" + assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ (\x. x \ space M \ (\i. U i x) \ u x) \ P u" shows "P u" proof - have "(\x. ereal (u x)) \ borel_measurable M" using u by auto @@ -199,13 +199,13 @@ intro!: real_of_ereal_positive_mono) next fix x assume x: "x \ space M" - have "(\i. U i x) ----> (SUP i. U i x)" + have "(\i. U i x) \ (SUP i. U i x)" using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) moreover have "(\i. U i x) = (\i. ereal (U' i x))" using x nn U(3) by (auto simp: fun_eq_iff U'_def ereal_real image_iff eq_commute) moreover have "(SUP i. U i x) = ereal (u x)" using sup u(2) by (simp add: max_def) - ultimately show "(\i. U' i x) ----> u x" + ultimately show "(\i. U' i x) \ u x" by simp next fix i @@ -516,8 +516,8 @@ for M f x where "f \ borel_measurable M \ (\i. simple_bochner_integrable M (s i)) \ - (\i. \\<^sup>+x. norm (f x - s i x) \M) ----> 0 \ - (\i. simple_bochner_integral M (s i)) ----> x \ + (\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0 \ + (\i. simple_bochner_integral M (s i)) \ x \ has_bochner_integral M f x" lemma has_bochner_integral_cong: @@ -530,7 +530,7 @@ "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ has_bochner_integral M f x \ has_bochner_integral M g x" unfolding has_bochner_integral.simps - by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\x. x ----> 0"] + by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\x. x \ 0"] nn_integral_cong_AE) auto @@ -572,8 +572,8 @@ has_bochner_integral M (\x. f x + g x) (x + y)" proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) fix sf sg - assume f_sf: "(\i. \\<^sup>+ x. norm (f x - sf i x) \M) ----> 0" - assume g_sg: "(\i. \\<^sup>+ x. norm (g x - sg i x) \M) ----> 0" + assume f_sf: "(\i. \\<^sup>+ x. norm (f x - sf i x) \M) \ 0" + assume g_sg: "(\i. \\<^sup>+ x. norm (g x - sg i x) \M) \ 0" assume sf: "\i. simple_bochner_integrable M (sf i)" and sg: "\i. simple_bochner_integrable M (sg i)" @@ -584,10 +584,10 @@ show "\i. simple_bochner_integrable M (\x. sf i x + sg i x)" using sf sg by (simp add: simple_bochner_integrable_compose2) - show "(\i. \\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \M) ----> 0" - (is "?f ----> 0") + show "(\i. \\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \M) \ 0" + (is "?f \ 0") proof (rule tendsto_sandwich) - show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) ----> 0" + show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by (auto simp: nn_integral_nonneg) show "eventually (\i. ?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) \M) + \\<^sup>+ x. (norm (g x - sg i x)) \M) sequentially" (is "eventually (\i. ?f i \ ?g i) sequentially") @@ -598,7 +598,7 @@ by (intro nn_integral_add) auto finally show "?f i \ ?g i" . qed - show "?g ----> 0" + show "?g \ 0" using tendsto_add_ereal[OF _ _ f_sf g_sg] by simp qed qed (auto simp: simple_bochner_integral_add tendsto_add) @@ -614,7 +614,7 @@ then show "(\x. T (f x)) \ borel_measurable M" by auto - fix s assume f_s: "(\i. \\<^sup>+ x. norm (f x - s i x) \M) ----> 0" + fix s assume f_s: "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" assume s: "\i. simple_bochner_integrable M (s i)" then show "\i. simple_bochner_integrable M (\x. T (s i x))" by (auto intro: simple_bochner_integrable_compose2 T.zero) @@ -625,10 +625,10 @@ obtain K where K: "K > 0" "\x i. norm (T (f x) - T (s i x)) \ norm (f x - s i x) * K" using T.pos_bounded by (auto simp: T.diff[symmetric]) - show "(\i. \\<^sup>+ x. norm (T (f x) - T (s i x)) \M) ----> 0" - (is "?f ----> 0") + show "(\i. \\<^sup>+ x. norm (T (f x) - T (s i x)) \M) \ 0" + (is "?f \ 0") proof (rule tendsto_sandwich) - show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) ----> 0" + show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) \ 0" by (auto simp: nn_integral_nonneg) show "eventually (\i. ?f i \ K * (\\<^sup>+ x. norm (f x - s i x) \M)) sequentially" @@ -640,12 +640,12 @@ using K by (intro nn_integral_cmult) auto finally show "?f i \ ?g i" . qed - show "?g ----> 0" + show "?g \ 0" using tendsto_cmult_ereal[OF _ f_s, of "ereal K"] by simp qed - assume "(\i. simple_bochner_integral M (s i)) ----> x" - with s show "(\i. simple_bochner_integral M (\x. T (s i x))) ----> T x" + assume "(\i. simple_bochner_integral M (s i)) \ x" + with s show "(\i. simple_bochner_integral M (\x. T (s i x))) \ T x" by (auto intro!: T.tendsto simp: T.simple_bochner_integral_linear) qed @@ -724,7 +724,7 @@ proof (elim has_bochner_integral.cases) fix s v assume [measurable]: "f \ borel_measurable M" and s: "\i. simple_bochner_integrable M (s i)" and - lim_0: "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" + lim_0: "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) \ 0" from order_tendstoD[OF lim_0, of "\"] obtain i where f_s_fin: "(\\<^sup>+ x. ereal (norm (f x - s i x)) \M) < \" by (metis (mono_tags, lifting) eventually_False_sequentially eventually_mono @@ -760,9 +760,9 @@ shows "norm x \ (\\<^sup>+x. norm (f x) \M)" using assms proof fix s assume - x: "(\i. simple_bochner_integral M (s i)) ----> x" (is "?s ----> x") and + x: "(\i. simple_bochner_integral M (s i)) \ x" (is "?s \ x") and s[simp]: "\i. simple_bochner_integrable M (s i)" and - lim: "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" and + lim: "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) \ 0" and f[measurable]: "f \ borel_measurable M" have [measurable]: "\i. s i \ borel_measurable M" @@ -770,7 +770,7 @@ show "norm x \ (\\<^sup>+x. norm (f x) \M)" proof (rule LIMSEQ_le) - show "(\i. ereal (norm (?s i))) ----> norm x" + show "(\i. ereal (norm (?s i))) \ norm x" using x by (intro tendsto_intros lim_ereal[THEN iffD2]) show "\N. \n\N. norm (?s n) \ (\\<^sup>+x. norm (f x - s n x) \M) + (\\<^sup>+x. norm (f x) \M)" (is "\N. \n\N. _ \ ?t n") @@ -788,10 +788,10 @@ by (rule nn_integral_add) auto finally show "norm (?s n) \ ?t n" . qed - have "?t ----> 0 + (\\<^sup>+ x. ereal (norm (f x)) \M)" + have "?t \ 0 + (\\<^sup>+ x. ereal (norm (f x)) \M)" using has_bochner_integral_implies_finite_norm[OF i] by (intro tendsto_add_ereal tendsto_const lim) auto - then show "?t ----> \\<^sup>+ x. ereal (norm (f x)) \M" + then show "?t \ \\<^sup>+ x. ereal (norm (f x)) \M" by simp qed qed @@ -802,8 +802,8 @@ assume f[measurable]: "f \ borel_measurable M" fix s t - assume "(\i. \\<^sup>+ x. norm (f x - s i x) \M) ----> 0" (is "?S ----> 0") - assume "(\i. \\<^sup>+ x. norm (f x - t i x) \M) ----> 0" (is "?T ----> 0") + assume "(\i. \\<^sup>+ x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") + assume "(\i. \\<^sup>+ x. norm (f x - t i x) \M) \ 0" (is "?T \ 0") assume s: "\i. simple_bochner_integrable M (s i)" assume t: "\i. simple_bochner_integrable M (t i)" @@ -812,22 +812,22 @@ let ?s = "\i. simple_bochner_integral M (s i)" let ?t = "\i. simple_bochner_integral M (t i)" - assume "?s ----> x" "?t ----> y" - then have "(\i. norm (?s i - ?t i)) ----> norm (x - y)" + assume "?s \ x" "?t \ y" + then have "(\i. norm (?s i - ?t i)) \ norm (x - y)" by (intro tendsto_intros) moreover - have "(\i. ereal (norm (?s i - ?t i))) ----> ereal 0" + have "(\i. ereal (norm (?s i - ?t i))) \ ereal 0" proof (rule tendsto_sandwich) - show "eventually (\i. 0 \ ereal (norm (?s i - ?t i))) sequentially" "(\_. 0) ----> ereal 0" + show "eventually (\i. 0 \ ereal (norm (?s i - ?t i))) sequentially" "(\_. 0) \ ereal 0" by (auto simp: nn_integral_nonneg zero_ereal_def[symmetric]) show "eventually (\i. norm (?s i - ?t i) \ ?S i + ?T i) sequentially" by (intro always_eventually allI simple_bochner_integral_bounded s t f) - show "(\i. ?S i + ?T i) ----> ereal 0" - using tendsto_add_ereal[OF _ _ \?S ----> 0\ \?T ----> 0\] + show "(\i. ?S i + ?T i) \ ereal 0" + using tendsto_add_ereal[OF _ _ \?S \ 0\ \?T \ 0\] by (simp add: zero_ereal_def[symmetric]) qed - then have "(\i. norm (?s i - ?t i)) ----> 0" + then have "(\i. norm (?s i - ?t i)) \ 0" by simp ultimately have "norm (x - y) = 0" by (rule LIMSEQ_unique) @@ -841,11 +841,11 @@ shows "has_bochner_integral M g x" using f proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) - fix s assume "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" + fix s assume "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) \ 0" also have "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) = (\i. \\<^sup>+ x. ereal (norm (g x - s i x)) \M)" using ae by (intro ext nn_integral_cong_AE, eventually_elim) simp - finally show "(\i. \\<^sup>+ x. ereal (norm (g x - s i x)) \M) ----> 0" . + finally show "(\i. \\<^sup>+ x. ereal (norm (g x - s i x)) \M) \ 0" . qed (auto intro: g) lemma has_bochner_integral_eq_AE: @@ -1144,12 +1144,12 @@ fixes f :: "'a \ 'b::{banach, second_countable_topology}" assumes f[measurable]: "f \ borel_measurable M" assumes s: "\i. simple_bochner_integrable M (s i)" - assumes lim: "(\i. \\<^sup>+x. norm (f x - s i x) \M) ----> 0" (is "?S ----> 0") + assumes lim: "(\i. \\<^sup>+x. norm (f x - s i x) \M) \ 0" (is "?S \ 0") shows "integrable M f" proof - let ?s = "\n. simple_bochner_integral M (s n)" - have "\x. ?s ----> x" + have "\x. ?s \ x" unfolding convergent_eq_cauchy proof (rule metric_CauchyI) fix e :: real assume "0 < e" @@ -1172,7 +1172,7 @@ by (simp add: dist_norm) qed qed - then obtain x where "?s ----> x" .. + then obtain x where "?s \ x" .. show ?thesis by (rule, rule) fact+ qed @@ -1183,14 +1183,14 @@ "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" and bound: "\j. AE x in M. norm (u j x) \ w x" and w: "(\\<^sup>+x. w x \M) < \" - and u': "AE x in M. (\i. u i x) ----> u' x" - shows "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) ----> 0" + and u': "AE x in M. (\i. u i x) \ u' x" + shows "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ 0" proof - have "AE x in M. \j. norm (u j x) \ w x" unfolding AE_all_countable by rule fact with u' have bnd: "AE x in M. \j. norm (u' x - u j x) \ 2 * w x" proof (eventually_elim, intro allI) - fix i x assume "(\i. u i x) ----> u' x" "\j. norm (u j x) \ w x" "\j. norm (u j x) \ w x" + fix i x assume "(\i. u i x) \ u' x" "\j. norm (u j x) \ w x" "\j. norm (u j x) \ w x" then have "norm (u' x) \ w x" "norm (u i x) \ w x" by (auto intro: LIMSEQ_le_const2 tendsto_norm) then have "norm (u' x) + norm (u i x) \ 2 * w x" @@ -1200,16 +1200,16 @@ finally (xtrans) show "norm (u' x - u i x) \ 2 * w x" . qed - have "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) ----> (\\<^sup>+x. 0 \M)" + have "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) \ (\\<^sup>+x. 0 \M)" proof (rule nn_integral_dominated_convergence) show "(\\<^sup>+x. 2 * w x \M) < \" by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) auto - show "AE x in M. (\i. ereal (norm (u' x - u i x))) ----> 0" + show "AE x in M. (\i. ereal (norm (u' x - u i x))) \ 0" using u' proof eventually_elim - fix x assume "(\i. u i x) ----> u' x" + fix x assume "(\i. u i x) \ u' x" from tendsto_diff[OF tendsto_const[of "u' x"] this] - show "(\i. ereal (norm (u' x - u i x))) ----> 0" + show "(\i. ereal (norm (u' x - u i x))) \ 0" by (simp add: zero_ereal_def tendsto_norm_zero_iff) qed qed (insert bnd, auto) @@ -1223,7 +1223,7 @@ proof - from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where s: "\i. simple_function M (s i)" and - pointwise: "\x. x \ space M \ (\i. s i x) ----> f x" and + pointwise: "\x. x \ space M \ (\i. s i x) \ f x" and bound: "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" by (simp add: norm_conv_dist) metis @@ -1241,7 +1241,7 @@ show "\i. simple_bochner_integrable M (s i)" by (rule simple_bochner_integrableI_bounded) fact+ - show "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" + show "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) \ 0" proof (rule nn_integral_dominated_convergence_norm) show "\j. AE x in M. norm (s j x) \ 2 * norm (f x)" using bound by auto @@ -1249,7 +1249,7 @@ using s by (auto intro: borel_measurable_simple_function) show "(\\<^sup>+ x. ereal (2 * norm (f x)) \M) < \" using fin unfolding times_ereal.simps(1)[symmetric] by (subst nn_integral_cmult) auto - show "AE x in M. (\i. s i x) ----> f x" + show "AE x in M. (\i. s i x) \ f x" using pointwise by auto qed fact qed fact @@ -1456,11 +1456,11 @@ lemma fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" assumes "f \ borel_measurable M" "\i. s i \ borel_measurable M" "integrable M w" - assumes lim: "AE x in M. (\i. s i x) ----> f x" + assumes lim: "AE x in M. (\i. s i x) \ f x" assumes bound: "\i. AE x in M. norm (s i x) \ w x" shows integrable_dominated_convergence: "integrable M f" and integrable_dominated_convergence2: "\i. integrable M (s i)" - and integral_dominated_convergence: "(\i. integral\<^sup>L M (s i)) ----> integral\<^sup>L M f" + and integral_dominated_convergence: "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" proof - have "AE x in M. 0 \ w x" using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) @@ -1487,16 +1487,16 @@ have "(\\<^sup>+ x. ereal (norm (f x)) \M) \ (\\<^sup>+x. w x \M)" using all_bound lim proof (intro nn_integral_mono_AE, eventually_elim) - fix x assume "\i. norm (s i x) \ w x" "(\i. s i x) ----> f x" + fix x assume "\i. norm (s i x) \ w x" "(\i. s i x) \ f x" then show "ereal (norm (f x)) \ ereal (w x)" by (intro LIMSEQ_le_const2[where X="\i. ereal (norm (s i x))"] tendsto_intros lim_ereal[THEN iffD2]) auto qed with w show "(\\<^sup>+ x. ereal (norm (f x)) \M) < \" by auto qed fact - have "(\n. ereal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) ----> ereal 0" (is "?d ----> ereal 0") + have "(\n. ereal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \ ereal 0" (is "?d \ ereal 0") proof (rule tendsto_sandwich) - show "eventually (\n. ereal 0 \ ?d n) sequentially" "(\_. ereal 0) ----> ereal 0" by auto + show "eventually (\n. ereal 0 \ ?d n) sequentially" "(\_. ereal 0) \ ereal 0" by auto show "eventually (\n. ?d n \ (\\<^sup>+x. norm (s n x - f x) \M)) sequentially" proof (intro always_eventually allI) fix n @@ -1506,7 +1506,7 @@ by (intro int_f int_s integrable_diff integral_norm_bound_ereal) finally show "?d n \ (\\<^sup>+x. norm (s n x - f x) \M)" . qed - show "(\n. \\<^sup>+x. norm (s n x - f x) \M) ----> ereal 0" + show "(\n. \\<^sup>+x. norm (s n x - f x) \M) \ ereal 0" unfolding zero_ereal_def[symmetric] apply (subst norm_minus_commute) proof (rule nn_integral_dominated_convergence_norm[where w=w]) @@ -1514,10 +1514,10 @@ using int_s unfolding integrable_iff_bounded by auto qed fact+ qed - then have "(\n. integral\<^sup>L M (s n) - integral\<^sup>L M f) ----> 0" + then have "(\n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \ 0" unfolding lim_ereal tendsto_norm_zero_iff . from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] - show "(\i. integral\<^sup>L M (s i)) ----> integral\<^sup>L M f" by simp + show "(\i. integral\<^sup>L M (s i)) \ integral\<^sup>L M f" by simp qed context @@ -1535,15 +1535,15 @@ obtain N where w: "\n. N \ n \ AE x in M. norm (s (X n) x) \ w x" by (auto simp: eventually_sequentially) - show "(\n. integral\<^sup>L M (s (X n))) ----> integral\<^sup>L M f" + show "(\n. integral\<^sup>L M (s (X n))) \ integral\<^sup>L M f" proof (rule LIMSEQ_offset, rule integral_dominated_convergence) show "AE x in M. norm (s (X (n + N)) x) \ w x" for n by (rule w) auto - show "AE x in M. (\n. s (X (n + N)) x) ----> f x" + show "AE x in M. (\n. s (X (n + N)) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) ---> f x) at_top" - then show "(\n. s (X (n + N)) x) ----> f x" + then show "(\n. s (X (n + N)) x) \ f x" by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X]) qed qed fact+ @@ -1557,11 +1557,11 @@ proof (rule integrable_dominated_convergence) show "AE x in M. norm (s (N + i) x) \ w x" for i :: nat by (intro w) auto - show "AE x in M. (\i. s (N + real i) x) ----> f x" + show "AE x in M. (\i. s (N + real i) x) \ f x" using lim proof eventually_elim fix x assume "((\i. s i x) ---> f x) at_top" - then show "(\n. s (N + n) x) ----> f x" + then show "(\n. s (N + n) x) \ f x" by (rule filterlim_compose) (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially) qed @@ -1606,7 +1606,7 @@ show ?case proof (rule LIMSEQ_unique) - show "(\i. ereal (integral\<^sup>L M (s i))) ----> ereal (integral\<^sup>L M f)" + show "(\i. ereal (integral\<^sup>L M (s i))) \ ereal (integral\<^sup>L M f)" unfolding lim_ereal proof (rule integral_dominated_convergence[where w=f]) show "integrable M f" by fact @@ -1615,13 +1615,13 @@ qed (insert seq, auto) have int_s: "\i. integrable M (s i)" using seq f s_le_f by (intro integrable_bound[OF f(3)]) auto - have "(\i. \\<^sup>+ x. s i x \M) ----> \\<^sup>+ x. f x \M" + have "(\i. \\<^sup>+ x. s i x \M) \ \\<^sup>+ x. f x \M" using seq s_le_f f by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded) also have "(\i. \\<^sup>+x. s i x \M) = (\i. \x. s i x \M)" using seq int_s by simp - finally show "(\i. \x. s i x \M) ----> \\<^sup>+x. f x \M" + finally show "(\i. \x. s i x \M) \ \\<^sup>+x. f x \M" by simp qed qed } @@ -1660,7 +1660,7 @@ by (simp add: suminf_ereal' sums) qed simp - have 2: "AE x in M. (\n. \i (\i. f i x)" + have 2: "AE x in M. (\n. \i (\i. f i x)" using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel) have 3: "\j. AE x in M. norm (\i (\i. norm (f i x))" @@ -1742,14 +1742,14 @@ assumes base: "\A c. A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R c)" assumes add: "\f g. integrable M f \ P f \ integrable M g \ P g \ P (\x. f x + g x)" assumes lim: "\f s. (\i. integrable M (s i)) \ (\i. P (s i)) \ - (\x. x \ space M \ (\i. s i x) ----> f x) \ + (\x. x \ space M \ (\i. s i x) \ f x) \ (\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)) \ integrable M f \ P f" shows "P f" proof - from \integrable M f\ have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" unfolding integrable_iff_bounded by auto from borel_measurable_implies_sequence_metric[OF f(1)] - obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) ----> f x" + obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) \ f x" "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" unfolding norm_conv_dist by metis @@ -1799,7 +1799,7 @@ then show "P (s' i)" by (subst s'_eq) (auto intro!: setsum base) - fix x assume "x \ space M" with s show "(\i. s' i x) ----> f x" + fix x assume "x \ space M" with s show "(\i. s' i x) \ f x" by (simp add: s'_eq_s) show "norm (s' i x) \ 2 * norm (f x)" using \x \ space M\ s by (simp add: s'_eq_s) @@ -1923,10 +1923,10 @@ case (lim f s) show ?case proof (rule LIMSEQ_unique) - show "(\i. integral\<^sup>L (restrict_space M \) (s i)) ----> integral\<^sup>L (restrict_space M \) f" + show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ integral\<^sup>L (restrict_space M \) f" using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) simp_all - show "(\i. integral\<^sup>L (restrict_space M \) (s i)) ----> (\ x. indicator \ x *\<^sub>R f x \M)" + show "(\i. integral\<^sup>L (restrict_space M \) (s i)) \ (\ x. indicator \ x *\<^sub>R f x \M)" unfolding lim using lim by (intro integral_dominated_convergence[where w="\x. 2 * norm (indicator \ x *\<^sub>R f x)"]) @@ -1998,16 +1998,16 @@ show ?case proof (rule LIMSEQ_unique) - show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) ----> integral\<^sup>L M (\x. g x *\<^sub>R f x)" + show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L M (\x. g x *\<^sub>R f x)" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (g x *\<^sub>R f x))" by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto - show "AE x in M. (\i. g x *\<^sub>R s i x) ----> g x *\<^sub>R f x" + show "AE x in M. (\i. g x *\<^sub>R s i x) \ g x *\<^sub>R f x" using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M]) show "\i. AE x in M. norm (g x *\<^sub>R s i x) \ 2 * norm (g x *\<^sub>R f x)" using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps) qed auto - show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) ----> integral\<^sup>L (density M g) f" + show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) \ integral\<^sup>L (density M g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) @@ -2077,16 +2077,16 @@ show ?case proof (rule LIMSEQ_unique) - show "(\i. integral\<^sup>L M (\x. s i (g x))) ----> integral\<^sup>L M (\x. f (g x))" + show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L M (\x. f (g x))" proof (rule integral_dominated_convergence) show "integrable M (\x. 2 * norm (f (g x)))" using lim by (auto simp: integrable_distr_eq) - show "AE x in M. (\i. s i (g x)) ----> f (g x)" + show "AE x in M. (\i. s i (g x)) \ f (g x)" using lim(3) g[THEN measurable_space] by auto show "\i. AE x in M. norm (s i (g x)) \ 2 * norm (f (g x))" using lim(4) g[THEN measurable_space] by auto qed auto - show "(\i. integral\<^sup>L M (\x. s i (g x))) ----> integral\<^sup>L (distr M N g) f" + show "(\i. integral\<^sup>L M (\x. s i (g x))) \ integral\<^sup>L (distr M N g) f" unfolding lim(2)[symmetric] by (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) (insert lim(3-5), auto) @@ -2266,8 +2266,8 @@ fixes f :: "nat \ 'a \ real" assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" and pos: "\i. AE x in M. 0 \ f i x" - and lim: "AE x in M. (\i. f i x) ----> u x" - and ilim: "(\i. integral\<^sup>L M (f i)) ----> x" + and lim: "AE x in M. (\i. f i x) \ u x" + and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows "integrable M u" and "integral\<^sup>L M u = x" @@ -2295,7 +2295,7 @@ using u by auto from mono pos[of 0] lim show "AE x in M. ereal (- u x) \ 0" proof eventually_elim - fix x assume "mono (\n. f n x)" "0 \ f 0 x" "(\i. f i x) ----> u x" + fix x assume "mono (\n. f n x)" "0 \ f 0 x" "(\i. f i x) \ u x" then show "ereal (- u x) \ 0" using incseq_le[of "\n. f n x" "u x" 0] by (simp add: mono_def incseq_def) qed @@ -2307,8 +2307,8 @@ lemma fixes f :: "nat \ 'a \ real" assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" - and lim: "AE x in M. (\i. f i x) ----> u x" - and ilim: "(\i. integral\<^sup>L M (f i)) ----> x" + and lim: "AE x in M. (\i. f i x) \ u x" + and ilim: "(\i. integral\<^sup>L M (f i)) \ x" and u: "u \ borel_measurable M" shows integrable_monotone_convergence: "integrable M u" and integral_monotone_convergence: "integral\<^sup>L M u = x" @@ -2320,9 +2320,9 @@ using mono by (auto simp: mono_def le_fun_def) have 3: "\n. AE x in M. 0 \ f n x - f 0 x" using mono by (auto simp: field_simps mono_def le_fun_def) - have 4: "AE x in M. (\i. f i x - f 0 x) ----> u x - f 0 x" + have 4: "AE x in M. (\i. f i x - f 0 x) \ u x - f 0 x" using lim by (auto intro!: tendsto_diff) - have 5: "(\i. (\x. f i x - f 0 x \M)) ----> x - integral\<^sup>L M (f 0)" + have 5: "(\i. (\x. f i x - f 0 x \M)) \ x - integral\<^sup>L M (f 0)" using f ilim by (auto intro!: tendsto_diff) have 6: "(\x. u x - f 0 x) \ borel_measurable M" using f[of 0] u by auto @@ -2462,17 +2462,17 @@ shows "((\y. \ x. indicator {.. y} x *\<^sub>R f x \M) ---> \ x. f x \M) at_top" proof (rule tendsto_at_topI_sequentially) fix X :: "nat \ real" assume "filterlim X at_top sequentially" - show "(\n. \x. indicator {..X n} x *\<^sub>R f x \M) ----> integral\<^sup>L M f" + show "(\n. \x. indicator {..X n} x *\<^sub>R f x \M) \ integral\<^sup>L M f" proof (rule integral_dominated_convergence) show "integrable M (\x. norm (f x))" by (rule integrable_norm) fact - show "AE x in M. (\n. indicator {..X n} x *\<^sub>R f x) ----> f x" + show "AE x in M. (\n. indicator {..X n} x *\<^sub>R f x) \ f x" proof fix x from \filterlim X at_top sequentially\ have "eventually (\n. x \ X n) sequentially" unfolding filterlim_at_top_ge[where c=x] by auto - then show "(\n. indicator {..X n} x *\<^sub>R f x) ----> f x" + then show "(\n. indicator {..X n} x *\<^sub>R f x) \ f x" by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono) qed fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \ norm (f x)" @@ -2497,9 +2497,9 @@ by (rule eventually_sequentiallyI[of "nat \x\"]) (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) } from filterlim_cong[OF refl refl this] - have "AE x in M. (\i. f x * indicator {..real i} x) ----> f x" + have "AE x in M. (\i. f x * indicator {..real i} x) \ f x" by simp - have "(\i. \ x. f x * indicator {..real i} x \M) ----> x" + have "(\i. \ x. f x * indicator {..real i} x \M) \ x" using conv filterlim_real_sequentially by (rule filterlim_compose) have M_measure[simp]: "borel_measurable M = borel_measurable borel" using M by (simp add: sets_eq_imp_space_eq measurable_def) @@ -2539,7 +2539,7 @@ proof - from borel_measurable_implies_sequence_metric[OF f, of 0] guess s .. then have s: "\i. simple_function (N \\<^sub>M M) (s i)" - "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) ----> f x y" + "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) \ f x y" "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" by (auto simp: space_pair_measure norm_conv_dist) @@ -2569,12 +2569,12 @@ { assume int_f: "integrable M (f x)" have int_2f: "integrable M (\y. 2 * norm (f x y))" by (intro integrable_norm integrable_mult_right int_f) - have "(\i. integral\<^sup>L M (\y. s i (x, y))) ----> integral\<^sup>L M (f x)" + have "(\i. integral\<^sup>L M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" proof (rule integral_dominated_convergence) from int_f show "f x \ borel_measurable M" by auto show "\i. (\y. s i (x, y)) \ borel_measurable M" using x by simp - show "AE xa in M. (\i. s i (x, xa)) ----> f x xa" + show "AE xa in M. (\i. s i (x, xa)) \ f x xa" using x s(2) by auto show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f x xa)" using x s(3) by auto @@ -2597,10 +2597,10 @@ qed then have "integral\<^sup>L M (\y. s i (x, y)) = simple_bochner_integral M (\y. s i (x, y))" by (rule simple_bochner_integrable_eq_integral[symmetric]) } - ultimately have "(\i. simple_bochner_integral M (\y. s i (x, y))) ----> integral\<^sup>L M (f x)" + ultimately have "(\i. simple_bochner_integral M (\y. s i (x, y))) \ integral\<^sup>L M (f x)" by simp } then - show "(\i. f' i x) ----> integral\<^sup>L M (f x)" + show "(\i. f' i x) \ integral\<^sup>L M (f x)" unfolding f'_def by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq) qed @@ -2779,25 +2779,25 @@ show ?case proof (rule LIMSEQ_unique) - show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) ----> integral\<^sup>L (M1 \\<^sub>M M2) f" + show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ integral\<^sup>L (M1 \\<^sub>M M2) f" proof (rule integral_dominated_convergence) show "integrable (M1 \\<^sub>M M2) (\x. 2 * norm (f x))" using lim(5) by auto qed (insert lim, auto) - have "(\i. \ x. \ y. s i (x, y) \M2 \M1) ----> \ x. \ y. f (x, y) \M2 \M1" + have "(\i. \ x. \ y. s i (x, y) \M2 \M1) \ \ x. \ y. f (x, y) \M2 \M1" proof (rule integral_dominated_convergence) have "AE x in M1. \i. integrable M2 (\y. s i (x, y))" unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] .. with AE_space AE_integrable_fst'[OF lim(5)] - show "AE x in M1. (\i. \ y. s i (x, y) \M2) ----> \ y. f (x, y) \M2" + show "AE x in M1. (\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof eventually_elim fix x assume x: "x \ space M1" and s: "\i. integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" - show "(\i. \ y. s i (x, y) \M2) ----> \ y. f (x, y) \M2" + show "(\i. \ y. s i (x, y) \M2) \ \ y. f (x, y) \M2" proof (rule integral_dominated_convergence) show "integrable M2 (\y. 2 * norm (f (x, y)))" using f by auto - show "AE xa in M2. (\i. s i (x, xa)) ----> f (x, xa)" + show "AE xa in M2. (\i. s i (x, xa)) \ f (x, xa)" using x lim(3) by (auto simp: space_pair_measure) show "\i. AE xa in M2. norm (s i (x, xa)) \ 2 * norm (f (x, xa))" using x lim(4) by (auto simp: space_pair_measure) @@ -2820,7 +2820,7 @@ by simp qed qed simp_all - then show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) ----> \ x. \ y. f (x, y) \M2 \M1" + then show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) \ \ x. \ y. f (x, y) \M2 \M1" using lim by simp qed qed @@ -2991,12 +2991,12 @@ using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all show ?case proof (intro LIMSEQ_unique) - show "(\i. integral\<^sup>L N (s i)) ----> integral\<^sup>L N f" + show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L N f" apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim apply auto done - show "(\i. integral\<^sup>L N (s i)) ----> integral\<^sup>L M f" + show "(\i. integral\<^sup>L N (s i)) \ integral\<^sup>L M f" unfolding lim apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) using lim M N(2)