diff -r f9ff311992b6 -r 3e39b0e730d6 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Wed Feb 02 10:35:41 2011 +0100 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Wed Feb 02 12:34:45 2011 +0100 @@ -7,6 +7,7 @@ begin lemma sums_If_finite: + fixes f :: "nat \ 'a::real_normed_vector" assumes finite: "finite {r. P r}" shows "(\r. if P r then f r else 0) sums (\r\{r. P r}. f r)" (is "?F sums _") proof cases @@ -24,7 +25,8 @@ qed lemma sums_single: - "(\r. if r = i then f r else 0) sums f i" + fixes f :: "nat \ 'a::real_normed_vector" + shows "(\r. if r = i then f r else 0) sums f i" using sums_If_finite[of "\r. r = i" f] by simp section "Simple function" @@ -37,12 +39,12 @@ *} -definition (in sigma_algebra) "simple_function g \ +definition "simple_function M g \ finite (g ` space M) \ (\x \ g ` space M. g -` {x} \ space M \ sets M)" lemma (in sigma_algebra) simple_functionD: - assumes "simple_function g" + assumes "simple_function M g" shows "finite (g ` space M)" and "g -` X \ space M \ sets M" proof - show "finite (g ` space M)" @@ -55,7 +57,7 @@ lemma (in sigma_algebra) simple_function_indicator_representation: fixes f ::"'a \ pextreal" - assumes f: "simple_function f" and x: "x \ space M" + assumes f: "simple_function M f" and x: "x \ space M" shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" (is "?l = ?r") proof - @@ -69,7 +71,7 @@ qed lemma (in measure_space) simple_function_notspace: - "simple_function (\x. h x * indicator (- space M) x::pextreal)" (is "simple_function ?h") + "simple_function M (\x. h x * indicator (- space M) x::pextreal)" (is "simple_function M ?h") proof - have "?h ` space M \ {0}" unfolding indicator_def by auto hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) @@ -79,7 +81,7 @@ lemma (in sigma_algebra) simple_function_cong: assumes "\t. t \ space M \ f t = g t" - shows "simple_function f \ simple_function g" + shows "simple_function M f \ simple_function M g" proof - have "f ` space M = g ` space M" "\x. f -` {x} \ space M = g -` {x} \ space M" @@ -87,15 +89,21 @@ thus ?thesis unfolding simple_function_def using assms by simp qed +lemma (in sigma_algebra) simple_function_cong_algebra: + assumes "sets N = sets M" "space N = space M" + shows "simple_function M f \ simple_function N f" + unfolding simple_function_def assms .. + lemma (in sigma_algebra) borel_measurable_simple_function: - assumes "simple_function f" + assumes "simple_function M f" shows "f \ borel_measurable M" proof (rule borel_measurableI) fix S let ?I = "f ` (f -` S \ space M)" have *: "(\x\?I. f -` {x} \ space M) = f -` S \ space M" (is "?U = _") by auto have "finite ?I" - using assms unfolding simple_function_def by (auto intro: finite_subset) + using assms unfolding simple_function_def + using finite_subset[of "f ` (f -` S \ space M)" "f ` space M"] by auto hence "?U \ sets M" apply (rule finite_UN) using assms unfolding simple_function_def by auto @@ -105,17 +113,17 @@ lemma (in sigma_algebra) simple_function_borel_measurable: fixes f :: "'a \ 'x::t2_space" assumes "f \ borel_measurable M" and "finite (f ` space M)" - shows "simple_function f" + shows "simple_function M f" using assms unfolding simple_function_def by (auto intro: borel_measurable_vimage) lemma (in sigma_algebra) simple_function_const[intro, simp]: - "simple_function (\x. c)" + "simple_function M (\x. c)" by (auto intro: finite_subset simp: simple_function_def) lemma (in sigma_algebra) simple_function_compose[intro, simp]: - assumes "simple_function f" - shows "simple_function (g \ f)" + assumes "simple_function M f" + shows "simple_function M (g \ f)" unfolding simple_function_def proof safe show "finite ((g \ f) ` space M)" @@ -132,7 +140,7 @@ lemma (in sigma_algebra) simple_function_indicator[intro, simp]: assumes "A \ sets M" - shows "simple_function (indicator A)" + shows "simple_function M (indicator A)" proof - have "indicator A ` space M \ {0, 1}" (is "?S \ _") by (auto simp: indicator_def) @@ -143,9 +151,9 @@ qed lemma (in sigma_algebra) simple_function_Pair[intro, simp]: - assumes "simple_function f" - assumes "simple_function g" - shows "simple_function (\x. (f x, g x))" (is "simple_function ?p") + assumes "simple_function M f" + assumes "simple_function M g" + shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") unfolding simple_function_def proof safe show "finite (?p ` space M)" @@ -161,16 +169,16 @@ qed lemma (in sigma_algebra) simple_function_compose1: - assumes "simple_function f" - shows "simple_function (\x. g (f x))" + assumes "simple_function M f" + shows "simple_function M (\x. g (f x))" using simple_function_compose[OF assms, of g] by (simp add: comp_def) lemma (in sigma_algebra) simple_function_compose2: - assumes "simple_function f" and "simple_function g" - shows "simple_function (\x. h (f x) (g x))" + assumes "simple_function M f" and "simple_function M g" + shows "simple_function M (\x. h (f x) (g x))" proof - - have "simple_function ((\(x, y). h x y) \ (\x. (f x, g x)))" + have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" using assms by auto thus ?thesis by (simp_all add: comp_def) qed @@ -183,14 +191,14 @@ and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] lemma (in sigma_algebra) simple_function_setsum[intro, simp]: - assumes "\i. i \ P \ simple_function (f i)" - shows "simple_function (\x. \i\P. f i x)" + assumes "\i. i \ P \ simple_function M (f i)" + shows "simple_function M (\x. \i\P. f i x)" proof cases assume "finite P" from this assms show ?thesis by induct auto qed auto lemma (in sigma_algebra) simple_function_le_measurable: - assumes "simple_function f" "simple_function g" + assumes "simple_function M f" "simple_function M g" shows "{x \ space M. f x \ g x} \ sets M" proof - have *: "{x \ space M. f x \ g x} = @@ -214,7 +222,7 @@ lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ pextreal" assumes u: "u \ borel_measurable M" - shows "\f. (\i. simple_function (f i) \ (\x\space M. f i x \ \)) \ f \ u" + shows "\f. (\i. simple_function M (f i) \ (\x\space M. f i x \ \)) \ f \ u" proof - have "\f. \x j. (of_nat j \ u x \ f x j = j*2^j) \ (u x < of_nat j \ of_nat (f x j) \ u x * 2^j \ u x * 2^j < of_nat (Suc (f x j)))" @@ -406,10 +414,10 @@ lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence': fixes u :: "'a \ pextreal" assumes "u \ borel_measurable M" - obtains (x) f where "f \ u" "\i. simple_function (f i)" "\i. \\f i`space M" + obtains (x) f where "f \ u" "\i. simple_function M (f i)" "\i. \\f i`space M" proof - from borel_measurable_implies_simple_function_sequence[OF assms] - obtain f where x: "\i. simple_function (f i)" "f \ u" + obtain f where x: "\i. simple_function M (f i)" "f \ u" and fin: "\i. \x. x\space M \ f i x \ \" by auto { fix i from fin[of _ i] have "\ \ f i`space M" by fastsimp } with x show thesis by (auto intro!: that[of f]) @@ -417,7 +425,7 @@ lemma (in sigma_algebra) simple_function_eq_borel_measurable: fixes f :: "'a \ pextreal" - shows "simple_function f \ + shows "simple_function M f \ finite (f`space M) \ f \ borel_measurable M" using simple_function_borel_measurable[of f] borel_measurable_simple_function[of f] @@ -425,8 +433,8 @@ lemma (in measure_space) simple_function_restricted: fixes f :: "'a \ pextreal" assumes "A \ sets M" - shows "sigma_algebra.simple_function (restricted_space A) f \ simple_function (\x. f x * indicator A x)" - (is "sigma_algebra.simple_function ?R f \ simple_function ?f") + shows "simple_function (restricted_space A) f \ simple_function M (\x. f x * indicator A x)" + (is "simple_function ?R f \ simple_function M ?f") proof - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) have "finite (f`A) \ finite (?f`space M)" @@ -463,29 +471,26 @@ qed lemma (in sigma_algebra) simple_function_subalgebra: - assumes "sigma_algebra.simple_function N f" - and N_subalgebra: "sets N \ sets M" "space N = space M" "sigma_algebra N" - shows "simple_function f" - using assms - unfolding simple_function_def - unfolding sigma_algebra.simple_function_def[OF N_subalgebra(3)] - by auto + assumes "simple_function N f" + and N_subalgebra: "sets N \ sets M" "space N = space M" + shows "simple_function M f" + using assms unfolding simple_function_def by auto lemma (in measure_space) simple_function_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" - and f: "sigma_algebra.simple_function M' f" - shows "simple_function (\x. f (T x))" + and f: "simple_function M' f" + shows "simple_function M (\x. f (T x))" proof (intro simple_function_def[THEN iffD2] conjI ballI) interpret T: sigma_algebra M' by fact have "(\x. f (T x)) ` space M \ f ` space M'" using T unfolding measurable_def by auto then show "finite ((\x. f (T x)) ` space M)" - using f unfolding T.simple_function_def by (auto intro: finite_subset) + using f unfolding simple_function_def by (auto intro: finite_subset) fix i assume i: "i \ (\x. f (T x)) ` space M" then have "i \ f ` space M'" using T unfolding measurable_def by auto then have "f -` {i} \ space M' \ sets M'" - using f unfolding T.simple_function_def by auto + using f unfolding simple_function_def by auto then have "T -` (f -` {i} \ space M') \ space M \ sets M" using T unfolding measurable_def by auto also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" @@ -495,12 +500,18 @@ section "Simple integral" -definition (in measure_space) simple_integral (binder "\\<^isup>S " 10) where - "simple_integral f = (\x \ f ` space M. x * \ (f -` {x} \ space M))" +definition simple_integral_def: + "integral\<^isup>S M f = (\x \ f ` space M. x * measure M (f -` {x} \ space M))" + +syntax + "_simple_integral" :: "'a \ ('a \ pextreal) \ ('a, 'b) measure_space_scheme \ pextreal" ("\\<^isup>S _. _ \_" [60,61] 110) + +translations + "\\<^isup>S x. f \M" == "CONST integral\<^isup>S M (%x. f)" lemma (in measure_space) simple_integral_cong: assumes "\t. t \ space M \ f t = g t" - shows "simple_integral f = simple_integral g" + shows "integral\<^isup>S M f = integral\<^isup>S M g" proof - have "f ` space M = g ` space M" "\x. f -` {x} \ space M = g -` {x} \ space M" @@ -509,18 +520,18 @@ qed lemma (in measure_space) simple_integral_cong_measure: - assumes "\A. A \ sets M \ \ A = \ A" and "simple_function f" - shows "measure_space.simple_integral M \ f = simple_integral f" + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + and "simple_function M f" + shows "integral\<^isup>S N f = integral\<^isup>S M f" proof - - interpret v: measure_space M \ - by (rule measure_space_cong) fact - from simple_functionD[OF `simple_function f`] assms show ?thesis - unfolding simple_integral_def v.simple_integral_def - by (auto intro!: setsum_cong) + interpret v: measure_space N + by (rule measure_space_cong) fact+ + from simple_functionD[OF `simple_function M f`] assms show ?thesis + by (auto intro!: setsum_cong simp: simple_integral_def) qed lemma (in measure_space) simple_integral_const[simp]: - "(\\<^isup>Sx. c) = c * \ (space M)" + "(\\<^isup>Sx. c \M) = c * \ (space M)" proof (cases "space M = {}") case True thus ?thesis unfolding simple_integral_def by simp next @@ -529,8 +540,8 @@ qed lemma (in measure_space) simple_function_partition: - assumes "simple_function f" and "simple_function g" - shows "simple_integral f = (\A\(\x. f -` {f x} \ g -` {g x} \ space M) ` space M. the_elem (f`A) * \ A)" + assumes "simple_function M f" and "simple_function M g" + shows "integral\<^isup>S M f = (\A\(\x. f -` {f x} \ g -` {g x} \ space M) ` space M. the_elem (f`A) * \ A)" (is "_ = setsum _ (?p ` space M)") proof- let "?sub x" = "?p ` (f -` {x} \ space M)" @@ -561,7 +572,7 @@ ultimately have "\ (f -` {f x} \ space M) = setsum (\) (?sub (f x))" by (subst measure_finitely_additive) auto } - hence "simple_integral f = (\(x,A)\?SIGMA. x * \ A)" + hence "integral\<^isup>S M f = (\(x,A)\?SIGMA. x * \ A)" unfolding simple_integral_def by (subst setsum_Sigma[symmetric], auto intro!: setsum_cong simp: setsum_right_distrib[symmetric]) @@ -584,8 +595,8 @@ qed lemma (in measure_space) simple_integral_add[simp]: - assumes "simple_function f" and "simple_function g" - shows "(\\<^isup>Sx. f x + g x) = simple_integral f + simple_integral g" + assumes "simple_function M f" and "simple_function M g" + shows "(\\<^isup>Sx. f x + g x \M) = integral\<^isup>S M f + integral\<^isup>S M g" proof - { fix x let ?S = "g -` {g x} \ f -` {f x} \ space M" assume "x \ space M" @@ -595,15 +606,15 @@ thus ?thesis unfolding simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]] - simple_function_partition[OF `simple_function f` `simple_function g`] - simple_function_partition[OF `simple_function g` `simple_function f`] + simple_function_partition[OF `simple_function M f` `simple_function M g`] + simple_function_partition[OF `simple_function M g` `simple_function M f`] apply (subst (3) Int_commute) by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong) qed lemma (in measure_space) simple_integral_setsum[simp]: - assumes "\i. i \ P \ simple_function (f i)" - shows "(\\<^isup>Sx. \i\P. f i x) = (\i\P. simple_integral (f i))" + assumes "\i. i \ P \ simple_function M (f i)" + shows "(\\<^isup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^isup>S M (f i))" proof cases assume "finite P" from this assms show ?thesis @@ -611,8 +622,8 @@ qed auto lemma (in measure_space) simple_integral_mult[simp]: - assumes "simple_function f" - shows "(\\<^isup>Sx. c * f x) = c * simple_integral f" + assumes "simple_function M f" + shows "(\\<^isup>Sx. c * f x \M) = c * integral\<^isup>S M 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" @@ -626,8 +637,8 @@ qed lemma (in sigma_algebra) simple_function_If: - assumes sf: "simple_function f" "simple_function g" and A: "A \ sets M" - shows "simple_function (\x. if x \ A then f x else g x)" (is "simple_function ?IF") + assumes sf: "simple_function M f" "simple_function M g" and A: "A \ sets M" + shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") proof - def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" show ?thesis unfolding simple_function_def @@ -648,17 +659,17 @@ qed lemma (in measure_space) simple_integral_mono_AE: - assumes "simple_function f" and "simple_function g" + assumes "simple_function M f" and "simple_function M g" and mono: "AE x. f x \ g x" - shows "simple_integral f \ simple_integral g" + shows "integral\<^isup>S M f \ integral\<^isup>S M g" proof - let "?S x" = "(g -` {g x} \ space M) \ (f -` {f x} \ space M)" have *: "\x. g -` {g x} \ f -` {f x} \ space M = ?S x" "\x. f -` {f x} \ g -` {g x} \ space M = ?S x" by auto show ?thesis unfolding * - simple_function_partition[OF `simple_function f` `simple_function g`] - simple_function_partition[OF `simple_function g` `simple_function f`] + simple_function_partition[OF `simple_function M f` `simple_function M g`] + simple_function_partition[OF `simple_function M g` `simple_function M f`] proof (safe intro!: setsum_mono) fix x assume "x \ space M" then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto @@ -680,23 +691,23 @@ qed lemma (in measure_space) simple_integral_mono: - assumes "simple_function f" and "simple_function g" + assumes "simple_function M f" and "simple_function M g" and mono: "\ x. x \ space M \ f x \ g x" - shows "simple_integral f \ simple_integral g" + shows "integral\<^isup>S M f \ integral\<^isup>S M g" proof (rule simple_integral_mono_AE[OF assms(1, 2)]) show "AE x. f x \ g x" using mono by (rule AE_cong) auto qed lemma (in measure_space) simple_integral_cong_AE: - assumes "simple_function f" "simple_function g" and "AE x. f x = g x" - shows "simple_integral f = simple_integral g" + assumes "simple_function M f" "simple_function M g" and "AE x. f x = g x" + shows "integral\<^isup>S M f = integral\<^isup>S M g" using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) lemma (in measure_space) simple_integral_cong': - assumes sf: "simple_function f" "simple_function g" + assumes sf: "simple_function M f" "simple_function M g" and mea: "\ {x\space M. f x \ g x} = 0" - shows "simple_integral f = simple_integral g" + shows "integral\<^isup>S M f = integral\<^isup>S M g" proof (intro simple_integral_cong_AE sf AE_I) show "\ {x\space M. f x \ g x} = 0" by fact show "{x \ space M. f x \ g x} \ sets M" @@ -705,12 +716,12 @@ lemma (in measure_space) simple_integral_indicator: assumes "A \ sets M" - assumes "simple_function f" - shows "(\\<^isup>Sx. f x * indicator A x) = + assumes "simple_function M f" + shows "(\\<^isup>Sx. f x * indicator A x \M) = (\x \ f ` space M. x * \ (f -` {x} \ space M \ A))" proof cases assume "A = space M" - moreover hence "(\\<^isup>Sx. f x * indicator A x) = simple_integral f" + moreover hence "(\\<^isup>Sx. f x * indicator A x \M) = integral\<^isup>S M 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) @@ -726,7 +737,7 @@ next show "0 \ ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x]) qed - have *: "(\\<^isup>Sx. f x * indicator A x) = + have *: "(\\<^isup>Sx. f x * indicator A x \M) = (\x \ f ` space M \ {0}. x * \ (f -` {x} \ space M \ A))" unfolding simple_integral_def I proof (rule setsum_mono_zero_cong_left) @@ -752,7 +763,7 @@ lemma (in measure_space) simple_integral_indicator_only[simp]: assumes "A \ sets M" - shows "simple_integral (indicator A) = \ A" + shows "integral\<^isup>S M (indicator A) = \ A" proof cases assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto thus ?thesis unfolding simple_integral_def using `space M = {}` by auto @@ -765,22 +776,22 @@ qed lemma (in measure_space) simple_integral_null_set: - assumes "simple_function u" "N \ null_sets" - shows "(\\<^isup>Sx. u x * indicator N x) = 0" + assumes "simple_function M u" "N \ null_sets" + shows "(\\<^isup>Sx. u x * indicator N x \M) = 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 "(\\<^isup>Sx. u x * indicator N x) = (\\<^isup>Sx. 0)" + then have "(\\<^isup>Sx. u x * indicator N x \M) = (\\<^isup>Sx. 0 \M)" 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 = (\\<^isup>Sx. f x * indicator S x)" + assumes sf: "simple_function M f" and eq: "AE x. x \ S" and "S \ sets M" + shows "integral\<^isup>S M f = (\\<^isup>Sx. f x * indicator S x \M)" proof (rule simple_integral_cong_AE) - show "simple_function f" by fact - show "simple_function (\x. f x * indicator S x)" + show "simple_function M f" by fact + show "simple_function M (\x. f x * indicator S x)" using sf `S \ sets M` by auto from eq show "AE x. f x = f x * indicator S x" by (rule AE_mp) simp @@ -788,10 +799,9 @@ 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 = (\\<^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`]] + assumes sf: "simple_function M (\x. f x * indicator A x)" + shows "integral\<^isup>S (restricted_space A) f = (\\<^isup>Sx. f x * indicator A x \M)" + (is "_ = integral\<^isup>S M ?f") unfolding simple_integral_def proof (simp, safe intro!: setsum_mono_zero_cong_left) from sf show "finite (?f ` space M)" @@ -816,64 +826,71 @@ qed lemma (in measure_space) simple_integral_subalgebra: - assumes N: "measure_space N \" and [simp]: "space N = space M" - shows "measure_space.simple_integral N \ = simple_integral" - unfolding simple_integral_def_raw - unfolding measure_space.simple_integral_def_raw[OF N] by simp + assumes N: "measure_space N" and [simp]: "space N = space M" "measure N = measure M" + shows "integral\<^isup>S N = integral\<^isup>S M" + unfolding simple_integral_def_raw by simp lemma (in measure_space) simple_integral_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" - and f: "sigma_algebra.simple_function M' f" - shows "measure_space.simple_integral M' (\A. \ (T -` A \ space M)) f = (\\<^isup>S x. f (T x))" - (is "measure_space.simple_integral M' ?nu f = _") + and f: "simple_function M' f" + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + shows "integral\<^isup>S M' f = (\\<^isup>S x. f (T x) \M)" proof - - interpret T: measure_space M' ?nu using T by (rule measure_space_vimage) auto - show "T.simple_integral f = (\\<^isup>S x. f (T x))" - unfolding simple_integral_def T.simple_integral_def + interpret T: measure_space M' using \ by (rule measure_space_vimage[OF T]) + show "integral\<^isup>S M' f = (\\<^isup>S x. f (T x) \M)" + unfolding simple_integral_def proof (intro setsum_mono_zero_cong_right ballI) show "(\x. f (T x)) ` space M \ f ` space M'" using T unfolding measurable_def by auto show "finite (f ` space M')" - using f unfolding T.simple_function_def by auto + using f unfolding simple_function_def by auto next fix i assume "i \ f ` space M' - (\x. f (T x)) ` space M" then have "T -` (f -` {i} \ space M') \ space M = {}" by (auto simp: image_iff) - then show "i * \ (T -` (f -` {i} \ space M') \ space M) = 0" by simp + with f[THEN T.simple_functionD(2), THEN \] + show "i * T.\ (f -` {i} \ space M') = 0" by simp next fix i assume "i \ (\x. f (T x)) ` space M" then have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" using T unfolding measurable_def by auto - then show "i * \ (T -` (f -` {i} \ space M') \ space M) = i * \ ((\x. f (T x)) -` {i} \ space M)" + with f[THEN T.simple_functionD(2), THEN \] + show "i * T.\ (f -` {i} \ space M') = i * \ ((\x. f (T x)) -` {i} \ space M)" by auto qed qed -section "Continuous posititve integration" +section "Continuous positive integration" + +definition positive_integral_def: + "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ f}. integral\<^isup>S M g)" -definition (in measure_space) positive_integral (binder "\\<^isup>+ " 10) where - "positive_integral f = SUPR {g. simple_function g \ g \ f} simple_integral" +syntax + "_positive_integral" :: "'a \ ('a \ pextreal) \ ('a, 'b) measure_space_scheme \ pextreal" ("\\<^isup>+ _. _ \_" [60,61] 110) + +translations + "\\<^isup>+ x. f \M" == "CONST integral\<^isup>P M (%x. f)" -lemma (in measure_space) positive_integral_alt: - "positive_integral f = - (SUPR {g. simple_function g \ g \ f \ \ \ g`space M} simple_integral)" (is "_ = ?alt") +lemma (in measure_space) positive_integral_alt: "integral\<^isup>P M f = + (SUP g : {g. simple_function M g \ g \ f \ \ \ g`space M}. integral\<^isup>S M g)" + (is "_ = ?alt") proof (rule antisym SUP_leI) - show "positive_integral f \ ?alt" unfolding positive_integral_def + show "integral\<^isup>P M f \ ?alt" unfolding positive_integral_def proof (safe intro!: SUP_leI) - fix g assume g: "simple_function g" "g \ f" + fix g assume g: "simple_function M g" "g \ f" let ?G = "g -` {\} \ space M" - show "simple_integral g \ - SUPR {g. simple_function g \ g \ f \ \ \ g ` space M} simple_integral" - (is "simple_integral g \ SUPR ?A simple_integral") + show "integral\<^isup>S M g \ + (SUP h : {i. simple_function M i \ i \ f \ \ \ i ` space M}. integral\<^isup>S M h)" + (is "integral\<^isup>S M g \ SUPR ?A _") proof cases let ?g = "\x. indicator (space M - ?G) x * g x" - have g': "simple_function ?g" + have g': "simple_function M ?g" using g by (auto intro: simple_functionD) moreover assume "\ ?G = 0" then have "AE x. g x = ?g x" using g by (intro AE_I[where N="?G"]) (auto intro: simple_functionD simp: indicator_def) - with g(1) g' have "simple_integral g = simple_integral ?g" + with g(1) g' have "integral\<^isup>S M g = integral\<^isup>S M ?g" by (rule simple_integral_cong_AE) moreover have "?g \ g" by (auto simp: le_fun_def indicator_def) from this `g \ f` have "?g \ f" by (rule order_trans) @@ -885,15 +902,15 @@ then have "?G \ {}" by auto then have "\ \ g`space M" by force then have "space M \ {}" by auto - have "SUPR ?A simple_integral = \" + have "SUPR ?A (integral\<^isup>S M) = \" proof (intro SUP_\[THEN iffD2] allI impI) fix x assume "x < \" then guess n unfolding less_\_Ex_of_nat .. note n = this then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp let ?g = "\x. (of_nat n / (if \ ?G = \ then 1 else \ ?G)) * indicator ?G x" - show "\i\?A. x < simple_integral i" + show "\i\?A. x < integral\<^isup>S M i" proof (intro bexI impI CollectI conjI) - show "simple_function ?g" using g + show "simple_function M ?g" using g by (auto intro!: simple_functionD simple_function_add) have "?g \ g" by (auto simp: le_fun_def indicator_def) from this g(2) show "?g \ f" by (rule order_trans) @@ -902,10 +919,10 @@ have "x < (of_nat n / (if \ ?G = \ then 1 else \ ?G)) * \ ?G" using n `\ ?G \ 0` `0 < n` by (auto simp: pextreal_noteq_omega_Ex field_simps) - also have "\ = simple_integral ?g" using g `space M \ {}` + also have "\ = integral\<^isup>S M ?g" using g `space M \ {}` by (subst simple_integral_indicator) (auto simp: image_constant ac_simps dest: simple_functionD) - finally show "x < simple_integral ?g" . + finally show "x < integral\<^isup>S M ?g" . qed qed then show ?thesis by simp @@ -914,40 +931,41 @@ qed (auto intro!: SUP_subset simp: positive_integral_def) lemma (in measure_space) positive_integral_cong_measure: - assumes "\A. A \ sets M \ \ A = \ A" - shows "measure_space.positive_integral M \ f = positive_integral f" + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - - interpret v: measure_space M \ - by (rule measure_space_cong) fact + interpret v: measure_space N + by (rule measure_space_cong) fact+ with assms show ?thesis - unfolding positive_integral_def v.positive_integral_def SUPR_def + unfolding positive_integral_def SUPR_def by (auto intro!: arg_cong[where f=Sup] image_cong - simp: simple_integral_cong_measure[of \]) + simp: simple_integral_cong_measure[OF assms] + simple_function_cong_algebra[OF assms(2,3)]) qed lemma (in measure_space) positive_integral_alt1: - "positive_integral f = - (SUP g : {g. simple_function g \ (\x\space M. g x \ f x \ g x \ \)}. simple_integral g)" + "integral\<^isup>P M f = + (SUP g : {g. simple_function M g \ (\x\space M. g x \ f x \ g x \ \)}. integral\<^isup>S M g)" unfolding positive_integral_alt SUPR_def proof (safe intro!: arg_cong[where f=Sup]) fix g let ?g = "\x. if x \ space M then g x else f x" - assume "simple_function g" "\x\space M. g x \ f x \ g x \ \" - hence "?g \ f" "simple_function ?g" "simple_integral g = simple_integral ?g" + assume "simple_function M g" "\x\space M. g x \ f x \ g x \ \" + hence "?g \ f" "simple_function M ?g" "integral\<^isup>S M g = integral\<^isup>S M ?g" "\ \ g`space M" unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong) - thus "simple_integral g \ simple_integral ` {g. simple_function g \ g \ f \ \ \ g`space M}" + thus "integral\<^isup>S M g \ integral\<^isup>S M ` {g. simple_function M g \ g \ f \ \ \ g`space M}" by auto next - fix g assume "simple_function g" "g \ f" "\ \ g`space M" - hence "simple_function g" "\x\space M. g x \ f x \ g x \ \" + fix g assume "simple_function M g" "g \ f" "\ \ g`space M" + hence "simple_function M g" "\x\space M. g x \ f x \ g x \ \" by (auto simp add: le_fun_def image_iff) - thus "simple_integral g \ simple_integral ` {g. simple_function g \ (\x\space M. g x \ f x \ g x \ \)}" + thus "integral\<^isup>S M g \ integral\<^isup>S M ` {g. simple_function M g \ (\x\space M. g x \ f x \ g x \ \)}" by auto qed lemma (in measure_space) positive_integral_cong: assumes "\x. x \ space M \ f x = g x" - shows "positive_integral f = positive_integral g" + shows "integral\<^isup>P M f = integral\<^isup>P M g" proof - have "\h. (\x\space M. h x \ f x \ h x \ \) = (\x\space M. h x \ g x \ h x \ \)" using assms by auto @@ -955,30 +973,30 @@ qed lemma (in measure_space) positive_integral_eq_simple_integral: - assumes "simple_function f" - shows "positive_integral f = simple_integral f" + assumes "simple_function M f" + shows "integral\<^isup>P M f = integral\<^isup>S M f" unfolding positive_integral_def proof (safe intro!: pextreal_SUPI) - fix g assume "simple_function g" "g \ f" - with assms show "simple_integral g \ simple_integral f" + fix g assume "simple_function M g" "g \ f" + with assms show "integral\<^isup>S M g \ integral\<^isup>S M f" by (auto intro!: simple_integral_mono simp: le_fun_def) next - fix y assume "\x. x\{g. simple_function g \ g \ f} \ simple_integral x \ y" - with assms show "simple_integral f \ y" by auto + fix y assume "\x. x\{g. simple_function M g \ g \ f} \ integral\<^isup>S M x \ y" + with assms show "integral\<^isup>S M f \ y" by auto qed lemma (in measure_space) positive_integral_mono_AE: assumes ae: "AE x. u x \ v x" - shows "positive_integral u \ positive_integral v" + shows "integral\<^isup>P M u \ integral\<^isup>P M v" unfolding positive_integral_alt1 proof (safe intro!: SUPR_mono) - fix a assume a: "simple_function a" and mono: "\x\space M. a x \ u x \ a x \ \" + fix a assume a: "simple_function M a" and mono: "\x\space M. a x \ u x \ a x \ \" from ae obtain N where N: "{x\space M. \ u x \ v x} \ N" "N \ sets M" "\ N = 0" by (auto elim!: AE_E) - have "simple_function (\x. a x * indicator (space M - N) x)" + have "simple_function M (\x. a x * indicator (space M - N) x)" using `N \ sets M` a by auto - with a show "\b\{g. simple_function g \ (\x\space M. g x \ v x \ g x \ \)}. - simple_integral a \ simple_integral b" + with a show "\b\{g. simple_function M g \ (\x\space M. g x \ v x \ g x \ \)}. + integral\<^isup>S M a \ integral\<^isup>S M b" proof (safe intro!: bexI[of _ "\x. a x * indicator (space M - N) x"] simple_integral_mono_AE) show "AE x. a x \ a x * indicator (space M - N) x" @@ -987,7 +1005,7 @@ N \ {x \ space M. a x \ 0}" (is "?N = _") using `N \ sets M`[THEN sets_into_space] by (auto simp: indicator_def) then show "?N \ sets M" - using `N \ sets M` `simple_function a`[THEN borel_measurable_simple_function] + using `N \ sets M` `simple_function M a`[THEN borel_measurable_simple_function] by (auto intro!: measure_mono Int) then have "\ ?N \ \ N" unfolding * using `N \ sets M` by (auto intro!: measure_mono) @@ -1010,12 +1028,12 @@ qed lemma (in measure_space) positive_integral_cong_AE: - "AE x. u x = v x \ positive_integral u = positive_integral v" + "AE x. u x = v x \ integral\<^isup>P M u = integral\<^isup>P M v" by (auto simp: eq_iff intro!: positive_integral_mono_AE) lemma (in measure_space) positive_integral_mono: assumes mono: "\x. x \ space M \ u x \ v x" - shows "positive_integral u \ positive_integral v" + shows "integral\<^isup>P M u \ integral\<^isup>P M v" using mono by (auto intro!: AE_cong positive_integral_mono_AE) lemma image_set_cong: @@ -1027,15 +1045,15 @@ lemma (in measure_space) positive_integral_SUP_approx: assumes "f \ s" and f: "\i. f i \ borel_measurable M" - and "simple_function u" + and "simple_function M u" and le: "u \ s" and real: "\ \ u`space M" - shows "simple_integral u \ (SUP i. positive_integral (f i))" (is "_ \ ?S") + shows "integral\<^isup>S M u \ (SUP i. integral\<^isup>P M (f i))" (is "_ \ ?S") proof (rule pextreal_le_mult_one_interval) fix a :: pextreal assume "0 < a" "a < 1" hence "a \ 0" by auto let "?B i" = "{x \ space M. a * u x \ f i x}" have B: "\i. ?B i \ sets M" - using f `simple_function u` by (auto simp: borel_measurable_simple_function) + using f `simple_function M u` by (auto simp: borel_measurable_simple_function) let "?uB i x" = "u x * indicator (?B i) x" @@ -1049,7 +1067,7 @@ note B_mono = this have u: "\x. x \ space M \ u -` {u x} \ space M \ sets M" - using `simple_function u` by (auto simp add: simple_function_def) + using `simple_function M u` by (auto simp add: simple_function_def) have "\i. (\n. (u -` {i} \ space M) \ ?B n) \ (u -` {i} \ space M)" using B_mono unfolding isoton_def proof safe @@ -1071,8 +1089,8 @@ qed auto note measure_conv = measure_up[OF Int[OF u B] this] - have "simple_integral u = (SUP i. simple_integral (?uB i))" - unfolding simple_integral_indicator[OF B `simple_function u`] + have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))" + unfolding simple_integral_indicator[OF B `simple_function M u`] proof (subst SUPR_pextreal_setsum, safe) fix x n assume "x \ space M" have "\ (u -` {u x} \ space M \ {x \ space M. a * u x \ f n x}) @@ -1082,52 +1100,51 @@ \ u x * \ (u -` {u x} \ space M \ ?B (Suc n))" by (auto intro: mult_left_mono) next - show "simple_integral u = + show "integral\<^isup>S M u = (\i\u ` space M. SUP n. i * \ (u -` {i} \ space M \ ?B n))" using measure_conv unfolding simple_integral_def isoton_def by (auto intro!: setsum_cong simp: pextreal_SUP_cmult) qed moreover - have "a * (SUP i. simple_integral (?uB i)) \ ?S" + have "a * (SUP i. integral\<^isup>S M (?uB i)) \ ?S" unfolding pextreal_SUP_cmult[symmetric] proof (safe intro!: SUP_mono bexI) fix i - have "a * simple_integral (?uB i) = (\\<^isup>Sx. a * ?uB i x)" - using B `simple_function u` + have "a * integral\<^isup>S M (?uB i) = (\\<^isup>Sx. a * ?uB i x \M)" + using B `simple_function M u` by (subst simple_integral_mult[symmetric]) (auto simp: field_simps) - also have "\ \ positive_integral (f i)" + also have "\ \ integral\<^isup>P M (f i)" proof - have "\x. a * ?uB i x \ f i x" unfolding indicator_def by auto - hence *: "simple_function (\x. a * ?uB i x)" using B assms(3) + hence *: "simple_function M (\x. a * ?uB i x)" using B assms(3) by (auto intro!: simple_integral_mono) show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric] by (auto intro!: positive_integral_mono simp: indicator_def) qed - finally show "a * simple_integral (?uB i) \ positive_integral (f i)" + finally show "a * integral\<^isup>S M (?uB i) \ integral\<^isup>P M (f i)" by auto qed simp - ultimately show "a * simple_integral u \ ?S" by simp + ultimately show "a * integral\<^isup>S M u \ ?S" by simp qed text {* Beppo-Levi monotone convergence theorem *} lemma (in measure_space) positive_integral_isoton: assumes "f \ u" "\i. f i \ borel_measurable M" - shows "(\i. positive_integral (f i)) \ positive_integral u" + shows "(\i. integral\<^isup>P M (f i)) \ integral\<^isup>P M u" unfolding isoton_def proof safe - fix i show "positive_integral (f i) \ positive_integral (f (Suc i))" + fix i show "integral\<^isup>P M (f i) \ integral\<^isup>P M (f (Suc i))" apply (rule positive_integral_mono) using `f \ u` unfolding isoton_def le_fun_def by auto next have u: "u = (SUP i. f i)" using `f \ u` unfolding isoton_def by auto - - show "(SUP i. positive_integral (f i)) = positive_integral u" + show "(SUP i. integral\<^isup>P M (f i)) = integral\<^isup>P M u" proof (rule antisym) from `f \ u`[THEN isoton_Sup, unfolded le_fun_def] - show "(SUP j. positive_integral (f j)) \ positive_integral u" + show "(SUP j. integral\<^isup>P M (f j)) \ integral\<^isup>P M u" by (auto intro!: SUP_leI positive_integral_mono) next - show "positive_integral u \ (SUP i. positive_integral (f i))" + show "integral\<^isup>P M u \ (SUP i. integral\<^isup>P M (f i))" unfolding positive_integral_alt[of u] by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms]) qed @@ -1136,12 +1153,12 @@ 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)) = (\\<^isup>+ x. SUP i. f i x)" - (is "_ = positive_integral ?u") + shows "(SUP i. integral\<^isup>P M (f i)) = (\\<^isup>+ x. (SUP i. f i x) \M)" + (is "_ = integral\<^isup>P M ?u") proof - show ?thesis proof (rule antisym) - show "(SUP j. positive_integral (f j)) \ positive_integral ?u" + show "(SUP j. integral\<^isup>P M (f j)) \ integral\<^isup>P M ?u" by (auto intro!: SUP_leI positive_integral_mono le_SUPI) next def rf \ "\i. \x\space M. f i x" and ru \ "\x\space M. ?u x" @@ -1151,26 +1168,26 @@ unfolding isoton_def le_fun_def fun_eq_iff SUPR_apply using SUP_const[OF UNIV_not_empty] by (auto simp: restrict_def le_fun_def fun_eq_iff) - ultimately have "positive_integral ru \ (SUP i. positive_integral (rf i))" + ultimately have "integral\<^isup>P M ru \ (SUP i. integral\<^isup>P M (rf i))" unfolding positive_integral_alt[of ru] by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx) - then show "positive_integral ?u \ (SUP i. positive_integral (f i))" + then show "integral\<^isup>P M ?u \ (SUP i. integral\<^isup>P M (f i))" unfolding ru_def rf_def by (simp cong: positive_integral_cong) qed qed lemma (in measure_space) SUP_simple_integral_sequences: - assumes f: "f \ u" "\i. simple_function (f i)" - and g: "g \ u" "\i. simple_function (g i)" - shows "(SUP i. simple_integral (f i)) = (SUP i. simple_integral (g i))" + assumes f: "f \ u" "\i. simple_function M (f i)" + and g: "g \ u" "\i. simple_function M (g i)" + shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))" (is "SUPR _ ?F = SUPR _ ?G") proof - - have "(SUP i. ?F i) = (SUP i. positive_integral (f i))" + have "(SUP i. ?F i) = (SUP i. integral\<^isup>P M (f i))" using assms by (simp add: positive_integral_eq_simple_integral) - also have "\ = positive_integral u" + also have "\ = integral\<^isup>P M u" using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]] unfolding isoton_def by simp - also have "\ = (SUP i. positive_integral (g i))" + also have "\ = (SUP i. integral\<^isup>P M (g i))" using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]] unfolding isoton_def by simp also have "\ = (SUP i. ?G i)" @@ -1179,38 +1196,36 @@ qed lemma (in measure_space) positive_integral_const[simp]: - "(\\<^isup>+ x. c) = c * \ (space M)" + "(\\<^isup>+ x. c \M) = c * \ (space M)" by (subst positive_integral_eq_simple_integral) auto lemma (in measure_space) positive_integral_isoton_simple: - assumes "f \ u" and e: "\i. simple_function (f i)" - shows "(\i. simple_integral (f i)) \ positive_integral u" + assumes "f \ u" and e: "\i. simple_function M (f i)" + shows "(\i. integral\<^isup>S M (f i)) \ integral\<^isup>P M u" using positive_integral_isoton[OF `f \ u` e(1)[THEN borel_measurable_simple_function]] unfolding positive_integral_eq_simple_integral[OF e] . lemma (in measure_space) positive_integral_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" and f: "f \ borel_measurable M'" - shows "measure_space.positive_integral M' (\A. \ (T -` A \ space M)) f = (\\<^isup>+ x. f (T x))" - (is "measure_space.positive_integral M' ?nu f = _") + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + shows "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" proof - - interpret T: measure_space M' ?nu using T by (rule measure_space_vimage) auto - obtain f' where f': "f' \ f" "\i. T.simple_function (f' i)" + interpret T: measure_space M' using \ by (rule measure_space_vimage[OF T]) + obtain f' where f': "f' \ f" "\i. simple_function M' (f' i)" using T.borel_measurable_implies_simple_function_sequence[OF f] by blast - then have f: "(\i x. f' i (T x)) \ (\x. f (T x))" "\i. simple_function (\x. f' i (T x))" + then have f: "(\i x. f' i (T x)) \ (\x. f (T x))" "\i. simple_function M (\x. f' i (T x))" using simple_function_vimage[OF T] unfolding isoton_fun_expand by auto - show "T.positive_integral f = (\\<^isup>+ x. f (T x))" + show "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" using positive_integral_isoton_simple[OF f] using T.positive_integral_isoton_simple[OF f'] - unfolding simple_integral_vimage[OF T f'(2)] isoton_def - by simp + by (simp add: simple_integral_vimage[OF T f'(2) \] isoton_def) qed lemma (in measure_space) positive_integral_linear: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" - shows "(\\<^isup>+ x. a * f x + g x) = - a * positive_integral f + positive_integral g" - (is "positive_integral ?L = _") + shows "(\\<^isup>+ x. a * f x + g x \M) = a * integral\<^isup>P M f + integral\<^isup>P M g" + (is "integral\<^isup>P M ?L = _") proof - from borel_measurable_implies_simple_function_sequence'[OF f] guess u . note u = this positive_integral_isoton_simple[OF this(1-2)] @@ -1222,46 +1237,45 @@ using assms by simp from borel_measurable_implies_simple_function_sequence'[OF this] guess l . note positive_integral_isoton_simple[OF this(1-2)] and l = this - moreover have - "(SUP i. simple_integral (l i)) = (SUP i. simple_integral (?L' i))" + moreover have "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))" proof (rule SUP_simple_integral_sequences[OF l(1-2)]) - show "?L' \ ?L" "\i. simple_function (?L' i)" + show "?L' \ ?L" "\i. simple_function M (?L' i)" using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right) qed moreover from u v have L'_isoton: - "(\i. simple_integral (?L' i)) \ a * positive_integral f + positive_integral g" + "(\i. integral\<^isup>S M (?L' i)) \ a * integral\<^isup>P M f + integral\<^isup>P M g" by (simp add: isoton_add isoton_cmult_right) ultimately show ?thesis by (simp add: isoton_def) qed lemma (in measure_space) positive_integral_cmult: assumes "f \ borel_measurable M" - shows "(\\<^isup>+ x. c * f x) = c * positive_integral f" + shows "(\\<^isup>+ x. c * f x \M) = c * integral\<^isup>P M 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 "(\\<^isup>+ x. f x * c) = positive_integral f * c" + shows "(\\<^isup>+ x. f x * c \M) = integral\<^isup>P M 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 \ (\\<^isup>+ x. indicator A x) = \ A" + "A \ sets M \ (\\<^isup>+ x. indicator A x\M) = \ 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 \ (\\<^isup>+ x. c * indicator A x) = c * \ A" + "A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = 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 "(\\<^isup>+ x. f x + g x) = positive_integral f + positive_integral g" + shows "(\\<^isup>+ x. f x + g x \M) = integral\<^isup>P M f + integral\<^isup>P M 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 "(\\<^isup>+ x. \i\P. f i x) = (\i\P. positive_integral (f i))" + shows "(\\<^isup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^isup>P M (f i))" proof cases assume "finite P" from this assms show ?thesis @@ -1277,14 +1291,13 @@ lemma (in measure_space) positive_integral_diff: assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" - and fin: "positive_integral g \ \" + and fin: "integral\<^isup>P M g \ \" and mono: "\x. x \ space M \ g x \ f x" - shows "(\\<^isup>+ x. f x - g x) = positive_integral f - positive_integral g" + shows "(\\<^isup>+ x. f x - g x \M) = integral\<^isup>P M f - integral\<^isup>P M g" proof - have borel: "(\x. f x - g x) \ borel_measurable M" using f g by (rule borel_measurable_pextreal_diff) - have "(\\<^isup>+x. f x - g x) + positive_integral g = - positive_integral f" + have "(\\<^isup>+x. f x - g x \M) + integral\<^isup>P M g = integral\<^isup>P M f" unfolding positive_integral_add[OF borel g, symmetric] proof (rule positive_integral_cong) fix x assume "x \ space M" @@ -1297,9 +1310,9 @@ lemma (in measure_space) positive_integral_psuminf: assumes "\i. f i \ borel_measurable M" - shows "(\\<^isup>+ x. \\<^isub>\ i. f i x) = (\\<^isub>\ i. positive_integral (f i))" + shows "(\\<^isup>+ x. (\\<^isub>\ i. f i x) \M) = (\\<^isub>\ i. integral\<^isup>P M (f i))" proof - - have "(\i. (\\<^isup>+x. \i (\\<^isup>+x. \\<^isub>\i. f i x)" + have "(\i. (\\<^isup>+x. (\iM)) \ (\\<^isup>+x. (\\<^isub>\i. f i x) \M)" by (rule positive_integral_isoton) (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono arg_cong[where f=Sup] @@ -1312,79 +1325,86 @@ lemma (in measure_space) positive_integral_lim_INF: fixes u :: "nat \ 'a \ pextreal" assumes "\i. u i \ borel_measurable M" - shows "(\\<^isup>+ x. SUP n. INF m. u (m + n) x) \ - (SUP n. INF m. positive_integral (u (m + n)))" + shows "(\\<^isup>+ x. (SUP n. INF m. u (m + n) x) \M) \ + (SUP n. INF m. integral\<^isup>P M (u (m + n)))" proof - - have "(\\<^isup>+x. SUP n. INF m. u (m + n) x) - = (SUP n. (\\<^isup>+x. INF m. u (m + n) x))" + have "(\\<^isup>+x. (SUP n. INF m. u (m + n) x) \M) + = (SUP n. (\\<^isup>+x. (INF m. u (m + n) x) \M))" using assms by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono) (auto simp del: add_Suc simp add: add_Suc[symmetric]) - also have "\ \ (SUP n. INF m. positive_integral (u (m + n)))" + also have "\ \ (SUP n. INF m. integral\<^isup>P M (u (m + n)))" by (auto intro!: SUP_mono bexI le_INFI positive_integral_mono INF_leI) finally show ?thesis . qed lemma (in measure_space) measure_space_density: assumes borel: "u \ borel_measurable M" - 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" - unfolding countably_additive_def - proof safe - fix A :: "nat \ 'a set" - assume "range A \ sets M" - hence "\i. (\x. u x * indicator (A i) x) \ borel_measurable M" - using borel by (auto intro: borel_measurable_indicator) - moreover assume "disjoint_family A" - note psuminf_indicator[OF this] - ultimately show "(\\<^isub>\n. ?v (A n)) = ?v (\x. A x)" - by (simp add: positive_integral_psuminf[symmetric]) + and M'[simp]: "M' = (M\measure := \A. (\\<^isup>+ x. u x * indicator A x \M)\)" + shows "measure_space M'" +proof - + interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto + show ?thesis + proof + show "measure M' {} = 0" unfolding M' by simp + show "countably_additive M' (measure M')" + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume "range A \ sets M'" + then have "\i. (\x. u x * indicator (A i) x) \ borel_measurable M" + using borel by (auto intro: borel_measurable_indicator) + moreover assume "disjoint_family A" + note psuminf_indicator[OF this] + ultimately show "(\\<^isub>\n. measure M' (A n)) = measure M' (\x. A x)" + by (simp add: positive_integral_psuminf[symmetric]) + qed qed qed lemma (in measure_space) positive_integral_translated_density: assumes "f \ borel_measurable M" "g \ borel_measurable M" - 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 _ = _") + and M': "M' = (M\ measure := \A. (\\<^isup>+ x. f x * indicator A x \M)\)" + shows "integral\<^isup>P M' g = (\\<^isup>+ x. f x * g x \M)" proof - - from measure_space_density[OF assms(1)] - interpret T: measure_space M ?T . + from measure_space_density[OF assms(1) M'] + interpret T: measure_space M' . + have borel[simp]: + "borel_measurable M' = borel_measurable M" + "simple_function M' = simple_function M" + unfolding measurable_def simple_function_def_raw by (auto simp: M') from borel_measurable_implies_simple_function_sequence[OF assms(2)] - obtain G where G: "\i. simple_function (G i)" "G \ g" by blast + obtain G where G: "\i. simple_function M (G i)" "G \ g" by blast note G_borel = borel_measurable_simple_function[OF this(1)] - from T.positive_integral_isoton[OF `G \ g` G_borel] - have *: "(\i. T.positive_integral (G i)) \ T.positive_integral g" . + from T.positive_integral_isoton[unfolded borel, OF `G \ g` G_borel] + have *: "(\i. integral\<^isup>P M' (G i)) \ integral\<^isup>P M' g" . { fix i have [simp]: "finite (G i ` space M)" using G(1) unfolding simple_function_def by auto - have "T.positive_integral (G i) = T.simple_integral (G i)" + have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)" using G T.positive_integral_eq_simple_integral by simp - 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) + also have "\ = (\\<^isup>+x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x) \M)" + apply (simp add: simple_integral_def M') apply (subst positive_integral_cmult[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage) apply (subst positive_integral_setsum[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage) by (simp add: setsum_right_distrib field_simps) - also have "\ = (\\<^isup>+x. f x * G i x)" + also have "\ = (\\<^isup>+x. f x * G i x \M)" by (auto intro!: positive_integral_cong simp: indicator_def if_distrib setsum_cases) - 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 + finally have "integral\<^isup>P M' (G i) = (\\<^isup>+x. f x * G i x \M)" . } + with * have eq_Tg: "(\i. (\\<^isup>+x. f x * G i x \M)) \ integral\<^isup>P M' 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. (\\<^isup>+x. f x * G i x)) \ (\\<^isup>+x. f x * g x)" + then have "(\i. (\\<^isup>+x. f x * G i x \M)) \ (\\<^isup>+x. f x * g x \M)" using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times) - with eq_Tg show "T.positive_integral g = (\\<^isup>+x. f x * g x)" + with eq_Tg show "integral\<^isup>P M' g = (\\<^isup>+x. f x * g x \M)" unfolding isoton_def by simp qed lemma (in measure_space) positive_integral_null_set: - assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x) = 0" + assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x \M) = 0" proof - - have "(\\<^isup>+ x. u x * indicator N x) = (\\<^isup>+ x. 0)" + have "(\\<^isup>+ x. u x * indicator N x \M) = (\\<^isup>+ x. 0 \M)" proof (intro positive_integral_cong_AE AE_I) show "{x \ space M. u x * indicator N x \ 0} \ N" by (auto simp: indicator_def) @@ -1396,20 +1416,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 * (\\<^isup>+ x. u x * indicator A x)" + shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x \M)" (is "\ ?A \ _ * ?PI") proof - have "?A \ sets M" using `A \ sets M` borel by auto - hence "\ ?A = (\\<^isup>+ x. indicator ?A x)" + hence "\ ?A = (\\<^isup>+ x. indicator ?A x \M)" using positive_integral_indicator by simp - also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x))" + also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x) \M)" 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 * (\\<^isup>+ x. u x * indicator A x)" + also have "\ = c * (\\<^isup>+ x. u x * indicator A x \M)" using assms by (auto intro!: positive_integral_cmult borel_measurable_indicator) finally show ?thesis . @@ -1417,11 +1437,11 @@ lemma (in measure_space) positive_integral_0_iff: assumes borel: "u \ borel_measurable M" - shows "positive_integral u = 0 \ \ {x\space M. u x \ 0} = 0" + shows "integral\<^isup>P M u = 0 \ \ {x\space M. u x \ 0} = 0" (is "_ \ \ ?A = 0") proof - have A: "?A \ sets M" using borel by auto - have u: "(\\<^isup>+ x. u x * indicator ?A x) = positive_integral u" + have u: "(\\<^isup>+ x. u x * indicator ?A x \M) = integral\<^isup>P M u" by (auto intro!: positive_integral_cong simp: indicator_def) show ?thesis @@ -1429,10 +1449,10 @@ assume "\ ?A = 0" hence "?A \ null_sets" using `?A \ sets M` by auto from positive_integral_null_set[OF this] - have "0 = (\\<^isup>+ x. u x * indicator ?A x)" by simp - thus "positive_integral u = 0" unfolding u by simp + have "0 = (\\<^isup>+ x. u x * indicator ?A x \M)" by simp + thus "integral\<^isup>P M u = 0" unfolding u by simp next - assume *: "positive_integral u = 0" + assume *: "integral\<^isup>P M u = 0" let "?M n" = "{x \ space M. 1 \ of_nat n * u x}" have "0 = (SUP n. \ (?M n \ ?A))" proof - @@ -1469,34 +1489,34 @@ lemma (in measure_space) positive_integral_restricted: assumes "A \ sets M" - 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") + shows "integral\<^isup>P (restricted_space A) f = (\\<^isup>+ x. f x * indicator A x \M)" + (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f") proof - - have msR: "measure_space ?R \" by (rule restricted_measure_space[OF `A \ sets M`]) - then interpret R: measure_space ?R \ . + have msR: "measure_space ?R" by (rule restricted_measure_space[OF `A \ sets M`]) + then interpret R: measure_space ?R . have saR: "sigma_algebra ?R" by fact - have *: "R.positive_integral f = R.positive_integral ?f" + have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f" by (intro R.positive_integral_cong) auto show ?thesis - unfolding * R.positive_integral_def positive_integral_def + unfolding * positive_integral_def unfolding simple_function_restricted[OF `A \ sets M`] apply (simp add: SUPR_def) apply (rule arg_cong[where f=Sup]) proof (auto simp add: image_iff simple_integral_restricted[OF `A \ sets M`]) - fix g assume "simple_function (\x. g x * indicator A x)" + fix g assume "simple_function M (\x. g x * indicator A x)" "g \ f" - then show "\x. simple_function x \ x \ (\x. f x * indicator A x) \ - (\\<^isup>Sx. g x * indicator A x) = simple_integral x" + then show "\x. simple_function M x \ x \ (\x. f x * indicator A x) \ + (\\<^isup>Sx. g x * indicator A x \M) = integral\<^isup>S M x" apply (rule_tac exI[of _ "\x. g x * indicator A x"]) by (auto simp: indicator_def le_fun_def) next - fix g assume g: "simple_function g" "g \ (\x. f x * indicator A x)" + fix g assume g: "simple_function M g" "g \ (\x. f x * indicator A x)" then have *: "(\x. g x * indicator A x) = g" "\x. g x * indicator A x = g x" "\x. g x \ f x" by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm) - from g show "\x. simple_function (\xa. x xa * indicator A xa) \ x \ f \ - simple_integral g = simple_integral (\xa. x xa * indicator A xa)" + from g show "\x. simple_function M (\xa. x xa * indicator A xa) \ x \ f \ + integral\<^isup>S M g = integral\<^isup>S M (\xa. x xa * indicator A xa)" using `A \ sets M`[THEN sets_into_space] apply (rule_tac exI[of _ "\x. g x * indicator A x"]) by (fastsimp simp: le_fun_def *) @@ -1505,103 +1525,113 @@ lemma (in measure_space) positive_integral_subalgebra: assumes borel: "f \ borel_measurable N" - and N: "sets N \ sets M" "space N = space M" and sa: "sigma_algebra N" - shows "measure_space.positive_integral N \ f = positive_integral f" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" + and sa: "sigma_algebra N" + shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - - interpret N: measure_space N \ using measure_space_subalgebra[OF sa N] . + interpret N: measure_space N using measure_space_subalgebra[OF sa N] . from N.borel_measurable_implies_simple_function_sequence[OF borel] - obtain fs where Nsf: "\i. N.simple_function (fs i)" and "fs \ f" by blast - then have sf: "\i. simple_function (fs i)" - using simple_function_subalgebra[OF _ N sa] by blast - from positive_integral_isoton_simple[OF `fs \ f` sf] N.positive_integral_isoton_simple[OF `fs \ f` Nsf] - show ?thesis unfolding isoton_def simple_integral_def N.simple_integral_def `space N = space M` by simp + obtain fs where Nsf: "\i. simple_function N (fs i)" and "fs \ f" by blast + then have sf: "\i. simple_function M (fs i)" + using simple_function_subalgebra[OF _ N(1,2)] by blast + from N.positive_integral_isoton_simple[OF `fs \ f` Nsf] + have "integral\<^isup>P N f = (SUP i. \x\fs i ` space M. x * N.\ (fs i -` {x} \ space M))" + unfolding isoton_def simple_integral_def `space N = space M` by simp + also have "\ = (SUP i. \x\fs i ` space M. x * \ (fs i -` {x} \ space M))" + using N N.simple_functionD(2)[OF Nsf] unfolding `space N = space M` by auto + also have "\ = integral\<^isup>P M f" + using positive_integral_isoton_simple[OF `fs \ f` sf] + unfolding isoton_def simple_integral_def `space N = space M` by simp + finally show ?thesis . qed section "Lebesgue Integral" -definition (in measure_space) integrable where - "integrable f \ f \ borel_measurable M \ - (\\<^isup>+ x. Real (f x)) \ \ \ - (\\<^isup>+ x. Real (- f x)) \ \" +definition integrable where + "integrable M f \ f \ borel_measurable M \ + (\\<^isup>+ x. Real (f x) \M) \ \ \ + (\\<^isup>+ x. Real (- f x) \M) \ \" -lemma (in measure_space) integrableD[dest]: - assumes "integrable f" - shows "f \ borel_measurable M" - "(\\<^isup>+ x. Real (f x)) \ \" - "(\\<^isup>+ x. Real (- f x)) \ \" +lemma integrableD[dest]: + assumes "integrable M f" + shows "f \ borel_measurable M" "(\\<^isup>+ x. Real (f x) \M) \ \" "(\\<^isup>+ x. Real (- f x) \M) \ \" using assms unfolding integrable_def by auto -definition (in measure_space) integral (binder "\ " 10) where - "integral f = real ((\\<^isup>+ x. Real (f x))) - real ((\\<^isup>+ x. Real (- f x)))" +definition lebesgue_integral_def: + "integral\<^isup>L M f = real ((\\<^isup>+ x. Real (f x) \M)) - real ((\\<^isup>+ x. Real (- f x) \M))" + +syntax + "_lebesgue_integral" :: "'a \ ('a \ real) \ ('a, 'b) measure_space_scheme \ real" ("\ _. _ \_" [60,61] 110) + +translations + "\ x. f \M" == "CONST integral\<^isup>L M (%x. f)" lemma (in measure_space) integral_cong: - assumes cong: "\x. x \ space M \ f x = g x" - shows "integral f = integral g" - using assms by (simp cong: positive_integral_cong add: integral_def) + assumes "\x. x \ space M \ f x = g x" + shows "integral\<^isup>L M f = integral\<^isup>L M g" + using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def) lemma (in measure_space) integral_cong_measure: - assumes "\A. A \ sets M \ \ A = \ A" - shows "measure_space.integral M \ f = integral f" + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + shows "integral\<^isup>L N f = integral\<^isup>L M f" proof - - interpret v: measure_space M \ - by (rule measure_space_cong) fact + interpret v: measure_space N + by (rule measure_space_cong) fact+ show ?thesis - unfolding integral_def v.integral_def - by (simp add: positive_integral_cong_measure[OF assms]) + by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def) qed lemma (in measure_space) integral_cong_AE: assumes cong: "AE x. f x = g x" - shows "integral f = integral g" + shows "integral\<^isup>L M f = integral\<^isup>L M g" proof - have "AE x. Real (f x) = Real (g x)" using cong by (rule AE_mp) simp moreover have "AE x. Real (- f x) = Real (- g x)" using cong by (rule AE_mp) simp ultimately show ?thesis - by (simp cong: positive_integral_cong_AE add: integral_def) + by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def) qed lemma (in measure_space) integrable_cong: - "(\x. x \ space M \ f x = g x) \ integrable f \ integrable g" + "(\x. x \ space M \ f x = g x) \ integrable M f \ integrable M g" by (simp cong: positive_integral_cong measurable_cong add: integrable_def) lemma (in measure_space) integral_eq_positive_integral: assumes "\x. 0 \ f x" - shows "integral f = real ((\\<^isup>+ x. Real (f x)))" + shows "integral\<^isup>L M f = real (\\<^isup>+ x. Real (f x) \M)" proof - have "\x. Real (- f x) = 0" using assms by simp - thus ?thesis by (simp del: Real_eq_0 add: integral_def) + thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def) qed lemma (in measure_space) integral_vimage: assumes T: "sigma_algebra M'" "T \ measurable M M'" - assumes f: "measure_space.integrable M' (\A. \ (T -` A \ space M)) f" - shows "integrable (\x. f (T x))" (is ?P) - and "measure_space.integral M' (\A. \ (T -` A \ space M)) f = (\x. f (T x))" (is ?I) + and \: "\A. A \ sets M' \ measure M' A = \ (T -` A \ space M)" + assumes f: "integrable M' f" + shows "integrable M (\x. f (T x))" (is ?P) + and "integral\<^isup>L M' f = (\x. f (T x) \M)" (is ?I) proof - - interpret T: measure_space M' "\A. \ (T -` A \ space M)" - using T by (rule measure_space_vimage) auto + interpret T: measure_space M' using \ by (rule measure_space_vimage[OF T]) from measurable_comp[OF T(2), of f borel] have borel: "(\x. Real (f x)) \ borel_measurable M'" "(\x. Real (- f x)) \ borel_measurable M'" and "(\x. f (T x)) \ borel_measurable M" - using f unfolding T.integrable_def by (auto simp: comp_def) + using f unfolding integrable_def by (auto simp: comp_def) then show ?P ?I - using f unfolding T.integral_def integral_def T.integrable_def integrable_def - unfolding borel[THEN positive_integral_vimage[OF T]] by auto + using f unfolding lebesgue_integral_def integrable_def + by (auto simp: borel[THEN positive_integral_vimage[OF T], OF \]) qed lemma (in measure_space) integral_minus[intro, simp]: - assumes "integrable f" - shows "integrable (\x. - f x)" "(\x. - f x) = - integral f" - using assms by (auto simp: integrable_def integral_def) + assumes "integrable M f" + shows "integrable M (\x. - f x)" "(\x. - f x \M) = - integral\<^isup>L M f" + using assms by (auto simp: integrable_def lebesgue_integral_def) lemma (in measure_space) integral_of_positive_diff: - assumes integrable: "integrable u" "integrable v" + assumes integrable: "integrable M u" "integrable M v" and f_def: "\x. f x = u x - v x" and pos: "\x. 0 \ u x" "\x. 0 \ v x" - shows "integrable f" and "integral f = integral u - integral v" + shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" proof - - let ?PI = positive_integral let "?f x" = "Real (f x)" let "?mf x" = "Real (- f x)" let "?u x" = "Real (u x)" @@ -1615,38 +1645,39 @@ "f \ borel_measurable M" by (auto simp: f_def[symmetric] integrable_def) - have "?PI (\x. Real (u x - v x)) \ ?PI ?u" + have "(\\<^isup>+ x. Real (u x - v x) \M) \ integral\<^isup>P M ?u" using pos by (auto intro!: positive_integral_mono) - moreover have "?PI (\x. Real (v x - u x)) \ ?PI ?v" + moreover have "(\\<^isup>+ x. Real (v x - u x) \M) \ integral\<^isup>P M ?v" using pos by (auto intro!: positive_integral_mono) - ultimately show f: "integrable f" - using `integrable u` `integrable v` `f \ borel_measurable M` + ultimately show f: "integrable M f" + using `integrable M u` `integrable M v` `f \ borel_measurable M` by (auto simp: integrable_def f_def) - hence mf: "integrable (\x. - f x)" .. + hence mf: "integrable M (\x. - f x)" .. have *: "\x. Real (- v x) = 0" "\x. Real (- u x) = 0" using pos by auto have "\x. ?u x + ?mf x = ?v x + ?f x" unfolding f_def using pos by simp - hence "?PI (\x. ?u x + ?mf x) = ?PI (\x. ?v x + ?f x)" by simp - hence "real (?PI ?u + ?PI ?mf) = real (?PI ?v + ?PI ?f)" + hence "(\\<^isup>+ x. ?u x + ?mf x \M) = (\\<^isup>+ x. ?v x + ?f x \M)" by simp + hence "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) = + real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)" using positive_integral_add[OF u_borel mf_borel] using positive_integral_add[OF v_borel f_borel] by auto - then show "integral f = integral u - integral v" - using f mf `integrable u` `integrable v` - unfolding integral_def integrable_def * - by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?v", cases "?PI ?u") + then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" + using f mf `integrable M u` `integrable M v` + unfolding lebesgue_integral_def integrable_def * + by (cases "integral\<^isup>P M ?f", cases "integral\<^isup>P M ?mf", cases "integral\<^isup>P M ?v", cases "integral\<^isup>P M ?u") (auto simp add: field_simps) qed lemma (in measure_space) integral_linear: - assumes "integrable f" "integrable g" and "0 \ a" - shows "integrable (\t. a * f t + g t)" - and "(\ t. a * f t + g t) = a * integral f + integral g" + assumes "integrable M f" "integrable M g" and "0 \ a" + shows "integrable M (\t. a * f t + g t)" + and "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" proof - - let ?PI = positive_integral + let ?PI = "integral\<^isup>P M" let "?f x" = "Real (f x)" let "?g x" = "Real (g x)" let "?mf x" = "Real (- f x)" @@ -1670,37 +1701,36 @@ positive_integral_linear[OF pos] positive_integral_linear[OF neg] - have "integrable ?p" "integrable ?n" + have "integrable M ?p" "integrable M ?n" "\t. a * f t + g t = ?p t - ?n t" "\t. 0 \ ?p t" "\t. 0 \ ?n t" using assms p n unfolding integrable_def * linear by auto note diff = integral_of_positive_diff[OF this] - show "integrable (\t. a * f t + g t)" by (rule diff) + show "integrable M (\t. a * f t + g t)" by (rule diff) - from assms show "(\ t. a * f t + g t) = a * integral f + integral g" - unfolding diff(2) unfolding integral_def * linear integrable_def + from assms show "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" + unfolding diff(2) unfolding lebesgue_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) qed lemma (in measure_space) integral_add[simp, intro]: - assumes "integrable f" "integrable g" - shows "integrable (\t. f t + g t)" - and "(\ t. f t + g t) = integral f + integral g" + assumes "integrable M f" "integrable M g" + shows "integrable M (\t. f t + g t)" + and "(\ t. f t + g t \M) = integral\<^isup>L M f + integral\<^isup>L M g" using assms integral_linear[where a=1] by auto lemma (in measure_space) integral_zero[simp, intro]: - shows "integrable (\x. 0)" - and "(\ x.0) = 0" - unfolding integrable_def integral_def + shows "integrable M (\x. 0)" "(\ x.0 \M) = 0" + unfolding integrable_def lebesgue_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 "(\ t. a * f t) = a * integral f" (is ?I) + assumes "integrable M f" + shows "integrable M (\t. a * f t)" (is ?P) + and "(\ t. a * f t \M) = a * integral\<^isup>L M f" (is ?I) proof - - have "integrable (\t. a * f t) \ (\ t. a * f t) = a * integral f" + have "integrable M (\t. a * f t) \ (\ t. a * f t \M) = a * integral\<^isup>L M f" proof (cases rule: le_cases) assume "0 \ a" show ?thesis using integral_linear[OF assms integral_zero(1) `0 \ a`] @@ -1716,56 +1746,56 @@ qed lemma (in measure_space) integral_multc: - assumes "integrable f" - shows "(\ x. f x * c) = integral f * c" + assumes "integrable M f" + shows "(\ x. f x * c \M) = integral\<^isup>L M f * c" unfolding mult_commute[of _ c] integral_cmult[OF assms] .. lemma (in measure_space) integral_mono_AE: - assumes fg: "integrable f" "integrable g" + assumes fg: "integrable M f" "integrable M g" and mono: "AE t. f t \ g t" - shows "integral f \ integral g" + shows "integral\<^isup>L M f \ integral\<^isup>L M g" proof - have "AE x. Real (f x) \ Real (g x)" using mono by (rule AE_mp) (auto intro!: AE_cong) - moreover have "AE x. Real (- g x) \ Real (- f x)" + moreover have "AE x. Real (- g x) \ Real (- f x)" using mono by (rule AE_mp) (auto intro!: AE_cong) ultimately show ?thesis using fg - by (auto simp: integral_def integrable_def diff_minus + by (auto simp: lebesgue_integral_def integrable_def diff_minus intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE) qed lemma (in measure_space) integral_mono: - assumes fg: "integrable f" "integrable g" + assumes fg: "integrable M f" "integrable M g" and mono: "\t. t \ space M \ f t \ g t" - shows "integral f \ integral g" + shows "integral\<^isup>L M f \ integral\<^isup>L M g" apply (rule integral_mono_AE[OF fg]) using mono by (rule AE_cong) auto lemma (in measure_space) integral_diff[simp, intro]: - assumes f: "integrable f" and g: "integrable g" - shows "integrable (\t. f t - g t)" - and "(\ t. f t - g t) = integral f - integral g" + assumes f: "integrable M f" and g: "integrable M g" + shows "integrable M (\t. f t - g t)" + and "(\ t. f t - g t \M) = integral\<^isup>L M f - integral\<^isup>L M g" using integral_add[OF f integral_minus(1)[OF g]] unfolding diff_minus integral_minus(2)[OF g] by auto lemma (in measure_space) integral_indicator[simp, intro]: assumes "a \ sets M" and "\ a \ \" - shows "integral (indicator a) = real (\ a)" (is ?int) - and "integrable (indicator a)" (is ?able) + shows "integral\<^isup>L M (indicator a) = real (\ a)" (is ?int) + and "integrable M (indicator a)" (is ?able) proof - have *: "\A x. Real (indicator A x) = indicator A x" "\A x. Real (- indicator A x) = 0" unfolding indicator_def by auto show ?int ?able - using assms unfolding integral_def integrable_def + using assms unfolding lebesgue_integral_def integrable_def by (auto simp: * positive_integral_indicator borel_measurable_indicator) qed 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 "(\x. c * indicator A x) = c * real (\ A)" (is ?I) + shows "integrable M (\x. c * indicator A x)" (is ?P) + and "(\x. c * indicator A x \M) = c * real (\ A)" (is ?I) proof - show ?P proof (cases "c = 0") @@ -1779,9 +1809,9 @@ qed lemma (in measure_space) integral_setsum[simp, intro]: - assumes "\n. n \ S \ integrable (f n)" - 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") + assumes "\n. n \ S \ integrable M (f n)" + shows "(\x. (\ i \ S. f i x) \M) = (\ i \ S. integral\<^isup>L M (f i))" (is "?int S") + and "integrable M (\x. \ i \ S. f i x)" (is "?I S") proof - have "?int S \ ?I S" proof (cases "finite S") @@ -1792,8 +1822,8 @@ qed lemma (in measure_space) integrable_abs: - assumes "integrable f" - shows "integrable (\ x. \f x\)" + assumes "integrable M f" + shows "integrable M (\ x. \f x\)" proof - have *: "\x. Real \f x\ = Real (f x) + Real (- f x)" @@ -1808,56 +1838,55 @@ lemma (in measure_space) integral_subalgebra: assumes borel: "f \ borel_measurable N" - and N: "sets N \ sets M" "space N = space M" and sa: "sigma_algebra N" - shows "measure_space.integrable N \ f \ integrable f" (is ?P) - and "measure_space.integral N \ f = integral f" (is ?I) + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" and sa: "sigma_algebra N" + shows "integrable N f \ integrable M f" (is ?P) + and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I) proof - - interpret N: measure_space N \ using measure_space_subalgebra[OF sa N] . + interpret N: measure_space N using measure_space_subalgebra[OF sa N] . have "(\x. Real (f x)) \ borel_measurable N" "(\x. Real (- f x)) \ borel_measurable N" using borel by auto note * = this[THEN positive_integral_subalgebra[OF _ N sa]] have "f \ borel_measurable M \ f \ borel_measurable N" using assms unfolding measurable_def by auto - then show ?P ?I unfolding integrable_def N.integrable_def integral_def N.integral_def - unfolding * by auto + then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def) qed lemma (in measure_space) integrable_bound: - assumes "integrable f" + assumes "integrable M f" and f: "\x. x \ space M \ 0 \ f x" "\x. x \ space M \ \g x\ \ f x" assumes borel: "g \ borel_measurable M" - shows "integrable g" + shows "integrable M g" proof - - have "(\\<^isup>+ x. Real (g x)) \ (\\<^isup>+ x. Real \g x\)" + have "(\\<^isup>+ x. Real (g x) \M) \ (\\<^isup>+ x. Real \g x\ \M)" by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. Real (f x))" + also have "\ \ (\\<^isup>+ x. Real (f x) \M)" using f by (auto intro!: positive_integral_mono) also have "\ < \" - using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have pos: "(\\<^isup>+ x. Real (g x)) < \" . + using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\) + finally have pos: "(\\<^isup>+ x. Real (g x) \M) < \" . - have "(\\<^isup>+ x. Real (- g x)) \ (\\<^isup>+ x. Real (\g x\))" + have "(\\<^isup>+ x. Real (- g x) \M) \ (\\<^isup>+ x. Real (\g x\) \M)" by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. Real (f x))" + also have "\ \ (\\<^isup>+ x. Real (f x) \M)" using f by (auto intro!: positive_integral_mono) also have "\ < \" - using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have neg: "(\\<^isup>+ x. Real (- g x)) < \" . + using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\) + finally have neg: "(\\<^isup>+ x. Real (- g x) \M) < \" . from neg pos borel show ?thesis unfolding integrable_def by auto qed lemma (in measure_space) integrable_abs_iff: - "f \ borel_measurable M \ integrable (\ x. \f x\) \ integrable f" + "f \ borel_measurable M \ integrable M (\ x. \f x\) \ integrable M f" by (auto intro!: integrable_bound[where g=f] integrable_abs) lemma (in measure_space) integrable_max: - assumes int: "integrable f" "integrable g" - shows "integrable (\ x. max (f x) (g x))" + assumes int: "integrable M f" "integrable M g" + shows "integrable M (\ x. max (f x) (g x))" proof (rule integrable_bound) - show "integrable (\x. \f x\ + \g x\)" + show "integrable M (\x. \f x\ + \g x\)" using int by (simp add: integrable_abs) show "(\x. max (f x) (g x)) \ borel_measurable M" using int unfolding integrable_def by auto @@ -1868,10 +1897,10 @@ qed lemma (in measure_space) integrable_min: - assumes int: "integrable f" "integrable g" - shows "integrable (\ x. min (f x) (g x))" + assumes int: "integrable M f" "integrable M g" + shows "integrable M (\ x. min (f x) (g x))" proof (rule integrable_bound) - show "integrable (\x. \f x\ + \g x\)" + show "integrable M (\x. \f x\ + \g x\)" using int by (simp add: integrable_abs) show "(\x. min (f x) (g x)) \ borel_measurable M" using int unfolding integrable_def by auto @@ -1882,33 +1911,33 @@ qed lemma (in measure_space) integral_triangle_inequality: - assumes "integrable f" - shows "\integral f\ \ (\x. \f x\)" + assumes "integrable M f" + shows "\integral\<^isup>L M f\ \ (\x. \f x\ \M)" proof - - have "\integral f\ = max (integral f) (- integral f)" by auto - also have "\ \ (\x. \f x\)" + have "\integral\<^isup>L M f\ = max (integral\<^isup>L M f) (- integral\<^isup>L M f)" by auto + also have "\ \ (\x. \f x\ \M)" using assms integral_minus(2)[of f, symmetric] by (auto intro!: integral_mono integrable_abs simp del: integral_minus) finally show ?thesis . qed lemma (in measure_space) integral_positive: - assumes "integrable f" "\x. x \ space M \ 0 \ f x" - shows "0 \ integral f" + assumes "integrable M f" "\x. x \ space M \ 0 \ f x" + shows "0 \ integral\<^isup>L M f" proof - - have "0 = (\x. 0)" by (auto simp: integral_zero) - also have "\ \ integral f" + have "0 = (\x. 0 \M)" by (auto simp: integral_zero) + also have "\ \ integral\<^isup>L M f" using assms by (rule integral_mono[OF integral_zero(1)]) finally show ?thesis . qed lemma (in measure_space) integral_monotone_convergence_pos: - assumes i: "\i. integrable (f i)" and mono: "\x. mono (\n. f n x)" + assumes i: "\i. integrable M (f i)" and mono: "\x. mono (\n. f n x)" and pos: "\x i. 0 \ f i x" and lim: "\x. (\i. f i x) ----> u x" - and ilim: "(\i. integral (f i)) ----> x" - shows "integrable u" - and "integral u = x" + and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + shows "integrable M u" + and "integral\<^isup>L M u = x" proof - { fix x have "0 \ u x" using mono pos[of 0 x] incseq_le[OF _ lim, of x 0] @@ -1928,43 +1957,42 @@ hence borel_u: "u \ borel_measurable M" using pos_u by (auto simp: borel_measurable_Real_eq SUP_F) - 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 integral_eq: "\n. (\\<^isup>+ x. Real (f n x) \M) = Real (integral\<^isup>L M (f n))" + using i unfolding lebesgue_integral_def integrable_def by (auto simp: Real_real) - have pos_integral: "\n. 0 \ integral (f n)" + have pos_integral: "\n. 0 \ integral\<^isup>L M (f n)" using pos i by (auto simp: integral_positive) hence "0 \ x" using LIMSEQ_le_const[OF ilim, of 0] by auto - have "(\i. (\\<^isup>+ x. Real (f i x))) \ (\\<^isup>+ x. Real (u x))" + have "(\i. (\\<^isup>+ x. Real (f i x) \M)) \ (\\<^isup>+ x. Real (u x) \M)" 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: "(\\<^isup>+ x. Real (u x)) = - (SUP n. (\\<^isup>+ x. Real (f n x)))" + hence pI: "(\\<^isup>+ x. Real (u x) \M) = (SUP n. (\\<^isup>+ x. Real (f n x) \M))" unfolding isoton_def by simp also have "\ = Real x" unfolding integral_eq proof (rule SUP_eq_LIMSEQ[THEN iffD2]) - show "mono (\n. integral (f n))" + show "mono (\n. integral\<^isup>L M (f n))" using mono i by (auto simp: mono_def intro!: integral_mono) - show "\n. 0 \ integral (f n)" using pos_integral . + show "\n. 0 \ integral\<^isup>L M (f n)" using pos_integral . show "0 \ x" using `0 \ x` . - show "(\n. integral (f n)) ----> x" using ilim . + show "(\n. integral\<^isup>L M (f n)) ----> x" using ilim . qed - finally show "integrable u" "integral u = x" using borel_u `0 \ x` - unfolding integrable_def integral_def by auto + finally show "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \ x` + unfolding integrable_def lebesgue_integral_def by auto qed lemma (in measure_space) integral_monotone_convergence: - assumes f: "\i. integrable (f i)" and "mono f" + assumes f: "\i. integrable M (f i)" and "mono f" and lim: "\x. (\i. f i x) ----> u x" - and ilim: "(\i. integral (f i)) ----> x" - shows "integrable u" - and "integral u = x" + and ilim: "(\i. integral\<^isup>L M (f i)) ----> x" + shows "integrable M u" + and "integral\<^isup>L M u = x" proof - - have 1: "\i. integrable (\x. f i x - f 0 x)" + have 1: "\i. integrable M (\x. f i x - f 0 x)" using f by (auto intro!: integral_diff) have 2: "\x. mono (\n. f n x - f 0 x)" using `mono f` unfolding mono_def le_fun_def by auto @@ -1972,43 +2000,43 @@ 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. (\x. f i x - f 0 x)) ----> x - integral (f 0)" + have 5: "(\i. (\x. f i x - f 0 x \M)) ----> x - integral\<^isup>L M (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)" + have "integrable M (\x. (u x - f 0 x) + f 0 x)" using diff(1) f by (rule integral_add(1)) - with diff(2) f show "integrable u" "integral u = x" + with diff(2) f show "integrable M u" "integral\<^isup>L M u = x" by (auto simp: integral_diff) qed lemma (in measure_space) integral_0_iff: - assumes "integrable f" - shows "(\x. \f x\) = 0 \ \ {x\space M. f x \ 0} = 0" + assumes "integrable M f" + shows "(\x. \f x\ \M) = 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) + have "integrable M (\x. \f x\)" using assms by (rule integrable_abs) hence "(\x. Real (\f x\)) \ borel_measurable M" - "(\\<^isup>+ x. Real \f x\) \ \" unfolding integrable_def by auto + "(\\<^isup>+ x. Real \f x\ \M) \ \" unfolding integrable_def by auto from positive_integral_0_iff[OF this(1)] this(2) - show ?thesis unfolding integral_def * + show ?thesis unfolding lebesgue_integral_def * by (simp add: real_of_pextreal_eq_0) qed lemma (in measure_space) positive_integral_omega: assumes "f \ borel_measurable M" - and "positive_integral f \ \" + and "integral\<^isup>P M f \ \" shows "\ (f -` {\} \ space M) = 0" proof - - have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x)" + have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x \M)" using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \ \] by simp - also have "\ \ positive_integral f" + also have "\ \ integral\<^isup>P M f" by (auto intro!: positive_integral_mono simp: indicator_def) finally show ?thesis using assms(2) by (cases ?thesis) auto qed lemma (in measure_space) positive_integral_omega_AE: - assumes "f \ borel_measurable M" "positive_integral f \ \" shows "AE x. f x \ \" + assumes "f \ borel_measurable M" "integral\<^isup>P M f \ \" shows "AE x. f x \ \" proof (rule AE_I) show "\ (f -` {\} \ space M) = 0" by (rule positive_integral_omega[OF assms]) @@ -2017,39 +2045,39 @@ qed auto lemma (in measure_space) simple_integral_omega: - assumes "simple_function f" - and "simple_integral f \ \" + assumes "simple_function M f" + and "integral\<^isup>S M f \ \" shows "\ (f -` {\} \ space M) = 0" proof (rule positive_integral_omega) show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) - show "positive_integral f \ \" + show "integral\<^isup>P M f \ \" using assms by (simp add: positive_integral_eq_simple_integral) qed lemma (in measure_space) integral_real: fixes f :: "'a \ pextreal" assumes "AE x. f x \ \" - shows "(\x. real (f x)) = real (positive_integral f)" (is ?plus) - and "(\x. - real (f x)) = - real (positive_integral f)" (is ?minus) + shows "(\x. real (f x) \M) = real (integral\<^isup>P M f)" (is ?plus) + and "(\x. - real (f x) \M) = - real (integral\<^isup>P M f)" (is ?minus) proof - - have "(\\<^isup>+ x. Real (real (f x))) = positive_integral f" + have "(\\<^isup>+ x. Real (real (f x)) \M) = integral\<^isup>P M f" apply (rule positive_integral_cong_AE) apply (rule AE_mp[OF assms(1)]) by (auto intro!: AE_cong simp: Real_real) moreover - have "(\\<^isup>+ x. Real (- real (f x))) = (\\<^isup>+ x. 0)" + have "(\\<^isup>+ x. Real (- real (f x)) \M) = (\\<^isup>+ x. 0 \M)" by (intro positive_integral_cong) auto ultimately show ?plus ?minus - by (auto simp: integral_def integrable_def) + by (auto simp: lebesgue_integral_def integrable_def) qed lemma (in measure_space) integral_dominated_convergence: - assumes u: "\i. integrable (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" - and w: "integrable w" "\x. x \ space M \ 0 \ w x" + assumes u: "\i. integrable M (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" + and w: "integrable M 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. (\x. \u i x - u' x\)) ----> 0" (is "?lim_diff") - and "(\i. integral (u i)) ----> integral u'" (is ?lim) + shows "integrable M u'" + and "(\i. (\x. \u i x - u' x\ \M)) ----> 0" (is "?lim_diff") + and "(\i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim) proof - { fix x j assume x: "x \ space M" from u'[OF x] have "(\i. \u i x\) ----> \u' x\" by (rule LIMSEQ_imp_rabs) @@ -2061,9 +2089,9 @@ have u'_borel: "u' \ borel_measurable M" using u' by (blast intro: borel_measurable_LIMSEQ[of u]) - show "integrable u'" + show "integrable M u'" proof (rule integrable_bound) - show "integrable w" by fact + show "integrable M w" by fact show "u' \ borel_measurable M" by fact next fix x assume x: "x \ space M" @@ -2072,8 +2100,8 @@ qed let "?diff n x" = "2 * w x - \u n x - u' x\" - have diff: "\n. integrable (\x. \u n x - u' x\)" - using w u `integrable u'` + have diff: "\n. integrable M (\x. \u n x - u' x\)" + using w u `integrable M u'` by (auto intro!: integral_add integral_diff integral_cmult integrable_abs) { fix j x assume x: "x \ space M" @@ -2083,31 +2111,31 @@ finally have "\u j x - u' x\ \ 2 * w x" by simp } note diff_less_2w = this - 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\)" + have PI_diff: "\m n. (\\<^isup>+ x. Real (?diff (m + n) x) \M) = + (\\<^isup>+ x. Real (2 * w x) \M) - (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)" 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)" + have "integrable M (\x. 2 * w x)" using w by (auto intro: integral_cmult) - hence I2w_fin: "(\\<^isup>+ x. Real (2 * w x)) \ \" and + hence I2w_fin: "(\\<^isup>+ x. Real (2 * w x) \M) \ \" and borel_2w: "(\x. Real (2 * w x)) \ borel_measurable M" unfolding integrable_def by auto - have "(INF n. SUP m. (\\<^isup>+ 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\ \M)) = 0" (is "?lim_SUP = 0") proof cases - 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))" + assume eq_0: "(\\<^isup>+ x. Real (2 * w x) \M) = 0" + have "\i. (\\<^isup>+ x. Real \u i x - u' x\ \M) \ (\\<^isup>+ x. Real (2 * w x) \M)" 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. (\\<^isup>+ x. Real \u i x - u' x\) = 0" using eq_0 by auto + hence "\i. (\\<^isup>+ x. Real \u i x - u' x\ \M) = 0" using eq_0 by auto thus ?thesis by simp next - 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))" + assume neq_0: "(\\<^isup>+ x. Real (2 * w x) \M) \ 0" + have "(\\<^isup>+ x. Real (2 * w x) \M) = (\\<^isup>+ x. (SUP n. INF m. Real (?diff (m + n) x)) \M)" 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))" @@ -2119,22 +2147,22 @@ thus "(\i. ?diff i x) ----> 2 * w x" by simp qed qed - also have "\ \ (SUP n. INF m. (\\<^isup>+ x. Real (?diff (m + n) x)))" + also have "\ \ (SUP n. INF m. (\\<^isup>+ x. Real (?diff (m + n) x) \M))" using u'_borel w u unfolding integrable_def by (auto intro!: positive_integral_lim_INF) - also have "\ = (\\<^isup>+ x. Real (2 * w x)) - - (INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\))" + also have "\ = (\\<^isup>+ x. Real (2 * w x) \M) - + (INF n. SUP m. \\<^isup>+ x. Real \u (m + n) x - u' x\ \M)" 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. (\\<^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 [simp]: "\n m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M) = + Real ((\x. \u (n + m) x - u' x\ \M))" + using diff by (subst add_commute) (simp add: lebesgue_integral_def integrable_def Real_real) - have "(SUP n. INF m. (\\<^isup>+ x. Real \u (m + n) x - u' x\)) \ ?lim_SUP" + have "(SUP n. INF m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)) \ ?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) @@ -2143,28 +2171,28 @@ proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from LIMSEQ_D[OF `?lim_diff` this] - obtain N where N: "\n. N \ n \ (\x. \u n x - u' x\) < r" + obtain N where N: "\n. N \ n \ (\x. \u n x - u' x\ \M) < r" using diff by (auto simp: integral_positive) - show "\N. \n\N. norm (integral (u n) - integral u') < r" + show "\N. \n\N. norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" proof (safe intro!: exI[of _ N]) fix n assume "N \ n" - have "\integral (u n) - integral u'\ = \(\x. u n x - u' x)\" - using u `integrable u'` by (auto simp: integral_diff) - also have "\ \ (\x. \u n x - u' x\)" using u `integrable u'` + have "\integral\<^isup>L M (u n) - integral\<^isup>L M u'\ = \(\x. u n x - u' x \M)\" + using u `integrable M u'` by (auto simp: integral_diff) + also have "\ \ (\x. \u n x - u' x\ \M)" using u `integrable M 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 + finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp qed qed qed lemma (in measure_space) integral_sums: - assumes borel: "\i. integrable (f i)" + assumes borel: "\i. integrable M (f i)" and summable: "\x. x \ space M \ summable (\i. \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 (\x. (\i. f i x))" (is ?integral) + and sums: "summable (\i. (\x. \f i x\ \M))" + shows "integrable M (\x. (\i. f i x))" (is "integrable M ?S") + and "(\i. integral\<^isup>L M (f i)) sums (\x. (\i. f i x) \M)" (is ?integral) proof - have "\x\space M. \w. (\i. \f i x\) sums w" using summable unfolding summable_def by auto @@ -2173,10 +2201,10 @@ let "?w y" = "if y \ space M then w y else 0" - obtain x where abs_sum: "(\i. (\x. \f i x\)) sums x" + obtain x where abs_sum: "(\i. (\x. \f i x\ \M)) sums x" using sums unfolding summable_def .. - have 1: "\n. integrable (\x. \i = 0..n. integrable M (\x. \i = 0.. space M" @@ -2185,21 +2213,21 @@ finally have "\\i = 0.. \ ?w x" by simp } note 2 = this - have 3: "integrable ?w" + have 3: "integrable M ?w" proof (rule integral_monotone_convergence(1)) let "?F n y" = "(\i = 0..f i y\)" let "?w' n y" = "if y \ space M then ?F n y else 0" - have "\n. integrable (?F n)" + have "\n. integrable M (?F n)" using borel by (auto intro!: integral_setsum integrable_abs) - thus "\n. integrable (?w' n)" by (simp cong: integrable_cong) + thus "\n. integrable M (?w' n)" by (simp cong: integrable_cong) show "mono ?w'" by (auto simp: mono_def le_fun_def intro!: setsum_mono2) { fix x show "(\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. (\x. \f i x\))" + have *: "\n. integral\<^isup>L M (?w' n) = (\i = 0..< n. (\x. \f i x\ \M))" using borel by (simp add: integral_setsum integrable_abs cong: integral_cong) from abs_sum - show "(\i. integral (?w' i)) ----> x" unfolding * sums_def . + show "(\i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def . qed have 4: "\x. x \ space M \ 0 \ ?w x" using 2[of _ 0] by simp @@ -2210,7 +2238,7 @@ note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 5] - from int show "integrable ?S" by simp + from int show "integrable M ?S" by simp show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel] using int(2) by simp @@ -2224,12 +2252,12 @@ and enum_zero: "enum ` (-S) \ {0}" and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" and abs_summable: "summable (\r. \enum r * real (\ (f -` {enum r} \ space M))\)" - shows "integrable f" - and "(\r. enum r * real (\ (f -` {enum r} \ space M))) sums integral f" (is ?sums) + shows "integrable M f" + and "(\r. enum r * real (\ (f -` {enum r} \ space M))) sums integral\<^isup>L M f" (is ?sums) proof - let "?A r" = "f -` {enum r} \ space M" let "?F r x" = "enum r * indicator (?A r) x" - have enum_eq: "\r. enum r * real (\ (?A r)) = integral (?F r)" + have enum_eq: "\r. enum r * real (\ (?A r)) = integral\<^isup>L M (?F r)" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) { fix x assume "x \ space M" @@ -2250,19 +2278,19 @@ 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 = (\x. \r. ?F r x)" "integrable f = integrable (\x. \r. ?F r x)" + have int_f: "integral\<^isup>L M f = (\x. (\r. ?F r x) \M)" "integrable M f = integrable M (\x. \r. ?F r x)" using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff) { fix r - have "(\x. \?F r x\) = (\x. \enum r\ * indicator (?A r) x)" + have "(\x. \?F r x\ \M) = (\x. \enum r\ * indicator (?A r) x \M)" 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 "(\x. \?F r x\) = \enum r * real (\ (?A r))\" + finally have "(\x. \?F r x\ \M) = \enum r * real (\ (?A r))\" by (simp add: abs_mult_pos real_pextreal_pos) } note int_abs_F = this - have 1: "\i. integrable (\x. ?F i x)" + have 1: "\i. integrable M (\x. ?F i x)" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) have 2: "\x. x \ space M \ summable (\i. \?F i x\)" @@ -2272,7 +2300,7 @@ show ?sums unfolding enum_eq int_f by simp from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] - show "integrable f" unfolding int_f by simp + show "integrable M f" unfolding int_f by simp qed section "Lebesgue integration on finite space" @@ -2280,8 +2308,8 @@ lemma (in measure_space) integral_on_finite: 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 "(\x. f x) = + shows "integrable M f" + and "(\x. f x \M) = (\ r \ f`space M. r * real (\ (f -` {r} \ space M)))" (is "?integral") proof - let "?A r" = "f -` {r} \ space M" @@ -2295,40 +2323,40 @@ finally have "f x = ?S x" . } note f_eq = this - have f_eq_S: "integrable f \ integrable ?S" "integral f = integral ?S" + have f_eq_S: "integrable M f \ integrable M ?S" "integral\<^isup>L M f = integral\<^isup>L M ?S" by (auto intro!: integrable_cong integral_cong simp only: f_eq) - show "integrable f" ?integral using fin f f_eq_S + show "integrable M f" ?integral using fin f f_eq_S by (simp_all add: integral_cmul_indicator borel_measurable_vimage) qed -lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f" +lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function M f" unfolding simple_function_def using finite_space by auto lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \ borel_measurable M" by (auto intro: borel_measurable_simple_function) lemma (in finite_measure_space) positive_integral_finite_eq_setsum: - "positive_integral f = (\x \ space M. f x * \ {x})" + "integral\<^isup>P M f = (\x \ space M. f x * \ {x})" proof - - have *: "positive_integral f = (\\<^isup>+ x. \y\space M. f y * indicator {y} x)" + have *: "integral\<^isup>P M f = (\\<^isup>+ x. (\y\space M. f y * indicator {y} x) \M)" 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) qed lemma (in finite_measure_space) integral_finite_singleton: - shows "integrable f" - and "integral f = (\x \ space M. f x * real (\ {x}))" (is ?I) + shows "integrable M f" + and "integral\<^isup>L M f = (\x \ space M. f x * real (\ {x}))" (is ?I) proof - have [simp]: - "(\\<^isup>+ 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) \M) = (\x \ space M. Real (f x) * \ {x})" + "(\\<^isup>+ x. Real (- f x) \M) = (\x \ space M. Real (- f x) * \ {x})" unfolding positive_integral_finite_eq_setsum by auto - show "integrable f" using finite_space finite_measure + show "integrable M f" using finite_space finite_measure by (simp add: setsum_\ integrable_def) show ?I using finite_measure - apply (simp add: integral_def real_of_pextreal_setsum[symmetric] + apply (simp add: lebesgue_integral_def real_of_pextreal_setsum[symmetric] real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric]) by (rule setsum_cong) (simp_all split: split_if) qed