diff -r 646a1399e792 -r c3b977fee8a3 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Fri Jan 14 14:21:26 2011 +0100 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Fri Jan 14 14:21:48 2011 +0100 @@ -489,7 +489,7 @@ section "Simple integral" -definition (in measure_space) +definition (in measure_space) simple_integral (binder "\\<^isup>S " 10) where "simple_integral f = (\x \ f ` space M. x * \ (f -` {x} \ space M))" lemma (in measure_space) simple_integral_cong: @@ -514,7 +514,7 @@ qed lemma (in measure_space) simple_integral_const[simp]: - "simple_integral (\x. c) = c * \ (space M)" + "(\\<^isup>Sx. c) = c * \ (space M)" proof (cases "space M = {}") case True thus ?thesis unfolding simple_integral_def by simp next @@ -579,7 +579,7 @@ lemma (in measure_space) simple_integral_add[simp]: assumes "simple_function f" and "simple_function g" - shows "simple_integral (\x. f x + g x) = simple_integral f + simple_integral g" + shows "(\\<^isup>Sx. f x + g x) = simple_integral f + simple_integral g" proof - { fix x let ?S = "g -` {g x} \ f -` {f x} \ space M" assume "x \ space M" @@ -597,7 +597,7 @@ lemma (in measure_space) simple_integral_setsum[simp]: assumes "\i. i \ P \ simple_function (f i)" - shows "simple_integral (\x. \i\P. f i x) = (\i\P. simple_integral (f i))" + shows "(\\<^isup>Sx. \i\P. f i x) = (\i\P. simple_integral (f i))" proof cases assume "finite P" from this assms show ?thesis @@ -606,7 +606,7 @@ lemma (in measure_space) simple_integral_mult[simp]: assumes "simple_function f" - shows "simple_integral (\x. c * f x) = c * simple_integral f" + shows "(\\<^isup>Sx. c * f x) = c * simple_integral f" proof - note mult = simple_function_mult[OF simple_function_const[of c] assms] { fix x let ?S = "f -` {f x} \ (\x. c * f x) -` {c * f x} \ space M" @@ -700,11 +700,11 @@ lemma (in measure_space) simple_integral_indicator: assumes "A \ sets M" assumes "simple_function f" - shows "simple_integral (\x. f x * indicator A x) = + shows "(\\<^isup>Sx. f x * indicator A x) = (\x \ f ` space M. x * \ (f -` {x} \ space M \ A))" proof cases assume "A = space M" - moreover hence "simple_integral (\x. f x * indicator A x) = simple_integral f" + moreover hence "(\\<^isup>Sx. f x * indicator A x) = simple_integral f" by (auto intro!: simple_integral_cong) moreover have "\X. X \ space M \ space M = X \ space M" by auto ultimately show ?thesis by (simp add: simple_integral_def) @@ -720,7 +720,7 @@ next show "0 \ ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x]) qed - have *: "simple_integral (\x. f x * indicator A x) = + have *: "(\\<^isup>Sx. f x * indicator A x) = (\x \ f ` space M \ {0}. x * \ (f -` {x} \ space M \ A))" unfolding simple_integral_def I proof (rule setsum_mono_zero_cong_left) @@ -760,18 +760,18 @@ lemma (in measure_space) simple_integral_null_set: assumes "simple_function u" "N \ null_sets" - shows "simple_integral (\x. u x * indicator N x) = 0" + shows "(\\<^isup>Sx. u x * indicator N x) = 0" proof - have "AE x. indicator N x = (0 :: pextreal)" using `N \ null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N]) - then have "simple_integral (\x. u x * indicator N x) = simple_integral (\x. 0)" + then have "(\\<^isup>Sx. u x * indicator N x) = (\\<^isup>Sx. 0)" using assms by (intro simple_integral_cong_AE) (auto intro!: AE_disjI2) then show ?thesis by simp qed lemma (in measure_space) simple_integral_cong_AE_mult_indicator: assumes sf: "simple_function f" and eq: "AE x. x \ S" and "S \ sets M" - shows "simple_integral f = simple_integral (\x. f x * indicator S x)" + shows "simple_integral f = (\\<^isup>Sx. f x * indicator S x)" proof (rule simple_integral_cong_AE) show "simple_function f" by fact show "simple_function (\x. f x * indicator S x)" @@ -783,7 +783,7 @@ lemma (in measure_space) simple_integral_restricted: assumes "A \ sets M" assumes sf: "simple_function (\x. f x * indicator A x)" - shows "measure_space.simple_integral (restricted_space A) \ f = simple_integral (\x. f x * indicator A x)" + shows "measure_space.simple_integral (restricted_space A) \ f = (\\<^isup>Sx. f x * indicator A x)" (is "_ = simple_integral ?f") unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \ sets M`]] unfolding simple_integral_def @@ -842,7 +842,7 @@ section "Continuous posititve integration" -definition (in measure_space) +definition (in measure_space) positive_integral (binder "\\<^isup>+ " 10) where "positive_integral f = SUPR {g. simple_function g \ g \ f} simple_integral" lemma (in measure_space) positive_integral_alt: @@ -1071,7 +1071,7 @@ fixes g :: "'d \ pextreal" and f :: "'d \ 'a" assumes f: "bij_inv S (space M) f h" shows "measure_space.positive_integral (vimage_algebra S f) (\A. \ (f ` A)) g = - positive_integral (\x. g (h x))" + (\\<^isup>+x. g (h x))" proof - interpret v: measure_space "vimage_algebra S f" "\A. \ (f ` A)" using f by (rule measure_space_isomorphic[OF bij_inv_bij_betw(1)]) @@ -1149,7 +1149,7 @@ unfolding pextreal_SUP_cmult[symmetric] proof (safe intro!: SUP_mono bexI) fix i - have "a * simple_integral (?uB i) = simple_integral (\x. a * ?uB i x)" + have "a * simple_integral (?uB i) = (\\<^isup>Sx. a * ?uB i x)" using B `simple_function u` by (subst simple_integral_mult[symmetric]) (auto simp: field_simps) also have "\ \ positive_integral (f i)" @@ -1193,7 +1193,7 @@ lemma (in measure_space) positive_integral_monotone_convergence_SUP: assumes "\i x. x \ space M \ f i x \ f (Suc i) x" assumes "\i. f i \ borel_measurable M" - shows "(SUP i. positive_integral (f i)) = positive_integral (\x. SUP i. f i x)" + shows "(SUP i. positive_integral (f i)) = (\\<^isup>+ x. SUP i. f i x)" (is "_ = positive_integral ?u") proof - show ?thesis @@ -1236,7 +1236,7 @@ qed lemma (in measure_space) positive_integral_const[simp]: - "positive_integral (\x. c) = c * \ (space M)" + "(\\<^isup>+ x. c) = c * \ (space M)" by (subst positive_integral_eq_simple_integral) auto lemma (in measure_space) positive_integral_isoton_simple: @@ -1248,7 +1248,7 @@ lemma (in measure_space) positive_integral_linear: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" - shows "positive_integral (\x. a * f x + g x) = + shows "(\\<^isup>+ x. a * f x + g x) = a * positive_integral f + positive_integral g" (is "positive_integral ?L = _") proof - @@ -1276,30 +1276,32 @@ lemma (in measure_space) positive_integral_cmult: assumes "f \ borel_measurable M" - shows "positive_integral (\x. c * f x) = c * positive_integral f" + shows "(\\<^isup>+ x. c * f x) = c * positive_integral f" using positive_integral_linear[OF assms, of "\x. 0" c] by auto lemma (in measure_space) positive_integral_multc: assumes "f \ borel_measurable M" - shows "positive_integral (\x. f x * c) = positive_integral f * c" + shows "(\\<^isup>+ x. f x * c) = positive_integral f * c" unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp lemma (in measure_space) positive_integral_indicator[simp]: - "A \ sets M \ positive_integral (\x. indicator A x) = \ A" -by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) + "A \ sets M \ (\\<^isup>+ x. indicator A x) = \ A" + by (subst positive_integral_eq_simple_integral) + (auto simp: simple_function_indicator simple_integral_indicator) lemma (in measure_space) positive_integral_cmult_indicator: - "A \ sets M \ positive_integral (\x. c * indicator A x) = c * \ A" -by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) + "A \ sets M \ (\\<^isup>+ x. c * indicator A x) = c * \ A" + by (subst positive_integral_eq_simple_integral) + (auto simp: simple_function_indicator simple_integral_indicator) lemma (in measure_space) positive_integral_add: assumes "f \ borel_measurable M" "g \ borel_measurable M" - shows "positive_integral (\x. f x + g x) = positive_integral f + positive_integral g" + shows "(\\<^isup>+ x. f x + g x) = positive_integral f + positive_integral g" using positive_integral_linear[OF assms, of 1] by simp lemma (in measure_space) positive_integral_setsum: assumes "\i. i\P \ f i \ borel_measurable M" - shows "positive_integral (\x. \i\P. f i x) = (\i\P. positive_integral (f i))" + shows "(\\<^isup>+ x. \i\P. f i x) = (\i\P. positive_integral (f i))" proof cases assume "finite P" from this assms show ?thesis @@ -1317,11 +1319,11 @@ assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" and fin: "positive_integral g \ \" and mono: "\x. x \ space M \ g x \ f x" - shows "positive_integral (\x. f x - g x) = positive_integral f - positive_integral g" + shows "(\\<^isup>+ x. f x - g x) = positive_integral f - positive_integral g" proof - have borel: "(\x. f x - g x) \ borel_measurable M" using f g by (rule borel_measurable_pextreal_diff) - have "positive_integral (\x. f x - g x) + positive_integral g = + have "(\\<^isup>+x. f x - g x) + positive_integral g = positive_integral f" unfolding positive_integral_add[OF borel g, symmetric] proof (rule positive_integral_cong) @@ -1335,9 +1337,9 @@ lemma (in measure_space) positive_integral_psuminf: assumes "\i. f i \ borel_measurable M" - shows "positive_integral (\x. \\<^isub>\ i. f i x) = (\\<^isub>\ i. positive_integral (f i))" + shows "(\\<^isup>+ x. \\<^isub>\ i. f i x) = (\\<^isub>\ i. positive_integral (f i))" proof - - have "(\i. positive_integral (\x. \i positive_integral (\x. \\<^isub>\i. f i x)" + have "(\i. (\\<^isup>+x. \i (\\<^isup>+x. \\<^isub>\i. f i x)" by (rule positive_integral_isoton) (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono arg_cong[where f=Sup] @@ -1350,11 +1352,11 @@ lemma (in measure_space) positive_integral_lim_INF: fixes u :: "nat \ 'a \ pextreal" assumes "\i. u i \ borel_measurable M" - shows "positive_integral (\x. SUP n. INF m. u (m + n) x) \ + shows "(\\<^isup>+ x. SUP n. INF m. u (m + n) x) \ (SUP n. INF m. positive_integral (u (m + n)))" proof - - have "positive_integral (\x. SUP n. INF m. u (m + n) x) - = (SUP n. positive_integral (\x. INF m. u (m + n) x))" + have "(\\<^isup>+x. SUP n. INF m. u (m + n) x) + = (SUP n. (\\<^isup>+x. INF m. u (m + n) x))" using assms by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono) (auto simp del: add_Suc simp add: add_Suc[symmetric]) @@ -1365,7 +1367,7 @@ lemma (in measure_space) measure_space_density: assumes borel: "u \ borel_measurable M" - shows "measure_space M (\A. positive_integral (\x. u x * indicator A x))" (is "measure_space M ?v") + shows "measure_space M (\A. (\\<^isup>+ x. u x * indicator A x))" (is "measure_space M ?v") proof show "?v {} = 0" by simp show "countably_additive M ?v" @@ -1384,8 +1386,8 @@ lemma (in measure_space) positive_integral_translated_density: assumes "f \ borel_measurable M" "g \ borel_measurable M" - shows "measure_space.positive_integral M (\A. positive_integral (\x. f x * indicator A x)) g = - positive_integral (\x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") + shows "measure_space.positive_integral M (\A. (\\<^isup>+ x. f x * indicator A x)) g = + (\\<^isup>+ x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") proof - from measure_space_density[OF assms(1)] interpret T: measure_space M ?T . @@ -1399,30 +1401,30 @@ using G(1) unfolding simple_function_def by auto have "T.positive_integral (G i) = T.simple_integral (G i)" using G T.positive_integral_eq_simple_integral by simp - also have "\ = positive_integral (\x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x))" + also have "\ = (\\<^isup>+x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x))" apply (simp add: T.simple_integral_def) apply (subst positive_integral_cmult[symmetric]) using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) apply (subst positive_integral_setsum[symmetric]) using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) by (simp add: setsum_right_distrib field_simps) - also have "\ = positive_integral (\x. f x * G i x)" + also have "\ = (\\<^isup>+x. f x * G i x)" by (auto intro!: positive_integral_cong simp: indicator_def if_distrib setsum_cases) - finally have "T.positive_integral (G i) = positive_integral (\x. f x * G i x)" . } - with * have eq_Tg: "(\i. positive_integral (\x. f x * G i x)) \ T.positive_integral g" by simp + finally have "T.positive_integral (G i) = (\\<^isup>+x. f x * G i x)" . } + with * have eq_Tg: "(\i. (\\<^isup>+x. f x * G i x)) \ T.positive_integral g" by simp from G(2) have "(\i x. f x * G i x) \ (\x. f x * g x)" unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) - then have "(\i. positive_integral (\x. f x * G i x)) \ positive_integral (\x. f x * g x)" + then have "(\i. (\\<^isup>+x. f x * G i x)) \ (\\<^isup>+x. f x * g x)" using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times) - with eq_Tg show "T.positive_integral g = positive_integral (\x. f x * g x)" + with eq_Tg show "T.positive_integral g = (\\<^isup>+x. f x * g x)" unfolding isoton_def by simp qed lemma (in measure_space) positive_integral_null_set: - assumes "N \ null_sets" shows "positive_integral (\x. u x * indicator N x) = 0" + assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x) = 0" proof - - have "positive_integral (\x. u x * indicator N x) = positive_integral (\x. 0)" + have "(\\<^isup>+ x. u x * indicator N x) = (\\<^isup>+ x. 0)" proof (intro positive_integral_cong_AE AE_I) show "{x \ space M. u x * indicator N x \ 0} \ N" by (auto simp: indicator_def) @@ -1434,20 +1436,20 @@ lemma (in measure_space) positive_integral_Markov_inequality: assumes borel: "u \ borel_measurable M" and "A \ sets M" and c: "c \ \" - shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * positive_integral (\x. u x * indicator A x)" + shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x)" (is "\ ?A \ _ * ?PI") proof - have "?A \ sets M" using `A \ sets M` borel by auto - hence "\ ?A = positive_integral (\x. indicator ?A x)" + hence "\ ?A = (\\<^isup>+ x. indicator ?A x)" using positive_integral_indicator by simp - also have "\ \ positive_integral (\x. c * (u x * indicator A x))" + also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x))" proof (rule positive_integral_mono) fix x assume "x \ space M" show "indicator ?A x \ c * (u x * indicator A x)" by (cases "x \ ?A") auto qed - also have "\ = c * positive_integral (\x. u x * indicator A x)" + also have "\ = c * (\\<^isup>+ x. u x * indicator A x)" using assms by (auto intro!: positive_integral_cmult borel_measurable_indicator) finally show ?thesis . @@ -1459,7 +1461,7 @@ (is "_ \ \ ?A = 0") proof - have A: "?A \ sets M" using borel by auto - have u: "positive_integral (\x. u x * indicator ?A x) = positive_integral u" + have u: "(\\<^isup>+ x. u x * indicator ?A x) = positive_integral u" by (auto intro!: positive_integral_cong simp: indicator_def) show ?thesis @@ -1467,7 +1469,7 @@ assume "\ ?A = 0" hence "?A \ null_sets" using `?A \ sets M` by auto from positive_integral_null_set[OF this] - have "0 = positive_integral (\x. u x * indicator ?A x)" by simp + have "0 = (\\<^isup>+ x. u x * indicator ?A x)" by simp thus "positive_integral u = 0" unfolding u by simp next assume *: "positive_integral u = 0" @@ -1507,7 +1509,7 @@ lemma (in measure_space) positive_integral_restricted: assumes "A \ sets M" - shows "measure_space.positive_integral (restricted_space A) \ f = positive_integral (\x. f x * indicator A x)" + shows "measure_space.positive_integral (restricted_space A) \ f = (\\<^isup>+ x. f x * indicator A x)" (is "measure_space.positive_integral ?R \ f = positive_integral ?f") proof - have msR: "measure_space ?R \" by (rule restricted_measure_space[OF `A \ sets M`]) @@ -1524,7 +1526,7 @@ fix g assume "simple_function (\x. g x * indicator A x)" "g \ f" then show "\x. simple_function x \ x \ (\x. f x * indicator A x) \ - simple_integral (\x. g x * indicator A x) = simple_integral x" + (\\<^isup>Sx. g x * indicator A x) = simple_integral x" apply (rule_tac exI[of _ "\x. g x * indicator A x"]) by (auto simp: indicator_def le_fun_def) next @@ -1560,20 +1562,18 @@ definition (in measure_space) integrable where "integrable f \ f \ borel_measurable M \ - positive_integral (\x. Real (f x)) \ \ \ - positive_integral (\x. Real (- f x)) \ \" + (\\<^isup>+ x. Real (f x)) \ \ \ + (\\<^isup>+ x. Real (- f x)) \ \" lemma (in measure_space) integrableD[dest]: assumes "integrable f" shows "f \ borel_measurable M" - "positive_integral (\x. Real (f x)) \ \" - "positive_integral (\x. Real (- f x)) \ \" + "(\\<^isup>+ x. Real (f x)) \ \" + "(\\<^isup>+ x. Real (- f x)) \ \" using assms unfolding integrable_def by auto -definition (in measure_space) integral where - "integral f = - real (positive_integral (\x. Real (f x))) - - real (positive_integral (\x. Real (- f x)))" +definition (in measure_space) integral (binder "\ " 10) where + "integral f = real ((\\<^isup>+ x. Real (f x))) - real ((\\<^isup>+ x. Real (- f x)))" lemma (in measure_space) integral_cong: assumes cong: "\x. x \ space M \ f x = g x" @@ -1609,7 +1609,7 @@ lemma (in measure_space) integral_eq_positive_integral: assumes "\x. 0 \ f x" - shows "integral f = real (positive_integral (\x. Real (f x)))" + shows "integral f = real ((\\<^isup>+ x. Real (f x)))" proof - have "\x. Real (- f x) = 0" using assms by simp thus ?thesis by (simp del: Real_eq_0 add: integral_def) @@ -1617,7 +1617,7 @@ lemma (in measure_space) integral_vimage_inv: assumes f: "bij_betw f S (space M)" - shows "measure_space.integral (vimage_algebra S f) (\A. \ (f ` A)) (\x. g x) = integral (\x. g (the_inv_into S f x))" + shows "measure_space.integral (vimage_algebra S f) (\A. \ (f ` A)) (\x. g x) = (\x. g (the_inv_into S f x))" proof - interpret v: measure_space "vimage_algebra S f" "\A. \ (f ` A)" using f by (rule measure_space_isomorphic) @@ -1634,7 +1634,7 @@ lemma (in measure_space) integral_minus[intro, simp]: assumes "integrable f" - shows "integrable (\x. - f x)" "integral (\x. - f x) = - integral f" + shows "integrable (\x. - f x)" "(\x. - f x) = - integral f" using assms by (auto simp: integrable_def integral_def) lemma (in measure_space) integral_of_positive_diff: @@ -1685,7 +1685,7 @@ lemma (in measure_space) integral_linear: assumes "integrable f" "integrable g" and "0 \ a" shows "integrable (\t. a * f t + g t)" - and "integral (\t. a * f t + g t) = a * integral f + integral g" + and "(\ t. a * f t + g t) = a * integral f + integral g" proof - let ?PI = positive_integral let "?f x" = "Real (f x)" @@ -1718,7 +1718,7 @@ show "integrable (\t. a * f t + g t)" by (rule diff) - from assms show "integral (\t. a * f t + g t) = a * integral f + integral g" + from assms show "(\ t. a * f t + g t) = a * integral f + integral g" unfolding diff(2) unfolding integral_def * linear integrable_def by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg") (auto simp add: field_simps zero_le_mult_iff) @@ -1727,21 +1727,21 @@ lemma (in measure_space) integral_add[simp, intro]: assumes "integrable f" "integrable g" shows "integrable (\t. f t + g t)" - and "integral (\t. f t + g t) = integral f + integral g" + and "(\ t. f t + g t) = integral f + integral g" using assms integral_linear[where a=1] by auto lemma (in measure_space) integral_zero[simp, intro]: shows "integrable (\x. 0)" - and "integral (\x. 0) = 0" + and "(\ x.0) = 0" unfolding integrable_def integral_def by (auto simp add: borel_measurable_const) lemma (in measure_space) integral_cmult[simp, intro]: assumes "integrable f" shows "integrable (\t. a * f t)" (is ?P) - and "integral (\t. a * f t) = a * integral f" (is ?I) + and "(\ t. a * f t) = a * integral f" (is ?I) proof - - have "integrable (\t. a * f t) \ integral (\t. a * f t) = a * integral f" + have "integrable (\t. a * f t) \ (\ t. a * f t) = a * integral f" proof (cases rule: le_cases) assume "0 \ a" show ?thesis using integral_linear[OF assms integral_zero(1) `0 \ a`] @@ -1758,7 +1758,7 @@ lemma (in measure_space) integral_multc: assumes "integrable f" - shows "integral (\x. f x * c) = integral f * c" + shows "(\ x. f x * c) = integral f * c" unfolding mult_commute[of _ c] integral_cmult[OF assms] .. lemma (in measure_space) integral_mono_AE: @@ -1785,7 +1785,7 @@ lemma (in measure_space) integral_diff[simp, intro]: assumes f: "integrable f" and g: "integrable g" shows "integrable (\t. f t - g t)" - and "integral (\t. f t - g t) = integral f - integral g" + and "(\ t. f t - g t) = integral f - integral g" using integral_add[OF f integral_minus(1)[OF g]] unfolding diff_minus integral_minus(2)[OF g] by auto @@ -1806,7 +1806,7 @@ lemma (in measure_space) integral_cmul_indicator: assumes "A \ sets M" and "c \ 0 \ \ A \ \" shows "integrable (\x. c * indicator A x)" (is ?P) - and "integral (\x. c * indicator A x) = c * real (\ A)" (is ?I) + and "(\x. c * indicator A x) = c * real (\ A)" (is ?I) proof - show ?P proof (cases "c = 0") @@ -1821,7 +1821,7 @@ lemma (in measure_space) integral_setsum[simp, intro]: assumes "\n. n \ S \ integrable (f n)" - shows "integral (\x. \ i \ S. f i x) = (\ i \ S. integral (f i))" (is "?int S") + shows "(\x. \ i \ S. f i x) = (\ i \ S. integral (f i))" (is "?int S") and "integrable (\x. \ i \ S. f i x)" (is "?I S") proof - have "?int S \ ?I S" @@ -1854,21 +1854,21 @@ assumes borel: "g \ borel_measurable M" shows "integrable g" proof - - have "positive_integral (\x. Real (g x)) \ positive_integral (\x. Real \g x\)" + have "(\\<^isup>+ x. Real (g x)) \ (\\<^isup>+ x. Real \g x\)" by (auto intro!: positive_integral_mono) - also have "\ \ positive_integral (\x. Real (f x))" + also have "\ \ (\\<^isup>+ x. Real (f x))" using f by (auto intro!: positive_integral_mono) also have "\ < \" using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have pos: "positive_integral (\x. Real (g x)) < \" . + finally have pos: "(\\<^isup>+ x. Real (g x)) < \" . - have "positive_integral (\x. Real (- g x)) \ positive_integral (\x. Real (\g x\))" + have "(\\<^isup>+ x. Real (- g x)) \ (\\<^isup>+ x. Real (\g x\))" by (auto intro!: positive_integral_mono) - also have "\ \ positive_integral (\x. Real (f x))" + also have "\ \ (\\<^isup>+ x. Real (f x))" using f by (auto intro!: positive_integral_mono) also have "\ < \" using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have neg: "positive_integral (\x. Real (- g x)) < \" . + finally have neg: "(\\<^isup>+ x. Real (- g x)) < \" . from neg pos borel show ?thesis unfolding integrable_def by auto @@ -1908,10 +1908,10 @@ lemma (in measure_space) integral_triangle_inequality: assumes "integrable f" - shows "\integral f\ \ integral (\x. \f x\)" + shows "\integral f\ \ (\x. \f x\)" proof - have "\integral f\ = max (integral f) (- integral f)" by auto - also have "\ \ integral (\x. \f x\)" + also have "\ \ (\x. \f x\)" using assms integral_minus(2)[of f, symmetric] by (auto intro!: integral_mono integrable_abs simp del: integral_minus) finally show ?thesis . @@ -1921,7 +1921,7 @@ assumes "integrable f" "\x. x \ space M \ 0 \ f x" shows "0 \ integral f" proof - - have "0 = integral (\x. 0)" by (auto simp: integral_zero) + have "0 = (\x. 0)" by (auto simp: integral_zero) also have "\ \ integral f" using assms by (rule integral_mono[OF integral_zero(1)]) finally show ?thesis . @@ -1953,7 +1953,7 @@ hence borel_u: "u \ borel_measurable M" using pos_u by (auto simp: borel_measurable_Real_eq SUP_F) - have integral_eq: "\n. positive_integral (\x. Real (f n x)) = Real (integral (f n))" + have integral_eq: "\n. (\\<^isup>+ x. Real (f n x)) = Real (integral (f n))" using i unfolding integral_def integrable_def by (auto simp: Real_real) have pos_integral: "\n. 0 \ integral (f n)" @@ -1961,14 +1961,14 @@ hence "0 \ x" using LIMSEQ_le_const[OF ilim, of 0] by auto - have "(\i. positive_integral (\x. Real (f i x))) \ positive_integral (\x. Real (u x))" + have "(\i. (\\<^isup>+ x. Real (f i x))) \ (\\<^isup>+ x. Real (u x))" proof (rule positive_integral_isoton) from SUP_F mono pos show "(\i x. Real (f i x)) \ (\x. Real (u x))" unfolding isoton_fun_expand by (auto simp: isoton_def mono_def) qed (rule borel_f) - hence pI: "positive_integral (\x. Real (u x)) = - (SUP n. positive_integral (\x. Real (f n x)))" + hence pI: "(\\<^isup>+ x. Real (u x)) = + (SUP n. (\\<^isup>+ x. Real (f n x)))" unfolding isoton_def by simp also have "\ = Real x" unfolding integral_eq proof (rule SUP_eq_LIMSEQ[THEN iffD2]) @@ -1997,7 +1997,7 @@ unfolding mono_def le_fun_def by (auto simp: field_simps) have 4: "\x. (\i. f i x - f 0 x) ----> u x - f 0 x" using lim by (auto intro!: LIMSEQ_diff) - have 5: "(\i. integral (\x. f i x - f 0 x)) ----> x - integral (f 0)" + have 5: "(\i. (\x. f i x - f 0 x)) ----> x - integral (f 0)" using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff) note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5] have "integrable (\x. (u x - f 0 x) + f 0 x)" @@ -2008,12 +2008,12 @@ lemma (in measure_space) integral_0_iff: assumes "integrable f" - shows "integral (\x. \f x\) = 0 \ \ {x\space M. f x \ 0} = 0" + shows "(\x. \f x\) = 0 \ \ {x\space M. f x \ 0} = 0" proof - have *: "\x. Real (- \f x\) = 0" by auto have "integrable (\x. \f x\)" using assms by (rule integrable_abs) hence "(\x. Real (\f x\)) \ borel_measurable M" - "positive_integral (\x. Real \f x\) \ \" unfolding integrable_def by auto + "(\\<^isup>+ x. Real \f x\) \ \" unfolding integrable_def by auto from positive_integral_0_iff[OF this(1)] this(2) show ?thesis unfolding integral_def * by (simp add: real_of_pextreal_eq_0) @@ -2024,7 +2024,7 @@ and "positive_integral f \ \" shows "\ (f -` {\} \ space M) = 0" proof - - have "\ * \ (f -` {\} \ space M) = positive_integral (\x. \ * indicator (f -` {\} \ space M) x)" + have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x)" using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \ \] by simp also have "\ \ positive_integral f" by (auto intro!: positive_integral_mono simp: indicator_def) @@ -2054,15 +2054,15 @@ lemma (in measure_space) integral_real: fixes f :: "'a \ pextreal" assumes "AE x. f x \ \" - shows "integral (\x. real (f x)) = real (positive_integral f)" (is ?plus) - and "integral (\x. - real (f x)) = - real (positive_integral f)" (is ?minus) + shows "(\x. real (f x)) = real (positive_integral f)" (is ?plus) + and "(\x. - real (f x)) = - real (positive_integral f)" (is ?minus) proof - - have "positive_integral (\x. Real (real (f x))) = positive_integral f" + have "(\\<^isup>+ x. Real (real (f x))) = positive_integral f" apply (rule positive_integral_cong_AE) apply (rule AE_mp[OF assms(1)]) by (auto intro!: AE_cong simp: Real_real) moreover - have "positive_integral (\x. Real (- real (f x))) = positive_integral (\x. 0)" + have "(\\<^isup>+ x. Real (- real (f x))) = (\\<^isup>+ x. 0)" by (intro positive_integral_cong) auto ultimately show ?plus ?minus by (auto simp: integral_def integrable_def) @@ -2073,7 +2073,7 @@ and w: "integrable w" "\x. x \ space M \ 0 \ w x" and u': "\x. x \ space M \ (\i. u i x) ----> u' x" shows "integrable u'" - and "(\i. integral (\x. \u i x - u' x\)) ----> 0" (is "?lim_diff") + and "(\i. (\x. \u i x - u' x\)) ----> 0" (is "?lim_diff") and "(\i. integral (u i)) ----> integral u'" (is ?lim) proof - { fix x j assume x: "x \ space M" @@ -2108,31 +2108,31 @@ finally have "\u j x - u' x\ \ 2 * w x" by simp } note diff_less_2w = this - have PI_diff: "\m n. positive_integral (\x. Real (?diff (m + n) x)) = - positive_integral (\x. Real (2 * w x)) - positive_integral (\x. Real \u (m + n) x - u' x\)" + have PI_diff: "\m n. (\\<^isup>+ x. Real (?diff (m + n) x)) = + (\\<^isup>+ x. Real (2 * w x)) - (\\<^isup>+ x. Real \u (m + n) x - u' x\)" using diff w diff_less_2w by (subst positive_integral_diff[symmetric]) (auto simp: integrable_def intro!: positive_integral_cong) have "integrable (\x. 2 * w x)" using w by (auto intro: integral_cmult) - hence I2w_fin: "positive_integral (\x. Real (2 * w x)) \ \" and + hence I2w_fin: "(\\<^isup>+ x. Real (2 * w x)) \ \" and borel_2w: "(\x. Real (2 * w x)) \ borel_measurable M" unfolding integrable_def by auto - have "(INF n. SUP m. positive_integral (\x. Real \u (m + n) x - u' x\)) = 0" (is "?lim_SUP = 0") + have "(INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\)) = 0" (is "?lim_SUP = 0") proof cases - assume eq_0: "positive_integral (\x. Real (2 * w x)) = 0" - have "\i. positive_integral (\x. Real \u i x - u' x\) \ positive_integral (\x. Real (2 * w x))" + assume eq_0: "(\\<^isup>+ x. Real (2 * w x)) = 0" + have "\i. (\\<^isup>+ x. Real \u i x - u' x\) \ (\\<^isup>+ x. Real (2 * w x))" proof (rule positive_integral_mono) fix i x assume "x \ space M" from diff_less_2w[OF this, of i] show "Real \u i x - u' x\ \ Real (2 * w x)" by auto qed - hence "\i. positive_integral (\x. Real \u i x - u' x\) = 0" using eq_0 by auto + hence "\i. (\\<^isup>+ x. Real \u i x - u' x\) = 0" using eq_0 by auto thus ?thesis by simp next - assume neq_0: "positive_integral (\x. Real (2 * w x)) \ 0" - have "positive_integral (\x. Real (2 * w x)) = positive_integral (\x. SUP n. INF m. Real (?diff (m + n) x))" + assume neq_0: "(\\<^isup>+ x. Real (2 * w x)) \ 0" + have "(\\<^isup>+ x. Real (2 * w x)) = (\\<^isup>+ x. SUP n. INF m. Real (?diff (m + n) x))" proof (rule positive_integral_cong, subst add_commute) fix x assume x: "x \ space M" show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))" @@ -2144,22 +2144,22 @@ thus "(\i. ?diff i x) ----> 2 * w x" by simp qed qed - also have "\ \ (SUP n. INF m. positive_integral (\x. Real (?diff (m + n) x)))" + also have "\ \ (SUP n. INF m. (\\<^isup>+ x. Real (?diff (m + n) x)))" using u'_borel w u unfolding integrable_def by (auto intro!: positive_integral_lim_INF) - also have "\ = positive_integral (\x. Real (2 * w x)) - - (INF n. SUP m. positive_integral (\x. Real \u (m + n) x - u' x\))" + also have "\ = (\\<^isup>+ x. Real (2 * w x)) - + (INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\))" unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus .. finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0) qed have [simp]: "\n m x. Real (- \u (m + n) x - u' x\) = 0" by auto - have [simp]: "\n m. positive_integral (\x. Real \u (m + n) x - u' x\) = - Real (integral (\x. \u (n + m) x - u' x\))" + have [simp]: "\n m. (\\<^isup>+ x. Real \u (m + n) x - u' x\) = + Real ((\x. \u (n + m) x - u' x\))" using diff by (subst add_commute) (simp add: integral_def integrable_def Real_real) - have "(SUP n. INF m. positive_integral (\x. Real \u (m + n) x - u' x\)) \ ?lim_SUP" + have "(SUP n. INF m. (\\<^isup>+ x. Real \u (m + n) x - u' x\)) \ ?lim_SUP" (is "?lim_INF \ _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP) hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP) @@ -2168,15 +2168,15 @@ proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from LIMSEQ_D[OF `?lim_diff` this] - obtain N where N: "\n. N \ n \ integral (\x. \u n x - u' x\) < r" + obtain N where N: "\n. N \ n \ (\x. \u n x - u' x\) < r" using diff by (auto simp: integral_positive) show "\N. \n\N. norm (integral (u n) - integral u') < r" proof (safe intro!: exI[of _ N]) fix n assume "N \ n" - have "\integral (u n) - integral u'\ = \integral (\x. u n x - u' x)\" + have "\integral (u n) - integral u'\ = \(\x. u n x - u' x)\" using u `integrable u'` by (auto simp: integral_diff) - also have "\ \ integral (\x. \u n x - u' x\)" using u `integrable u'` + also have "\ \ (\x. \u n x - u' x\)" using u `integrable u'` by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff) also note N[OF `N \ n`] finally show "norm (integral (u n) - integral u') < r" by simp @@ -2187,9 +2187,9 @@ lemma (in measure_space) integral_sums: assumes borel: "\i. integrable (f i)" and summable: "\x. x \ space M \ summable (\i. \f i x\)" - and sums: "summable (\i. integral (\x. \f i x\))" + and sums: "summable (\i. (\x. \f i x\))" shows "integrable (\x. (\i. f i x))" (is "integrable ?S") - and "(\i. integral (f i)) sums integral (\x. (\i. f i x))" (is ?integral) + and "(\i. integral (f i)) sums (\x. (\i. f i x))" (is ?integral) proof - have "\x\space M. \w. (\i. \f i x\) sums w" using summable unfolding summable_def by auto @@ -2198,7 +2198,7 @@ let "?w y" = "if y \ space M then w y else 0" - obtain x where abs_sum: "(\i. integral (\x. \f i x\)) sums x" + obtain x where abs_sum: "(\i. (\x. \f i x\)) sums x" using sums unfolding summable_def .. have 1: "\n. integrable (\x. \i = 0..n. ?w' n x) ----> ?w x" using w by (cases "x \ space M") (simp_all add: LIMSEQ_const sums_def) } - have *: "\n. integral (?w' n) = (\i = 0..< n. integral (\x. \f i x\))" + have *: "\n. integral (?w' n) = (\i = 0..< n. (\x. \f i x\))" using borel by (simp add: integral_setsum integrable_abs cong: integral_cong) from abs_sum show "(\i. integral (?w' i)) ----> x" unfolding * sums_def . @@ -2275,15 +2275,15 @@ by (auto intro!: sums_single simp: F F_abs) } note F_sums_f = this(1) and F_abs_sums_f = this(2) - have int_f: "integral f = integral (\x. \r. ?F r x)" "integrable f = integrable (\x. \r. ?F r x)" + have int_f: "integral f = (\x. \r. ?F r x)" "integrable f = integrable (\x. \r. ?F r x)" using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff) { fix r - have "integral (\x. \?F r x\) = integral (\x. \enum r\ * indicator (?A r) x)" + have "(\x. \?F r x\) = (\x. \enum r\ * indicator (?A r) x)" by (auto simp: indicator_def intro!: integral_cong) also have "\ = \enum r\ * real (\ (?A r))" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) - finally have "integral (\x. \?F r x\) = \enum r * real (\ (?A r))\" + finally have "(\x. \?F r x\) = \enum r * real (\ (?A r))\" by (simp add: abs_mult_pos real_pextreal_pos) } note int_abs_F = this @@ -2306,7 +2306,7 @@ assumes f: "f \ borel_measurable M" and finite: "finite (f`space M)" and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" shows "integrable f" - and "integral (\x. f x) = + and "(\x. f x) = (\ r \ f`space M. r * real (\ (f -` {r} \ space M)))" (is "?integral") proof - let "?A r" = "f -` {r} \ space M" @@ -2336,7 +2336,7 @@ lemma (in finite_measure_space) positive_integral_finite_eq_setsum: "positive_integral f = (\x \ space M. f x * \ {x})" proof - - have *: "positive_integral f = positive_integral (\x. \y\space M. f y * indicator {y} x)" + have *: "positive_integral f = (\\<^isup>+ x. \y\space M. f y * indicator {y} x)" by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space]) show ?thesis unfolding * using borel_measurable_finite[of f] by (simp add: positive_integral_setsum positive_integral_cmult_indicator) @@ -2347,8 +2347,8 @@ and "integral f = (\x \ space M. f x * real (\ {x}))" (is ?I) proof - have [simp]: - "positive_integral (\x. Real (f x)) = (\x \ space M. Real (f x) * \ {x})" - "positive_integral (\x. Real (- f x)) = (\x \ space M. Real (- f x) * \ {x})" + "(\\<^isup>+ x. Real (f x)) = (\x \ space M. Real (f x) * \ {x})" + "(\\<^isup>+ x. Real (- f x)) = (\x \ space M. Real (- f x) * \ {x})" unfolding positive_integral_finite_eq_setsum by auto show "integrable f" using finite_space finite_measure by (simp add: setsum_\ integrable_def)