diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon May 19 12:04:45 2014 +0200 @@ -0,0 +1,2039 @@ +(* Title: HOL/Probability/Nonnegative_Lebesgue_Integration.thy + Author: Johannes Hölzl, TU München + Author: Armin Heller, TU München +*) + +header {* Lebesgue Integration for Nonnegative Functions *} + +theory Nonnegative_Lebesgue_Integration + imports Measure_Space Borel_Space +begin + +lemma indicator_less_ereal[simp]: + "indicator A x \ (indicator B x::ereal) \ (x \ A \ x \ B)" + by (simp add: indicator_def not_le) + +section "Simple function" + +text {* + +Our simple functions are not restricted to positive real numbers. Instead +they are just functions with a finite range and are measurable when singleton +sets are measurable. + +*} + +definition "simple_function M g \ + finite (g ` space M) \ + (\x \ g ` space M. g -` {x} \ space M \ sets M)" + +lemma simple_functionD: + assumes "simple_function M g" + shows "finite (g ` space M)" and "g -` X \ space M \ sets M" +proof - + show "finite (g ` space M)" + using assms unfolding simple_function_def by auto + have "g -` X \ space M = g -` (X \ g`space M) \ space M" by auto + also have "\ = (\x\X \ g`space M. g-`{x} \ space M)" by auto + finally show "g -` X \ space M \ sets M" using assms + by (auto simp del: UN_simps simp: simple_function_def) +qed + +lemma measurable_simple_function[measurable_dest]: + "simple_function M f \ f \ measurable M (count_space UNIV)" + unfolding simple_function_def measurable_def +proof safe + fix A assume "finite (f ` space M)" "\x\f ` space M. f -` {x} \ space M \ sets M" + then have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) \ sets M" + by (intro sets.finite_UN) auto + also have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) = f -` A \ space M" + by (auto split: split_if_asm) + finally show "f -` A \ space M \ sets M" . +qed simp + +lemma borel_measurable_simple_function: + "simple_function M f \ f \ borel_measurable M" + by (auto dest!: measurable_simple_function simp: measurable_def) + +lemma simple_function_measurable2[intro]: + assumes "simple_function M f" "simple_function M g" + shows "f -` A \ g -` B \ space M \ sets M" +proof - + have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" + by auto + then show ?thesis using assms[THEN simple_functionD(2)] by auto +qed + +lemma simple_function_indicator_representation: + fixes f ::"'a \ ereal" + 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 - + have "?r = (\y \ f ` space M. + (if y = f x then y * indicator (f -` {y} \ space M) x else 0))" + by (auto intro!: setsum_cong2) + also have "... = f x * indicator (f -` {f x} \ space M) x" + using assms by (auto dest: simple_functionD simp: setsum_delta) + also have "... = f x" using x by (auto simp: indicator_def) + finally show ?thesis by auto +qed + +lemma simple_function_notspace: + "simple_function M (\x. h x * indicator (- space M) x::ereal)" (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) + have "?h -` {0} \ space M = space M" by auto + thus ?thesis unfolding simple_function_def by auto +qed + +lemma simple_function_cong: + assumes "\t. t \ space M \ f t = g t" + 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" + using assms by (auto intro!: image_eqI) + thus ?thesis unfolding simple_function_def using assms by simp +qed + +lemma 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 simple_function_borel_measurable: + fixes f :: "'a \ 'x::{t2_space}" + assumes "f \ borel_measurable M" and "finite (f ` space M)" + shows "simple_function M f" + using assms unfolding simple_function_def + by (auto intro: borel_measurable_vimage) + +lemma simple_function_eq_measurable: + fixes f :: "'a \ ereal" + shows "simple_function M f \ finite (f`space M) \ f \ measurable M (count_space UNIV)" + using simple_function_borel_measurable[of f] measurable_simple_function[of M f] + by (fastforce simp: simple_function_def) + +lemma simple_function_const[intro, simp]: + "simple_function M (\x. c)" + by (auto intro: finite_subset simp: simple_function_def) +lemma simple_function_compose[intro, simp]: + assumes "simple_function M f" + shows "simple_function M (g \ f)" + unfolding simple_function_def +proof safe + show "finite ((g \ f) ` space M)" + using assms unfolding simple_function_def by (auto simp: image_comp [symmetric]) +next + fix x assume "x \ space M" + let ?G = "g -` {g (f x)} \ (f`space M)" + have *: "(g \ f) -` {(g \ f) x} \ space M = + (\x\?G. f -` {x} \ space M)" by auto + show "(g \ f) -` {(g \ f) x} \ space M \ sets M" + using assms unfolding simple_function_def * + by (rule_tac sets.finite_UN) auto +qed + +lemma simple_function_indicator[intro, simp]: + assumes "A \ sets M" + shows "simple_function M (indicator A)" +proof - + have "indicator A ` space M \ {0, 1}" (is "?S \ _") + by (auto simp: indicator_def) + hence "finite ?S" by (rule finite_subset) simp + moreover have "- A \ space M = space M - A" by auto + ultimately show ?thesis unfolding simple_function_def + using assms by (auto simp: indicator_def [abs_def]) +qed + +lemma simple_function_Pair[intro, simp]: + 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)" + using assms unfolding simple_function_def + by (rule_tac finite_subset[of _ "f`space M \ g`space M"]) auto +next + fix x assume "x \ space M" + have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = + (f -` {f x} \ space M) \ (g -` {g x} \ space M)" + by auto + with `x \ space M` show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" + using assms unfolding simple_function_def by auto +qed + +lemma simple_function_compose1: + 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 simple_function_compose2: + assumes "simple_function M f" and "simple_function M g" + shows "simple_function M (\x. h (f x) (g x))" +proof - + 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 + +lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] + and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] + and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] + and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] + and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] + and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] + and simple_function_max[intro, simp] = simple_function_compose2[where h=max] + +lemma simple_function_setsum[intro, simp]: + 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 simple_function_ereal[intro, simp]: + fixes f g :: "'a \ real" assumes sf: "simple_function M f" + shows "simple_function M (\x. ereal (f x))" + by (auto intro!: simple_function_compose1[OF sf]) + +lemma simple_function_real_of_nat[intro, simp]: + fixes f g :: "'a \ nat" assumes sf: "simple_function M f" + shows "simple_function M (\x. real (f x))" + by (auto intro!: simple_function_compose1[OF sf]) + +lemma borel_measurable_implies_simple_function_sequence: + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" + shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ + (\x. (SUP i. f i x) = max 0 (u x)) \ (\i x. 0 \ f i x)" +proof - + def f \ "\x i. if real i \ u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)" + { fix x j have "f x j \ j * 2 ^ j" unfolding f_def + proof (split split_if, intro conjI impI) + assume "\ real j \ u x" + then have "natfloor (real (u x) * 2 ^ j) \ natfloor (j * 2 ^ j)" + by (cases "u x") (auto intro!: natfloor_mono) + moreover have "real (natfloor (j * 2 ^ j)) \ j * 2^j" + by (intro real_natfloor_le) auto + ultimately show "natfloor (real (u x) * 2 ^ j) \ j * 2 ^ j" + unfolding real_of_nat_le_iff by auto + qed auto } + note f_upper = this + + have real_f: + "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))" + unfolding f_def by auto + + let ?g = "\j x. real (f x j) / 2^j :: ereal" + show ?thesis + proof (intro exI[of _ ?g] conjI allI ballI) + fix i + have "simple_function M (\x. real (f x i))" + proof (intro simple_function_borel_measurable) + show "(\x. real (f x i)) \ borel_measurable M" + using u by (auto simp: real_f) + have "(\x. real (f x i))`space M \ real`{..i*2^i}" + using f_upper[of _ i] by auto + then show "finite ((\x. real (f x i))`space M)" + by (rule finite_subset) auto + qed + then show "simple_function M (?g i)" + by (auto intro: simple_function_ereal simple_function_div) + next + show "incseq ?g" + proof (intro incseq_ereal incseq_SucI le_funI) + fix x and i :: nat + have "f x i * 2 \ f x (Suc i)" unfolding f_def + proof ((split split_if)+, intro conjI impI) + assume "ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" + then show "i * 2 ^ i * 2 \ natfloor (real (u x) * 2 ^ Suc i)" + by (cases "u x") (auto intro!: le_natfloor) + next + assume "\ ereal (real i) \ u x" "ereal (real (Suc i)) \ u x" + then show "natfloor (real (u x) * 2 ^ i) * 2 \ Suc i * 2 ^ Suc i" + by (cases "u x") auto + next + assume "\ ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" + have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2" + by simp + also have "\ \ natfloor (real (u x) * 2 ^ i * 2)" + proof cases + assume "0 \ u x" then show ?thesis + by (intro le_mult_natfloor) + next + assume "\ 0 \ u x" then show ?thesis + by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg) + qed + also have "\ = natfloor (real (u x) * 2 ^ Suc i)" + by (simp add: ac_simps) + finally show "natfloor (real (u x) * 2 ^ i) * 2 \ natfloor (real (u x) * 2 ^ Suc i)" . + qed simp + then show "?g i x \ ?g (Suc i) x" + by (auto simp: field_simps) + qed + next + fix x show "(SUP i. ?g i x) = max 0 (u x)" + proof (rule SUP_eqI) + fix i show "?g i x \ max 0 (u x)" unfolding max_def f_def + by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg + mult_nonpos_nonneg) + next + fix y assume *: "\i. i \ UNIV \ ?g i x \ y" + have "\i. 0 \ ?g i x" by auto + from order_trans[OF this *] have "0 \ y" by simp + show "max 0 (u x) \ y" + proof (cases y) + case (real r) + with * have *: "\i. f x i \ r * 2^i" by (auto simp: divide_le_eq) + from reals_Archimedean2[of r] * have "u x \ \" by (auto simp: f_def) (metis less_le_not_le) + then have "\p. max 0 (u x) = ereal p \ 0 \ p" by (cases "u x") (auto simp: max_def) + then guess p .. note ux = this + obtain m :: nat where m: "p < real m" using reals_Archimedean2 .. + have "p \ r" + proof (rule ccontr) + assume "\ p \ r" + with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] + obtain N where "\n\N. r * 2^n < p * 2^n - 1" by (auto simp: field_simps) + then have "r * 2^max N m < p * 2^max N m - 1" by simp + moreover + have "real (natfloor (p * 2 ^ max N m)) \ r * 2 ^ max N m" + using *[of "max N m"] m unfolding real_f using ux + by (cases "0 \ u x") (simp_all add: max_def split: split_if_asm) + then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" + by (metis real_natfloor_gt_diff_one less_le_trans) + ultimately show False by auto + qed + then show "max 0 (u x) \ y" using real ux by simp + qed (insert `0 \ y`, auto) + qed + qed auto +qed + +lemma borel_measurable_implies_simple_function_sequence': + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" + obtains f where "\i. simple_function M (f i)" "incseq f" "\i. \ \ range (f i)" + "\x. (SUP i. f i x) = max 0 (u x)" "\i x. 0 \ f i x" + using borel_measurable_implies_simple_function_sequence[OF u] by auto + +lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]: + fixes u :: "'a \ ereal" + assumes u: "simple_function M u" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. P u \ P (\x. c * u x)" + assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" + shows "P u" +proof (rule cong) + from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" + proof eventually_elim + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. + qed +next + from u have "finite (u ` space M)" + unfolding simple_function_def by auto + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" + proof induct + case empty show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + qed (auto intro!: add mult set simple_functionD u) +next + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" + apply (subst simple_function_cong) + apply (rule simple_function_indicator_representation[symmetric]) + apply (auto intro: u) + done +qed fact + +lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: + fixes u :: "'a \ ereal" + assumes u: "simple_function M u" and nn: "\x. 0 \ u x" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. 0 \ c \ simple_function M u \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" + assumes add: "\u v. simple_function M u \ (\x. 0 \ u x) \ P u \ simple_function M v \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" + shows "P u" +proof - + show ?thesis + proof (rule cong) + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. + next + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" + apply (subst simple_function_cong) + apply (rule simple_function_indicator_representation[symmetric]) + apply (auto intro: u) + done + next + + from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" + unfolding simple_function_def by auto + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" + proof induct + case empty show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + next + case (insert x S) + { fix z have "(\y\S. y * indicator (u -` {y} \ space M) z) = 0 \ + x * indicator (u -` {x} \ space M) z = 0" + using insert by (subst setsum_ereal_0) (auto simp: indicator_def) } + note disj = this + from insert show ?case + by (auto intro!: add mult set simple_functionD u setsum_nonneg simple_function_setsum disj) + qed + qed fact +qed + +lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" "\x. 0 \ u x" + assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult': "\u c. 0 \ c \ c < \ \ u \ borel_measurable M \ (\x. 0 \ u x) \ (\x. x \ space M \ u x < \) \ P u \ P (\x. c * u x)" + assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ (\x. x \ space M \ u x < \) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ v x < \) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" + assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i x. x \ space M \ U i x < \) \ (\i. P (U i)) \ incseq U \ u = (SUP i. U i) \ P (SUP i. U i)" + shows "P u" + using u +proof (induct rule: borel_measurable_implies_simple_function_sequence') + fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and + sup: "\x. (SUP i. U i x) = max 0 (u x)" and nn: "\i x. 0 \ U i x" + have u_eq: "u = (SUP i. U i)" + using nn u sup by (auto simp: max_def) + + have not_inf: "\x i. x \ space M \ U i x < \" + using U by (auto simp: image_iff eq_commute) + + from U have "\i. U i \ borel_measurable M" + by (simp add: borel_measurable_simple_function) + + show "P u" + unfolding u_eq + proof (rule seq) + fix i show "P (U i)" + using `simple_function M (U i)` nn[of i] not_inf[of _ i] + proof (induct rule: simple_function_induct_nn) + case (mult u c) + show ?case + proof cases + assume "c = 0 \ space M = {} \ (\x\space M. u x = 0)" + with mult(2) show ?thesis + by (intro cong[of "\x. c * u x" "indicator {}"] set) + (auto dest!: borel_measurable_simple_function) + next + assume "\ (c = 0 \ space M = {} \ (\x\space M. u x = 0))" + with mult obtain x where u_fin: "\x. x \ space M \ u x < \" + and x: "x \ space M" "u x \ 0" "c \ 0" + by auto + with mult have "P u" + by auto + from x mult(5)[OF `x \ space M`] mult(1) mult(3)[of x] have "c < \" + by auto + with u_fin mult + show ?thesis + by (intro mult') (auto dest!: borel_measurable_simple_function) + qed + qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) + qed fact+ +qed + +lemma simple_function_If_set: + assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ 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 + proof safe + have "?IF ` space M \ f ` space M \ g ` space M" by auto + from finite_subset[OF this] assms + show "finite (?IF ` space M)" unfolding simple_function_def by auto + next + fix x assume "x \ space M" + then have *: "?IF -` {?IF x} \ space M = (if x \ A + then ((F (f x) \ (A \ space M)) \ (G (f x) - (G (f x) \ (A \ space M)))) + else ((F (g x) \ (A \ space M)) \ (G (g x) - (G (g x) \ (A \ space M)))))" + using sets.sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) + have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" + unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto + show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto + qed +qed + +lemma simple_function_If: + assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" + shows "simple_function M (\x. if P x then f x else g x)" +proof - + have "{x\space M. P x} = {x. P x} \ space M" by auto + with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp +qed + +lemma simple_function_subalgebra: + 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 simple_function_comp: + assumes T: "T \ measurable M M'" + and f: "simple_function M' f" + shows "simple_function M (\x. f (T x))" +proof (intro simple_function_def[THEN iffD2] conjI ballI) + 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 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 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" + using T unfolding measurable_def by auto + finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . +qed + +section "Simple integral" + +definition simple_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>S") where + "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" + +syntax + "_simple_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>S _. _ \_" [60,61] 110) + +translations + "\\<^sup>S x. f \M" == "CONST simple_integral M (%x. f)" + +lemma simple_integral_cong: + assumes "\t. t \ space M \ f t = g t" + shows "integral\<^sup>S M f = integral\<^sup>S M g" +proof - + have "f ` space M = g ` space M" + "\x. f -` {x} \ space M = g -` {x} \ space M" + using assms by (auto intro!: image_eqI) + thus ?thesis unfolding simple_integral_def by simp +qed + +lemma simple_integral_const[simp]: + "(\\<^sup>Sx. c \M) = c * (emeasure M) (space M)" +proof (cases "space M = {}") + case True thus ?thesis unfolding simple_integral_def by simp +next + case False hence "(\x. c) ` space M = {c}" by auto + thus ?thesis unfolding simple_integral_def by simp +qed + +lemma simple_function_partition: + assumes f: "simple_function M f" and g: "simple_function M g" + assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" + assumes v: "\x. x \ space M \ f x = v (g x)" + shows "integral\<^sup>S M f = (\y\g ` space M. v y * emeasure M {x\space M. g x = y})" + (is "_ = ?r") +proof - + from f g have [simp]: "finite (f`space M)" "finite (g`space M)" + by (auto simp: simple_function_def) + from f g have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function) + + { fix y assume "y \ space M" + then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" + by (auto cong: sub simp: v[symmetric]) } + note eq = this + + have "integral\<^sup>S M f = + (\y\f`space M. y * (\z\g`space M. + if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" + unfolding simple_integral_def + proof (safe intro!: setsum_cong ereal_left_mult_cong) + fix y assume y: "y \ space M" "f y \ 0" + have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = + {z. \x\space M. f y = f x \ z = g x}" + by auto + have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = + f -` {f y} \ space M" + by (auto simp: eq_commute cong: sub rev_conj_cong) + have "finite (g`space M)" by simp + then have "finite {z. \x\space M. f y = f x \ z = g x}" + by (rule rev_finite_subset) auto + then show "emeasure M (f -` {f y} \ space M) = + (\z\g ` space M. if \x\space M. f y = f x \ z = g x then emeasure M {x \ space M. g x = z} else 0)" + apply (simp add: setsum_cases) + apply (subst setsum_emeasure) + apply (auto simp: disjoint_family_on_def eq) + done + qed + also have "\ = (\y\f`space M. (\z\g`space M. + if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" + by (auto intro!: setsum_cong simp: setsum_ereal_right_distrib emeasure_nonneg) + also have "\ = ?r" + by (subst setsum_commute) + (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq) + finally show "integral\<^sup>S M f = ?r" . +qed + +lemma simple_integral_add[simp]: + assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" + shows "(\\<^sup>Sx. f x + g x \M) = integral\<^sup>S M f + integral\<^sup>S M g" +proof - + have "(\\<^sup>Sx. f x + g x \M) = + (\y\(\x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\space M. (f x, g x) = y})" + by (intro simple_function_partition) (auto intro: f g) + also have "\ = (\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) + + (\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y})" + using assms(2,4) by (auto intro!: setsum_cong ereal_left_distrib simp: setsum_addf[symmetric]) + also have "(\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. f x \M)" + by (intro simple_function_partition[symmetric]) (auto intro: f g) + also have "(\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. g x \M)" + by (intro simple_function_partition[symmetric]) (auto intro: f g) + finally show ?thesis . +qed + +lemma simple_integral_setsum[simp]: + assumes "\i x. i \ P \ 0 \ f i x" + assumes "\i. i \ P \ simple_function M (f i)" + shows "(\\<^sup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^sup>S M (f i))" +proof cases + assume "finite P" + from this assms show ?thesis + by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) +qed auto + +lemma simple_integral_mult[simp]: + assumes f: "simple_function M f" "\x. 0 \ f x" + shows "(\\<^sup>Sx. c * f x \M) = c * integral\<^sup>S M f" +proof - + have "(\\<^sup>Sx. c * f x \M) = (\y\f ` space M. (c * y) * emeasure M {x\space M. f x = y})" + using f by (intro simple_function_partition) auto + also have "\ = c * integral\<^sup>S M f" + using f unfolding simple_integral_def + by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult_assoc Int_def conj_commute) + finally show ?thesis . +qed + +lemma simple_integral_mono_AE: + assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" + and mono: "AE x in M. f x \ g x" + shows "integral\<^sup>S M f \ integral\<^sup>S M g" +proof - + let ?\ = "\P. emeasure M {x\space M. P x}" + have "integral\<^sup>S M f = (\y\(\x. (f x, g x))`space M. fst y * ?\ (\x. (f x, g x) = y))" + using f g by (intro simple_function_partition) auto + also have "\ \ (\y\(\x. (f x, g x))`space M. snd y * ?\ (\x. (f x, g x) = y))" + proof (clarsimp intro!: setsum_mono) + fix x assume "x \ space M" + let ?M = "?\ (\y. f y = f x \ g y = g x)" + show "f x * ?M \ g x * ?M" + proof cases + assume "?M \ 0" + then have "0 < ?M" + by (simp add: less_le emeasure_nonneg) + also have "\ \ ?\ (\y. f x \ g x)" + using mono by (intro emeasure_mono_AE) auto + finally have "\ \ f x \ g x" + by (intro notI) auto + then show ?thesis + by (intro ereal_mult_right_mono) auto + qed simp + qed + also have "\ = integral\<^sup>S M g" + using f g by (intro simple_function_partition[symmetric]) auto + finally show ?thesis . +qed + +lemma simple_integral_mono: + assumes "simple_function M f" and "simple_function M g" + and mono: "\ x. x \ space M \ f x \ g x" + shows "integral\<^sup>S M f \ integral\<^sup>S M g" + using assms by (intro simple_integral_mono_AE) auto + +lemma simple_integral_cong_AE: + assumes "simple_function M f" and "simple_function M g" + and "AE x in M. f x = g x" + shows "integral\<^sup>S M f = integral\<^sup>S M g" + using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) + +lemma simple_integral_cong': + assumes sf: "simple_function M f" "simple_function M g" + and mea: "(emeasure M) {x\space M. f x \ g x} = 0" + shows "integral\<^sup>S M f = integral\<^sup>S M g" +proof (intro simple_integral_cong_AE sf AE_I) + show "(emeasure M) {x\space M. f x \ g x} = 0" by fact + show "{x \ space M. f x \ g x} \ sets M" + using sf[THEN borel_measurable_simple_function] by auto +qed simp + +lemma simple_integral_indicator: + assumes A: "A \ sets M" + assumes f: "simple_function M f" + shows "(\\<^sup>Sx. f x * indicator A x \M) = + (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" +proof - + have eq: "(\x. (f x, indicator A x)) ` space M \ {x. snd x = 1} = (\x. (f x, 1::ereal))`A" + using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: split_if_asm) + have eq2: "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" + by (auto simp: image_iff) + + have "(\\<^sup>Sx. f x * indicator A x \M) = + (\y\(\x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\space M. (f x, indicator A x) = y})" + using assms by (intro simple_function_partition) auto + also have "\ = (\y\(\x. (f x, indicator A x::ereal))`space M. + if snd y = 1 then fst y * emeasure M (f -` {fst y} \ space M \ A) else 0)" + by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum_cong) + also have "\ = (\y\(\x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \ space M \ A))" + using assms by (subst setsum_cases) (auto intro!: simple_functionD(1) simp: eq) + also have "\ = (\y\fst`(\x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \ space M \ A))" + by (subst setsum_reindex[where f=fst]) (auto simp: inj_on_def) + also have "\ = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" + using A[THEN sets.sets_into_space] + by (intro setsum_mono_zero_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) + finally show ?thesis . +qed + +lemma simple_integral_indicator_only[simp]: + assumes "A \ sets M" + shows "integral\<^sup>S M (indicator A) = emeasure M A" + using simple_integral_indicator[OF assms, of "\x. 1"] sets.sets_into_space[OF assms] + by (simp_all add: image_constant_conv Int_absorb1 split: split_if_asm) + +lemma simple_integral_null_set: + assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" + shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" +proof - + have "AE x in M. indicator N x = (0 :: ereal)" + using `N \ null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) + then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" + using assms apply (intro simple_integral_cong_AE) by auto + then show ?thesis by simp +qed + +lemma simple_integral_cong_AE_mult_indicator: + assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" + shows "integral\<^sup>S M f = (\\<^sup>Sx. f x * indicator S x \M)" + using assms by (intro simple_integral_cong_AE) auto + +lemma simple_integral_cmult_indicator: + assumes A: "A \ sets M" + shows "(\\<^sup>Sx. c * indicator A x \M) = c * emeasure M A" + using simple_integral_mult[OF simple_function_indicator[OF A]] + unfolding simple_integral_indicator_only[OF A] by simp + +lemma simple_integral_positive: + assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" + shows "0 \ integral\<^sup>S M f" +proof - + have "integral\<^sup>S M (\x. 0) \ integral\<^sup>S M f" + using simple_integral_mono_AE[OF _ f ae] by auto + then show ?thesis by simp +qed + +section "Continuous positive integration" + +definition positive_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>P") where + "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^sup>S M g)" + +syntax + "_positive_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>+ _. _ \_" [60,61] 110) + +translations + "\\<^sup>+ x. f \M" == "CONST positive_integral M (%x. f)" + +lemma positive_integral_positive: + "0 \ integral\<^sup>P M f" + by (auto intro!: SUP_upper2[of "\x. 0"] simp: positive_integral_def le_fun_def) + +lemma positive_integral_not_MInfty[simp]: "integral\<^sup>P M f \ -\" + using positive_integral_positive[of M f] by auto + +lemma positive_integral_def_finite: + "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f \ range g \ {0 ..< \}}. integral\<^sup>S M g)" + (is "_ = SUPREMUM ?A ?f") + unfolding positive_integral_def +proof (safe intro!: antisym SUP_least) + fix g assume g: "simple_function M g" "g \ max 0 \ f" + let ?G = "{x \ space M. \ g x \ \}" + note gM = g(1)[THEN borel_measurable_simple_function] + have \_G_pos: "0 \ (emeasure M) ?G" using gM by auto + let ?g = "\y x. if g x = \ then y else max 0 (g x)" + from g gM have g_in_A: "\y. 0 \ y \ y \ \ \ ?g y \ ?A" + apply (safe intro!: simple_function_max simple_function_If) + apply (force simp: max_def le_fun_def split: split_if_asm)+ + done + show "integral\<^sup>S M g \ SUPREMUM ?A ?f" + proof cases + have g0: "?g 0 \ ?A" by (intro g_in_A) auto + assume "(emeasure M) ?G = 0" + with gM have "AE x in M. x \ ?G" + by (auto simp add: AE_iff_null intro!: null_setsI) + with gM g show ?thesis + by (intro SUP_upper2[OF g0] simple_integral_mono_AE) + (auto simp: max_def intro!: simple_function_If) + next + assume \_G: "(emeasure M) ?G \ 0" + have "SUPREMUM ?A (integral\<^sup>S M) = \" + proof (intro SUP_PInfty) + fix n :: nat + let ?y = "ereal (real n) / (if (emeasure M) ?G = \ then 1 else (emeasure M) ?G)" + have "0 \ ?y" "?y \ \" using \_G \_G_pos by (auto simp: ereal_divide_eq) + then have "?g ?y \ ?A" by (rule g_in_A) + have "real n \ ?y * (emeasure M) ?G" + using \_G \_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) + also have "\ = (\\<^sup>Sx. ?y * indicator ?G x \M)" + using `0 \ ?y` `?g ?y \ ?A` gM + by (subst simple_integral_cmult_indicator) auto + also have "\ \ integral\<^sup>S M (?g ?y)" using `?g ?y \ ?A` gM + by (intro simple_integral_mono) auto + finally show "\i\?A. real n \ integral\<^sup>S M i" + using `?g ?y \ ?A` by blast + qed + then show ?thesis by simp + qed +qed (auto intro: SUP_upper) + +lemma positive_integral_mono_AE: + assumes ae: "AE x in M. u x \ v x" shows "integral\<^sup>P M u \ integral\<^sup>P M v" + unfolding positive_integral_def +proof (safe intro!: SUP_mono) + fix n assume n: "simple_function M n" "n \ max 0 \ u" + from ae[THEN AE_E] guess N . note N = this + then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) + let ?n = "\x. n x * indicator (space M - N) x" + have "AE x in M. n x \ ?n x" "simple_function M ?n" + using n N ae_N by auto + moreover + { fix x have "?n x \ max 0 (v x)" + proof cases + assume x: "x \ space M - N" + with N have "u x \ v x" by auto + with n(2)[THEN le_funD, of x] x show ?thesis + by (auto simp: max_def split: split_if_asm) + qed simp } + then have "?n \ max 0 \ v" by (auto simp: le_funI) + moreover have "integral\<^sup>S M n \ integral\<^sup>S M ?n" + using ae_N N n by (auto intro!: simple_integral_mono_AE) + ultimately show "\m\{g. simple_function M g \ g \ max 0 \ v}. integral\<^sup>S M n \ integral\<^sup>S M m" + by force +qed + +lemma positive_integral_mono: + "(\x. x \ space M \ u x \ v x) \ integral\<^sup>P M u \ integral\<^sup>P M v" + by (auto intro: positive_integral_mono_AE) + +lemma positive_integral_cong_AE: + "AE x in M. u x = v x \ integral\<^sup>P M u = integral\<^sup>P M v" + by (auto simp: eq_iff intro!: positive_integral_mono_AE) + +lemma positive_integral_cong: + "(\x. x \ space M \ u x = v x) \ integral\<^sup>P M u = integral\<^sup>P M v" + by (auto intro: positive_integral_cong_AE) + +lemma positive_integral_cong_strong: + "M = N \ (\x. x \ space M \ u x = v x) \ integral\<^sup>P M u = integral\<^sup>P N v" + by (auto intro: positive_integral_cong) + +lemma positive_integral_eq_simple_integral: + assumes f: "simple_function M f" "\x. 0 \ f x" shows "integral\<^sup>P M f = integral\<^sup>S M f" +proof - + let ?f = "\x. f x * indicator (space M) x" + have f': "simple_function M ?f" using f by auto + with f(2) have [simp]: "max 0 \ ?f = ?f" + by (auto simp: fun_eq_iff max_def split: split_indicator) + have "integral\<^sup>P M ?f \ integral\<^sup>S M ?f" using f' + by (force intro!: SUP_least simple_integral_mono simp: le_fun_def positive_integral_def) + moreover have "integral\<^sup>S M ?f \ integral\<^sup>P M ?f" + unfolding positive_integral_def + using f' by (auto intro!: SUP_upper) + ultimately show ?thesis + by (simp cong: positive_integral_cong simple_integral_cong) +qed + +lemma positive_integral_eq_simple_integral_AE: + assumes f: "simple_function M f" "AE x in M. 0 \ f x" shows "integral\<^sup>P M f = integral\<^sup>S M f" +proof - + have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) + with f have "integral\<^sup>P M f = integral\<^sup>S M (\x. max 0 (f x))" + by (simp cong: positive_integral_cong_AE simple_integral_cong_AE + add: positive_integral_eq_simple_integral) + with assms show ?thesis + by (auto intro!: simple_integral_cong_AE split: split_max) +qed + +lemma positive_integral_SUP_approx: + assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + and u: "simple_function M u" "u \ (SUP i. f i)" "u`space M \ {0..<\}" + shows "integral\<^sup>S M u \ (SUP i. integral\<^sup>P M (f i))" (is "_ \ ?S") +proof (rule ereal_le_mult_one_interval) + have "0 \ (SUP i. integral\<^sup>P M (f i))" + using f(3) by (auto intro!: SUP_upper2 positive_integral_positive) + then show "(SUP i. integral\<^sup>P M (f i)) \ -\" by auto + have u_range: "\x. x \ space M \ 0 \ u x \ u x \ \" + using u(3) by auto + fix a :: ereal 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 M u`[THEN borel_measurable_simple_function] by auto + + let ?uB = "\i x. u x * indicator (?B i) x" + + { fix i have "?B i \ ?B (Suc i)" + proof safe + fix i x assume "a * u x \ f i x" + also have "\ \ f (Suc i) x" + using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto + finally show "a * u x \ f (Suc i) x" . + qed } + note B_mono = this + + note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B] + + let ?B' = "\i n. (u -` {i} \ space M) \ ?B n" + have measure_conv: "\i. (emeasure M) (u -` {i} \ space M) = (SUP n. (emeasure M) (?B' i n))" + proof - + fix i + have 1: "range (?B' i) \ sets M" using B_u by auto + have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI) + have "(\n. ?B' i n) = u -` {i} \ space M" + proof safe + fix x i assume x: "x \ space M" + show "x \ (\i. ?B' (u x) i)" + proof cases + assume "u x = 0" thus ?thesis using `x \ space M` f(3) by simp + next + assume "u x \ 0" + with `a < 1` u_range[OF `x \ space M`] + have "a * u x < 1 * u x" + by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) + also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) + finally obtain i where "a * u x < f i x" unfolding SUP_def + by (auto simp add: less_SUP_iff) + hence "a * u x \ f i x" by auto + thus ?thesis using `x \ space M` by auto + qed + qed + then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp + qed + + have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" + unfolding simple_integral_indicator[OF B `simple_function M u`] + proof (subst SUP_ereal_setsum, safe) + fix x n assume "x \ space M" + with u_range show "incseq (\i. u x * (emeasure M) (?B' (u x) i))" "\i. 0 \ u x * (emeasure M) (?B' (u x) i)" + using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) + next + show "integral\<^sup>S M u = (\i\u ` space M. SUP n. i * (emeasure M) (?B' i n))" + using measure_conv u_range B_u unfolding simple_integral_def + by (auto intro!: setsum_cong SUP_ereal_cmult [symmetric]) + qed + moreover + have "a * (SUP i. integral\<^sup>S M (?uB i)) \ ?S" + apply (subst SUP_ereal_cmult [symmetric]) + proof (safe intro!: SUP_mono bexI) + fix i + have "a * integral\<^sup>S M (?uB i) = (\\<^sup>Sx. a * ?uB i x \M)" + using B `simple_function M u` u_range + by (subst simple_integral_mult) (auto split: split_indicator) + also have "\ \ integral\<^sup>P M (f i)" + proof - + have *: "simple_function M (\x. a * ?uB i x)" using B `0 < a` u(1) by auto + show ?thesis using f(3) * u_range `0 < a` + by (subst positive_integral_eq_simple_integral[symmetric]) + (auto intro!: positive_integral_mono split: split_indicator) + qed + finally show "a * integral\<^sup>S M (?uB i) \ integral\<^sup>P M (f i)" + by auto + next + fix i show "0 \ \\<^sup>S x. ?uB i x \M" using B `0 < a` u(1) u_range + by (intro simple_integral_positive) (auto split: split_indicator) + qed (insert `0 < a`, auto) + ultimately show "a * integral\<^sup>S M u \ ?S" by simp +qed + +lemma incseq_positive_integral: + assumes "incseq f" shows "incseq (\i. integral\<^sup>P M (f i))" +proof - + have "\i x. f i x \ f (Suc i) x" + using assms by (auto dest!: incseq_SucD simp: le_fun_def) + then show ?thesis + by (auto intro!: incseq_SucI positive_integral_mono) +qed + +text {* Beppo-Levi monotone convergence theorem *} +lemma positive_integral_monotone_convergence_SUP: + assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" +proof (rule antisym) + show "(SUP j. integral\<^sup>P M (f j)) \ (\\<^sup>+ x. (SUP i. f i x) \M)" + by (auto intro!: SUP_least SUP_upper positive_integral_mono) +next + show "(\\<^sup>+ x. (SUP i. f i x) \M) \ (SUP j. integral\<^sup>P M (f j))" + unfolding positive_integral_def_finite[of _ "\x. SUP i. f i x"] + proof (safe intro!: SUP_least) + fix g assume g: "simple_function M g" + and *: "g \ max 0 \ (\x. SUP i. f i x)" "range g \ {0..<\}" + then have "\x. 0 \ (SUP i. f i x)" and g': "g`space M \ {0..<\}" + using f by (auto intro!: SUP_upper2) + with * show "integral\<^sup>S M g \ (SUP j. integral\<^sup>P M (f j))" + by (intro positive_integral_SUP_approx[OF f g _ g']) + (auto simp: le_fun_def max_def) + qed +qed + +lemma positive_integral_monotone_convergence_SUP_AE: + assumes f: "\i. AE x in M. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" + shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" +proof - + from f have "AE x in M. \i. f i x \ f (Suc i) x \ 0 \ f i x" + by (simp add: AE_all_countable) + from this[THEN AE_E] guess N . note N = this + let ?f = "\i x. if x \ space M - N then f i x else 0" + have f_eq: "AE x in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) + then have "(\\<^sup>+ x. (SUP i. f i x) \M) = (\\<^sup>+ x. (SUP i. ?f i x) \M)" + by (auto intro!: positive_integral_cong_AE) + also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \M))" + proof (rule positive_integral_monotone_convergence_SUP) + show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) + { fix i show "(\x. if x \ space M - N then f i x else 0) \ borel_measurable M" + using f N(3) by (intro measurable_If_set) auto + fix x show "0 \ ?f i x" + using N(1) by auto } + qed + also have "\ = (SUP i. (\\<^sup>+ x. f i x \M))" + using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] positive_integral_cong_AE ext) + finally show ?thesis . +qed + +lemma positive_integral_monotone_convergence_SUP_AE_incseq: + assumes f: "incseq f" "\i. AE x in M. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" + shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" + using f[unfolded incseq_Suc_iff le_fun_def] + by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel]) + auto + +lemma positive_integral_monotone_convergence_simple: + assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" + shows "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" + using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1) + f(3)[THEN borel_measurable_simple_function] f(2)] + by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPREMUM UNIV"] ext) + +lemma positive_integral_max_0: + "(\\<^sup>+x. max 0 (f x) \M) = integral\<^sup>P M f" + by (simp add: le_fun_def positive_integral_def) + +lemma positive_integral_cong_pos: + assumes "\x. x \ space M \ f x \ 0 \ g x \ 0 \ f x = g x" + shows "integral\<^sup>P M f = integral\<^sup>P M g" +proof - + have "integral\<^sup>P M (\x. max 0 (f x)) = integral\<^sup>P M (\x. max 0 (g x))" + proof (intro positive_integral_cong) + fix x assume "x \ space M" + from assms[OF this] show "max 0 (f x) = max 0 (g x)" + by (auto split: split_max) + qed + then show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma SUP_simple_integral_sequences: + assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" + and g: "incseq g" "\i x. 0 \ g i x" "\i. simple_function M (g i)" + and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" + shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" + (is "SUPREMUM _ ?F = SUPREMUM _ ?G") +proof - + have "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" + using f by (rule positive_integral_monotone_convergence_simple) + also have "\ = (\\<^sup>+x. (SUP i. g i x) \M)" + unfolding eq[THEN positive_integral_cong_AE] .. + also have "\ = (SUP i. ?G i)" + using g by (rule positive_integral_monotone_convergence_simple[symmetric]) + finally show ?thesis by simp +qed + +lemma positive_integral_const[simp]: + "0 \ c \ (\\<^sup>+ x. c \M) = c * (emeasure M) (space M)" + by (subst positive_integral_eq_simple_integral) auto + +lemma positive_integral_linear: + assumes f: "f \ borel_measurable M" "\x. 0 \ f x" and "0 \ a" + and g: "g \ borel_measurable M" "\x. 0 \ g x" + shows "(\\<^sup>+ x. a * f x + g x \M) = a * integral\<^sup>P M f + integral\<^sup>P M g" + (is "integral\<^sup>P M ?L = _") +proof - + from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . + note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . + note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + let ?L' = "\i x. a * u i x + v i x" + + have "?L \ borel_measurable M" using assms by auto + from borel_measurable_implies_simple_function_sequence'[OF this] guess l . + note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + + have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" + using u v `0 \ a` + by (auto simp: incseq_Suc_iff le_fun_def + intro!: add_mono ereal_mult_left_mono simple_integral_mono) + have pos: "\i. 0 \ integral\<^sup>S M (u i)" "\i. 0 \ integral\<^sup>S M (v i)" "\i. 0 \ a * integral\<^sup>S M (u i)" + using u v `0 \ a` by (auto simp: simple_integral_positive) + { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \ -\" "integral\<^sup>S M (v i) \ -\" + by (auto split: split_if_asm) } + note not_MInf = this + + have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" + proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) + show "incseq ?L'" "\i x. 0 \ ?L' i x" "\i. simple_function M (?L' i)" + using u v `0 \ a` unfolding incseq_Suc_iff le_fun_def + by (auto intro!: add_mono ereal_mult_left_mono) + { fix x + { fix i have "a * u i x \ -\" "v i x \ -\" "u i x \ -\" using `0 \ a` u(6)[of i x] v(6)[of i x] + by auto } + then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" + using `0 \ a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] + by (subst SUP_ereal_cmult [symmetric, OF u(6) `0 \ a`]) + (auto intro!: SUP_ereal_add + simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } + then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" + unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) + by (intro AE_I2) (auto split: split_max) + qed + also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" + using u(2, 6) v(2, 6) `0 \ a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) + finally have "(\\<^sup>+ x. max 0 (a * f x + g x) \M) = a * (\\<^sup>+x. max 0 (f x) \M) + (\\<^sup>+x. max 0 (g x) \M)" + unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] + unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] + apply (subst SUP_ereal_cmult [symmetric, OF pos(1) `0 \ a`]) + apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) . + then show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma positive_integral_cmult: + assumes f: "f \ borel_measurable M" "0 \ c" + shows "(\\<^sup>+ x. c * f x \M) = c * integral\<^sup>P M f" +proof - + have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` + by (auto split: split_max simp: ereal_zero_le_0_iff) + have "(\\<^sup>+ x. c * f x \M) = (\\<^sup>+ x. c * max 0 (f x) \M)" + by (simp add: positive_integral_max_0) + then show ?thesis + using positive_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" _ "\x. 0"] f + by (auto simp: positive_integral_max_0) +qed + +lemma positive_integral_multc: + assumes "f \ borel_measurable M" "0 \ c" + shows "(\\<^sup>+ x. f x * c \M) = integral\<^sup>P M f * c" + unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp + +lemma positive_integral_indicator[simp]: + "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" + by (subst positive_integral_eq_simple_integral) + (auto simp: simple_integral_indicator) + +lemma positive_integral_cmult_indicator: + "0 \ c \ A \ sets M \ (\\<^sup>+ x. c * indicator A x \M) = c * (emeasure M) A" + by (subst positive_integral_eq_simple_integral) + (auto simp: simple_function_indicator simple_integral_indicator) + +lemma positive_integral_indicator': + assumes [measurable]: "A \ space M \ sets M" + shows "(\\<^sup>+ x. indicator A x \M) = emeasure M (A \ space M)" +proof - + have "(\\<^sup>+ x. indicator A x \M) = (\\<^sup>+ x. indicator (A \ space M) x \M)" + by (intro positive_integral_cong) (simp split: split_indicator) + also have "\ = emeasure M (A \ space M)" + by simp + finally show ?thesis . +qed + +lemma positive_integral_add: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + shows "(\\<^sup>+ x. f x + g x \M) = integral\<^sup>P M f + integral\<^sup>P M g" +proof - + have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" + using assms by (auto split: split_max) + have "(\\<^sup>+ x. f x + g x \M) = (\\<^sup>+ x. max 0 (f x + g x) \M)" + by (simp add: positive_integral_max_0) + also have "\ = (\\<^sup>+ x. max 0 (f x) + max 0 (g x) \M)" + unfolding ae[THEN positive_integral_cong_AE] .. + also have "\ = (\\<^sup>+ x. max 0 (f x) \M) + (\\<^sup>+ x. max 0 (g x) \M)" + using positive_integral_linear[of "\x. max 0 (f x)" _ 1 "\x. max 0 (g x)"] f g + by auto + finally show ?thesis + by (simp add: positive_integral_max_0) +qed + +lemma positive_integral_setsum: + assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x in M. 0 \ f i x" + shows "(\\<^sup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^sup>P M (f i))" +proof cases + assume f: "finite P" + from assms have "AE x in M. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto + from f this assms(1) show ?thesis + proof induct + case (insert i P) + then have "f i \ borel_measurable M" "AE x in M. 0 \ f i x" + "(\x. \i\P. f i x) \ borel_measurable M" "AE x in M. 0 \ (\i\P. f i x)" + by (auto intro!: setsum_nonneg) + from positive_integral_add[OF this] + show ?case using insert by auto + qed simp +qed simp + +lemma positive_integral_Markov_inequality: + assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" and "A \ sets M" and c: "0 \ c" + shows "(emeasure M) ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^sup>+ x. u x * indicator A x \M)" + (is "(emeasure M) ?A \ _ * ?PI") +proof - + have "?A \ sets M" + using `A \ sets M` u by auto + hence "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" + using positive_integral_indicator by simp + also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u c + by (auto intro!: positive_integral_mono_AE + simp: indicator_def ereal_zero_le_0_iff) + also have "\ = c * (\\<^sup>+ x. u x * indicator A x \M)" + using assms + by (auto intro!: positive_integral_cmult simp: ereal_zero_le_0_iff) + finally show ?thesis . +qed + +lemma positive_integral_noteq_infinite: + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + and "integral\<^sup>P M g \ \" + shows "AE x in M. g x \ \" +proof (rule ccontr) + assume c: "\ (AE x in M. g x \ \)" + have "(emeasure M) {x\space M. g x = \} \ 0" + using c g by (auto simp add: AE_iff_null) + moreover have "0 \ (emeasure M) {x\space M. g x = \}" using g by (auto intro: measurable_sets) + ultimately have "0 < (emeasure M) {x\space M. g x = \}" by auto + then have "\ = \ * (emeasure M) {x\space M. g x = \}" by auto + also have "\ \ (\\<^sup>+x. \ * indicator {x\space M. g x = \} x \M)" + using g by (subst positive_integral_cmult_indicator) auto + also have "\ \ integral\<^sup>P M g" + using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def) + finally show False using `integral\<^sup>P M g \ \` by auto +qed + +lemma positive_integral_PInf: + assumes f: "f \ borel_measurable M" + and not_Inf: "integral\<^sup>P M f \ \" + shows "(emeasure M) (f -` {\} \ space M) = 0" +proof - + have "\ * (emeasure M) (f -` {\} \ space M) = (\\<^sup>+ x. \ * indicator (f -` {\} \ space M) x \M)" + using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets) + also have "\ \ integral\<^sup>P M (\x. max 0 (f x))" + by (auto intro!: positive_integral_mono simp: indicator_def max_def) + finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^sup>P M f" + by (simp add: positive_integral_max_0) + moreover have "0 \ (emeasure M) (f -` {\} \ space M)" + by (rule emeasure_nonneg) + ultimately show ?thesis + using assms by (auto split: split_if_asm) +qed + +lemma positive_integral_PInf_AE: + assumes "f \ borel_measurable M" "integral\<^sup>P M f \ \" shows "AE x in M. f x \ \" +proof (rule AE_I) + show "(emeasure M) (f -` {\} \ space M) = 0" + by (rule positive_integral_PInf[OF assms]) + show "f -` {\} \ space M \ sets M" + using assms by (auto intro: borel_measurable_vimage) +qed auto + +lemma simple_integral_PInf: + assumes "simple_function M f" "\x. 0 \ f x" + and "integral\<^sup>S M f \ \" + shows "(emeasure M) (f -` {\} \ space M) = 0" +proof (rule positive_integral_PInf) + show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) + show "integral\<^sup>P M f \ \" + using assms by (simp add: positive_integral_eq_simple_integral) +qed + +lemma positive_integral_diff: + assumes f: "f \ borel_measurable M" + and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + and fin: "integral\<^sup>P M g \ \" + and mono: "AE x in M. g x \ f x" + shows "(\\<^sup>+ x. f x - g x \M) = integral\<^sup>P M f - integral\<^sup>P M g" +proof - + have diff: "(\x. f x - g x) \ borel_measurable M" "AE x in M. 0 \ f x - g x" + using assms by (auto intro: ereal_diff_positive) + have pos_f: "AE x in M. 0 \ f x" using mono g by auto + { fix a b :: ereal assume "0 \ a" "a \ \" "0 \ b" "a \ b" then have "b - a + a = b" + by (cases rule: ereal2_cases[of a b]) auto } + note * = this + then have "AE x in M. f x = f x - g x + g x" + using mono positive_integral_noteq_infinite[OF g fin] assms by auto + then have **: "integral\<^sup>P M f = (\\<^sup>+x. f x - g x \M) + integral\<^sup>P M g" + unfolding positive_integral_add[OF diff g, symmetric] + by (rule positive_integral_cong_AE) + show ?thesis unfolding ** + using fin positive_integral_positive[of M g] + by (cases rule: ereal2_cases[of "\\<^sup>+ x. f x - g x \M" "integral\<^sup>P M g"]) auto +qed + +lemma positive_integral_suminf: + assumes f: "\i. f i \ borel_measurable M" "\i. AE x in M. 0 \ f i x" + shows "(\\<^sup>+ x. (\i. f i x) \M) = (\i. integral\<^sup>P M (f i))" +proof - + have all_pos: "AE x in M. \i. 0 \ f i x" + using assms by (auto simp: AE_all_countable) + have "(\i. integral\<^sup>P M (f i)) = (SUP n. \iP M (f i))" + using positive_integral_positive by (rule suminf_ereal_eq_SUP) + also have "\ = (SUP n. \\<^sup>+x. (\iM)" + unfolding positive_integral_setsum[OF f] .. + also have "\ = \\<^sup>+x. (SUP n. \iM" using f all_pos + by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) + (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3) + also have "\ = \\<^sup>+x. (\i. f i x) \M" using all_pos + by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUP) + finally show ?thesis by simp +qed + +lemma positive_integral_mult_bounded_inf: + assumes f: "f \ borel_measurable M" "(\\<^sup>+x. f x \M) < \" + and c: "0 \ c" "c \ \" and ae: "AE x in M. g x \ c * f x" + shows "(\\<^sup>+x. g x \M) < \" +proof - + have "(\\<^sup>+x. g x \M) \ (\\<^sup>+x. c * f x \M)" + by (intro positive_integral_mono_AE ae) + also have "(\\<^sup>+x. c * f x \M) < \" + using c f by (subst positive_integral_cmult) auto + finally show ?thesis . +qed + +text {* Fatou's lemma: convergence theorem on limes inferior *} + +lemma positive_integral_liminf: + fixes u :: "nat \ 'a \ ereal" + assumes u: "\i. u i \ borel_measurable M" "\i. AE x in M. 0 \ u i x" + shows "(\\<^sup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>P M (u n))" +proof - + have pos: "AE x in M. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) + have "(\\<^sup>+ x. liminf (\n. u n x) \M) = + (SUP n. \\<^sup>+ x. (INF i:{n..}. u i x) \M)" + unfolding liminf_SUP_INF using pos u + by (intro positive_integral_monotone_convergence_SUP_AE) + (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) + also have "\ \ liminf (\n. integral\<^sup>P M (u n))" + unfolding liminf_SUP_INF + by (auto intro!: SUP_mono exI INF_greatest positive_integral_mono INF_lower) + finally show ?thesis . +qed + +lemma le_Limsup: + "F \ bot \ eventually (\x. c \ g x) F \ c \ Limsup F g" + using Limsup_mono[of "\_. c" g F] by (simp add: Limsup_const) + +lemma Limsup_le: + "F \ bot \ eventually (\x. f x \ c) F \ Limsup F f \ c" + using Limsup_mono[of f "\_. c" F] by (simp add: Limsup_const) + +lemma ereal_mono_minus_cancel: + fixes a b c :: ereal + shows "c - a \ c - b \ 0 \ c \ c < \ \ b \ a" + by (cases a b c rule: ereal3_cases) auto + +lemma positive_integral_limsup: + fixes u :: "nat \ 'a \ ereal" + assumes [measurable]: "\i. u i \ borel_measurable M" "w \ borel_measurable M" + assumes bounds: "\i. AE x in M. 0 \ u i x" "\i. AE x in M. u i x \ w x" and w: "(\\<^sup>+x. w x \M) < \" + shows "limsup (\n. integral\<^sup>P M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" +proof - + have bnd: "AE x in M. \i. 0 \ u i x \ u i x \ w x" + using bounds by (auto simp: AE_all_countable) + + from bounds[of 0] have w_nonneg: "AE x in M. 0 \ w x" + by auto + + have "(\\<^sup>+x. w x \M) - (\\<^sup>+x. limsup (\n. u n x) \M) = (\\<^sup>+x. w x - limsup (\n. u n x) \M)" + proof (intro positive_integral_diff[symmetric]) + show "AE x in M. 0 \ limsup (\n. u n x)" + using bnd by (auto intro!: le_Limsup) + show "AE x in M. limsup (\n. u n x) \ w x" + using bnd by (auto intro!: Limsup_le) + then have "(\\<^sup>+x. limsup (\n. u n x) \M) < \" + by (intro positive_integral_mult_bounded_inf[OF _ w, of 1]) auto + then show "(\\<^sup>+x. limsup (\n. u n x) \M) \ \" + by simp + qed auto + also have "\ = (\\<^sup>+x. liminf (\n. w x - u n x) \M)" + using w_nonneg + by (intro positive_integral_cong_AE, eventually_elim) + (auto intro!: liminf_ereal_cminus[symmetric]) + also have "\ \ liminf (\n. \\<^sup>+x. w x - u n x \M)" + proof (rule positive_integral_liminf) + fix i show "AE x in M. 0 \ w x - u i x" + using bounds[of i] by eventually_elim (auto intro: ereal_diff_positive) + qed simp + also have "(\n. \\<^sup>+x. w x - u n x \M) = (\n. (\\<^sup>+x. w x \M) - (\\<^sup>+x. u n x \M))" + proof (intro ext positive_integral_diff) + fix i have "(\\<^sup>+x. u i x \M) < \" + using bounds by (intro positive_integral_mult_bounded_inf[OF _ w, of 1]) auto + then show "(\\<^sup>+x. u i x \M) \ \" by simp + qed (insert bounds, auto) + also have "liminf (\n. (\\<^sup>+x. w x \M) - (\\<^sup>+x. u n x \M)) = (\\<^sup>+x. w x \M) - limsup (\n. \\<^sup>+x. u n x \M)" + using w by (intro liminf_ereal_cminus) auto + finally show ?thesis + by (rule ereal_mono_minus_cancel) (intro w positive_integral_positive)+ +qed + +lemma positive_integral_dominated_convergence: + assumes [measurable]: + "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" + and bound: "\j. AE x in M. 0 \ u j x" "\j. AE x in M. u j x \ w x" + and w: "(\\<^sup>+x. w x \M) < \" + and u': "AE x in M. (\i. u i x) ----> u' x" + shows "(\i. (\\<^sup>+x. u i x \M)) ----> (\\<^sup>+x. u' x \M)" +proof - + have "limsup (\n. integral\<^sup>P M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" + by (intro positive_integral_limsup[OF _ _ bound w]) auto + moreover have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" + using u' by (intro positive_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) + moreover have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" + using u' by (intro positive_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) + moreover have "(\\<^sup>+x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>P M (u n))" + by (intro positive_integral_liminf[OF _ bound(1)]) auto + moreover have "liminf (\n. integral\<^sup>P M (u n)) \ limsup (\n. integral\<^sup>P M (u n))" + by (intro Liminf_le_Limsup sequentially_bot) + ultimately show ?thesis + by (intro Liminf_eq_Limsup) auto +qed + +lemma positive_integral_null_set: + assumes "N \ null_sets M" shows "(\\<^sup>+ x. u x * indicator N x \M) = 0" +proof - + have "(\\<^sup>+ x. u x * indicator N x \M) = (\\<^sup>+ 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) + show "(emeasure M) N = 0" "N \ sets M" + using assms by auto + qed + then show ?thesis by simp +qed + +lemma positive_integral_0_iff: + assumes u: "u \ borel_measurable M" and pos: "AE x in M. 0 \ u x" + shows "integral\<^sup>P M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" + (is "_ \ (emeasure M) ?A = 0") +proof - + have u_eq: "(\\<^sup>+ x. u x * indicator ?A x \M) = integral\<^sup>P M u" + by (auto intro!: positive_integral_cong simp: indicator_def) + show ?thesis + proof + assume "(emeasure M) ?A = 0" + with positive_integral_null_set[of ?A M u] u + show "integral\<^sup>P M u = 0" by (simp add: u_eq null_sets_def) + next + { fix r :: ereal and n :: nat assume gt_1: "1 \ real n * r" + then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) + then have "0 \ r" by (auto simp add: ereal_zero_less_0_iff) } + note gt_1 = this + assume *: "integral\<^sup>P M u = 0" + let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" + have "0 = (SUP n. (emeasure M) (?M n \ ?A))" + proof - + { fix n :: nat + from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] + have "(emeasure M) (?M n \ ?A) \ 0" unfolding u_eq * using u by simp + moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto + ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } + thus ?thesis by simp + qed + also have "\ = (emeasure M) (\n. ?M n \ ?A)" + proof (safe intro!: SUP_emeasure_incseq) + fix n show "?M n \ ?A \ sets M" + using u by (auto intro!: sets.Int) + next + show "incseq (\n. {x \ space M. 1 \ real n * u x} \ {x \ space M. u x \ 0})" + proof (safe intro!: incseq_SucI) + fix n :: nat and x + assume *: "1 \ real n * u x" + also from gt_1[OF *] have "real n * u x \ real (Suc n) * u x" + using `0 \ u x` by (auto intro!: ereal_mult_right_mono) + finally show "1 \ real (Suc n) * u x" by auto + qed + qed + also have "\ = (emeasure M) {x\space M. 0 < u x}" + proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1) + fix x assume "0 < u x" and [simp, intro]: "x \ space M" + show "x \ (\n. ?M n \ ?A)" + proof (cases "u x") + case (real r) with `0 < u x` have "0 < r" by auto + obtain j :: nat where "1 / r \ real j" using real_arch_simple .. + hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using `0 < r` by auto + hence "1 \ real j * r" using real `0 < r` by auto + thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) + qed (insert `0 < u x`, auto) + qed auto + finally have "(emeasure M) {x\space M. 0 < u x} = 0" by simp + moreover + from pos have "AE x in M. \ (u x < 0)" by auto + then have "(emeasure M) {x\space M. u x < 0} = 0" + using AE_iff_null[of M] u by auto + moreover have "(emeasure M) {x\space M. u x \ 0} = (emeasure M) {x\space M. u x < 0} + (emeasure M) {x\space M. 0 < u x}" + using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) + ultimately show "(emeasure M) ?A = 0" by simp + qed +qed + +lemma positive_integral_0_iff_AE: + assumes u: "u \ borel_measurable M" + shows "integral\<^sup>P M u = 0 \ (AE x in M. u x \ 0)" +proof - + have sets: "{x\space M. max 0 (u x) \ 0} \ sets M" + using u by auto + from positive_integral_0_iff[of "\x. max 0 (u x)"] + have "integral\<^sup>P M u = 0 \ (AE x in M. max 0 (u x) = 0)" + unfolding positive_integral_max_0 + using AE_iff_null[OF sets] u by auto + also have "\ \ (AE x in M. u x \ 0)" by (auto split: split_max) + finally show ?thesis . +qed + +lemma AE_iff_positive_integral: + "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>P M (indicator {x. \ P x}) = 0" + by (subst positive_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def + sets.sets_Collect_neg indicator_def[abs_def] measurable_If) + +lemma positive_integral_const_If: + "(\\<^sup>+x. a \M) = (if 0 \ a then a * (emeasure M) (space M) else 0)" + by (auto intro!: positive_integral_0_iff_AE[THEN iffD2]) + +lemma positive_integral_subalgebra: + assumes f: "f \ borel_measurable N" "\x. 0 \ f x" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" + shows "integral\<^sup>P N f = integral\<^sup>P M f" +proof - + have [simp]: "\f :: 'a \ ereal. f \ borel_measurable N \ f \ borel_measurable M" + using N by (auto simp: measurable_def) + have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" + using N by (auto simp add: eventually_ae_filter null_sets_def) + have [simp]: "\A. A \ sets N \ A \ sets M" + using N by auto + from f show ?thesis + apply induct + apply (simp_all add: positive_integral_add positive_integral_cmult positive_integral_monotone_convergence_SUP N) + apply (auto intro!: positive_integral_cong cong: positive_integral_cong simp: N(2)[symmetric]) + done +qed + +lemma positive_integral_nat_function: + fixes f :: "'a \ nat" + assumes "f \ measurable M (count_space UNIV)" + shows "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\t. emeasure M {x\space M. t < f x})" +proof - + def F \ "\i. {x\space M. i < f x}" + with assms have [measurable]: "\i. F i \ sets M" + by auto + + { fix x assume "x \ space M" + have "(\i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" + using sums_If_finite[of "\i. i < f x" "\_. 1::real"] by simp + then have "(\i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))" + unfolding sums_ereal . + moreover have "\i. ereal (if i < f x then 1 else 0) = indicator (F i) x" + using `x \ space M` by (simp add: one_ereal_def F_def) + ultimately have "ereal(of_nat(f x)) = (\i. indicator (F i) x)" + by (simp add: sums_iff) } + then have "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" + by (simp cong: positive_integral_cong) + also have "\ = (\i. emeasure M (F i))" + by (simp add: positive_integral_suminf) + finally show ?thesis + by (simp add: F_def) +qed + +section {* Distributions *} + +lemma positive_integral_distr': + assumes T: "T \ measurable M M'" + and f: "f \ borel_measurable (distr M M' T)" "\x. 0 \ f x" + shows "integral\<^sup>P (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" + using f +proof induct + case (cong f g) + with T show ?case + apply (subst positive_integral_cong[of _ f g]) + apply simp + apply (subst positive_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) + apply (simp add: measurable_def Pi_iff) + apply simp + done +next + case (set A) + then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" + by (auto simp: indicator_def) + from set T show ?case + by (subst positive_integral_cong[OF eq]) + (auto simp add: emeasure_distr intro!: positive_integral_indicator[symmetric] measurable_sets) +qed (simp_all add: measurable_compose[OF T] T positive_integral_cmult positive_integral_add + positive_integral_monotone_convergence_SUP le_fun_def incseq_def) + +lemma positive_integral_distr: + "T \ measurable M M' \ f \ borel_measurable M' \ integral\<^sup>P (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" + by (subst (1 2) positive_integral_max_0[symmetric]) + (simp add: positive_integral_distr') + +section {* Lebesgue integration on @{const count_space} *} + +lemma simple_function_count_space[simp]: + "simple_function (count_space A) f \ finite (f ` A)" + unfolding simple_function_def by simp + +lemma positive_integral_count_space: + assumes A: "finite {a\A. 0 < f a}" + shows "integral\<^sup>P (count_space A) f = (\a|a\A \ 0 < f a. f a)" +proof - + have *: "(\\<^sup>+x. max 0 (f x) \count_space A) = + (\\<^sup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" + by (auto intro!: positive_integral_cong + simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less) + also have "\ = (\a|a\A \ 0 < f a. \\<^sup>+ x. f a * indicator {a} x \count_space A)" + by (subst positive_integral_setsum) + (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) + also have "\ = (\a|a\A \ 0 < f a. f a)" + by (auto intro!: setsum_cong simp: positive_integral_cmult_indicator one_ereal_def[symmetric]) + finally show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma positive_integral_count_space_finite: + "finite A \ (\\<^sup>+x. f x \count_space A) = (\a\A. max 0 (f a))" + by (subst positive_integral_max_0[symmetric]) + (auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le) + +lemma emeasure_UN_countable: + assumes sets: "\i. i \ I \ X i \ sets M" and I: "countable I" + assumes disj: "disjoint_family_on X I" + shows "emeasure M (UNION I X) = (\\<^sup>+i. emeasure M (X i) \count_space I)" +proof cases + assume "finite I" with sets disj show ?thesis + by (subst setsum_emeasure[symmetric]) + (auto intro!: setsum_cong simp add: max_def subset_eq positive_integral_count_space_finite emeasure_nonneg) +next + assume f: "\ finite I" + then have [intro]: "I \ {}" by auto + from from_nat_into_inj_infinite[OF I f] from_nat_into[OF this] disj + have disj2: "disjoint_family (\i. X (from_nat_into I i))" + unfolding disjoint_family_on_def by metis + + from f have "bij_betw (from_nat_into I) UNIV I" + using bij_betw_from_nat_into[OF I] by simp + then have "(\i\I. X i) = (\i. (X \ from_nat_into I) i)" + unfolding SUP_def image_comp [symmetric] by (simp add: bij_betw_def) + then have "emeasure M (UNION I X) = emeasure M (\i. X (from_nat_into I i))" + by simp + also have "\ = (\i. emeasure M (X (from_nat_into I i)))" + by (intro suminf_emeasure[symmetric] disj disj2) (auto intro!: sets from_nat_into[OF `I \ {}`]) + also have "\ = (\n. \\<^sup>+i. emeasure M (X i) * indicator {from_nat_into I n} i \count_space I)" + proof (intro arg_cong[where f=suminf] ext) + fix i + have eq: "{a \ I. 0 < emeasure M (X a) * indicator {from_nat_into I i} a} + = (if 0 < emeasure M (X (from_nat_into I i)) then {from_nat_into I i} else {})" + using ereal_0_less_1 + by (auto simp: ereal_zero_less_0_iff indicator_def from_nat_into `I \ {}` simp del: ereal_0_less_1) + have "(\\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I) = + (if 0 < emeasure M (X (from_nat_into I i)) then emeasure M (X (from_nat_into I i)) else 0)" + by (subst positive_integral_count_space) (simp_all add: eq) + also have "\ = emeasure M (X (from_nat_into I i))" + by (simp add: less_le emeasure_nonneg) + finally show "emeasure M (X (from_nat_into I i)) = + \\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I" .. + qed + also have "\ = (\\<^sup>+i. emeasure M (X i) \count_space I)" + apply (subst positive_integral_suminf[symmetric]) + apply (auto simp: emeasure_nonneg intro!: positive_integral_cong) + proof - + fix x assume "x \ I" + then have "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = (\i\{to_nat_on I x}. emeasure M (X x) * indicator {from_nat_into I i} x)" + by (intro suminf_finite) (auto simp: indicator_def I f) + also have "\ = emeasure M (X x)" + by (simp add: I f `x\I`) + finally show "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = emeasure M (X x)" . + qed + finally show ?thesis . +qed + +section {* Measures with Restricted Space *} + +lemma positive_integral_restrict_space: + assumes \: "\ \ sets M" and f: "f \ borel_measurable M" "\x. 0 \ f x" "\x. x \ space M - \ \ f x = 0" + shows "positive_integral (restrict_space M \) f = positive_integral M f" +using f proof (induct rule: borel_measurable_induct) + case (cong f g) then show ?case + using positive_integral_cong[of M f g] positive_integral_cong[of "restrict_space M \" f g] + sets.sets_into_space[OF `\ \ sets M`] + by (simp add: subset_eq space_restrict_space) +next + case (set A) + then have "A \ \" + unfolding indicator_eq_0_iff by (auto dest: sets.sets_into_space) + with set `\ \ sets M` sets.sets_into_space[OF `\ \ sets M`] show ?case + by (subst positive_integral_indicator') + (auto simp add: sets_restrict_space_iff space_restrict_space + emeasure_restrict_space Int_absorb2 + dest: sets.sets_into_space) +next + case (mult f c) then show ?case + by (cases "c = 0") (simp_all add: measurable_restrict_space1 \ positive_integral_cmult) +next + case (add f g) then show ?case + by (simp add: measurable_restrict_space1 \ positive_integral_add ereal_add_nonneg_eq_0_iff) +next + case (seq F) then show ?case + by (auto simp add: SUP_eq_iff measurable_restrict_space1 \ positive_integral_monotone_convergence_SUP) +qed + +section {* Measure spaces with an associated density *} + +definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where + "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" + +lemma + shows sets_density[simp]: "sets (density M f) = sets M" + and space_density[simp]: "space (density M f) = space M" + by (auto simp: density_def) + +(* FIXME: add conversion to simplify space, sets and measurable *) +lemma space_density_imp[measurable_dest]: + "\x M f. x \ space (density M f) \ x \ space M" by auto + +lemma + shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" + and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" + and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" + unfolding measurable_def simple_function_def by simp_all + +lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ + (AE x in M. f x = f' x) \ density M f = density M f'" + unfolding density_def by (auto intro!: measure_of_eq positive_integral_cong_AE sets.space_closed) + +lemma density_max_0: "density M f = density M (\x. max 0 (f x))" +proof - + have "\x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" + by (auto simp: indicator_def) + then show ?thesis + unfolding density_def by (simp add: positive_integral_max_0) +qed + +lemma density_ereal_max_0: "density M (\x. ereal (f x)) = density M (\x. ereal (max 0 (f x)))" + by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) + +lemma emeasure_density: + assumes f[measurable]: "f \ borel_measurable M" and A[measurable]: "A \ sets M" + shows "emeasure (density M f) A = (\\<^sup>+ x. f x * indicator A x \M)" + (is "_ = ?\ A") + unfolding density_def +proof (rule emeasure_measure_of_sigma) + show "sigma_algebra (space M) (sets M)" .. + show "positive (sets M) ?\" + using f by (auto simp: positive_def intro!: positive_integral_positive) + have \_eq: "?\ = (\A. \\<^sup>+ x. max 0 (f x) * indicator A x \M)" (is "?\ = ?\'") + apply (subst positive_integral_max_0[symmetric]) + apply (intro ext positive_integral_cong_AE AE_I2) + apply (auto simp: indicator_def) + done + show "countably_additive (sets M) ?\" + unfolding \_eq + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume "range A \ sets M" + then have "\i. A i \ sets M" by auto + then have *: "\i. (\x. max 0 (f x) * indicator (A i) x) \ borel_measurable M" + by (auto simp: set_eq_iff) + assume disj: "disjoint_family A" + have "(\n. ?\' (A n)) = (\\<^sup>+ x. (\n. max 0 (f x) * indicator (A n) x) \M)" + using f * by (simp add: positive_integral_suminf) + also have "\ = (\\<^sup>+ x. max 0 (f x) * (\n. indicator (A n) x) \M)" using f + by (auto intro!: suminf_cmult_ereal positive_integral_cong_AE) + also have "\ = (\\<^sup>+ x. max 0 (f x) * indicator (\n. A n) x \M)" + unfolding suminf_indicator[OF disj] .. + finally show "(\n. ?\' (A n)) = ?\' (\x. A x)" by simp + qed +qed fact + +lemma null_sets_density_iff: + assumes f: "f \ borel_measurable M" + shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x \ 0)" +proof - + { assume "A \ sets M" + have eq: "(\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. max 0 (f x) * indicator A x \M)" + apply (subst positive_integral_max_0[symmetric]) + apply (intro positive_integral_cong) + apply (auto simp: indicator_def) + done + have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ + emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" + unfolding eq + using f `A \ sets M` + by (intro positive_integral_0_iff) auto + also have "\ \ (AE x in M. max 0 (f x) * indicator A x = 0)" + using f `A \ sets M` + by (intro AE_iff_measurable[OF _ refl, symmetric]) auto + also have "(AE x in M. max 0 (f x) * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" + by (auto simp add: indicator_def max_def split: split_if_asm) + finally have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } + with f show ?thesis + by (simp add: null_sets_def emeasure_density cong: conj_cong) +qed + +lemma AE_density: + assumes f: "f \ borel_measurable M" + shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" +proof + assume "AE x in density M f. P x" + with f obtain N where "{x \ space M. \ P x} \ N" "N \ sets M" and ae: "AE x in M. x \ N \ f x \ 0" + by (auto simp: eventually_ae_filter null_sets_density_iff) + then have "AE x in M. x \ N \ P x" by auto + with ae show "AE x in M. 0 < f x \ P x" + by (rule eventually_elim2) auto +next + fix N assume ae: "AE x in M. 0 < f x \ P x" + then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" + by (auto simp: eventually_ae_filter) + then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. \ 0 < f x}" + "N \ {x\space M. \ 0 < f x} \ sets M" and ae2: "AE x in M. x \ N" + using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in) + show "AE x in density M f. P x" + using ae2 + unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] + by (intro exI[of _ "N \ {x\space M. \ 0 < f x}"] conjI *) + (auto elim: eventually_elim2) +qed + +lemma positive_integral_density': + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g: "g \ borel_measurable M" "\x. 0 \ g x" + shows "integral\<^sup>P (density M f) g = (\\<^sup>+ x. f x * g x \M)" +using g proof induct + case (cong u v) + then show ?case + apply (subst positive_integral_cong[OF cong(3)]) + apply (simp_all cong: positive_integral_cong) + done +next + case (set A) then show ?case + by (simp add: emeasure_density f) +next + case (mult u c) + moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) + ultimately show ?case + using f by (simp add: positive_integral_cmult) +next + case (add u v) + then have "\x. f x * (v x + u x) = f x * v x + f x * u x" + by (simp add: ereal_right_distrib) + with add f show ?case + by (auto simp add: positive_integral_add ereal_zero_le_0_iff intro!: positive_integral_add[symmetric]) +next + case (seq U) + from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" + by eventually_elim (simp add: SUP_ereal_cmult seq) + from seq f show ?case + apply (simp add: positive_integral_monotone_convergence_SUP) + apply (subst positive_integral_cong_AE[OF eq]) + apply (subst positive_integral_monotone_convergence_SUP_AE) + apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) + done +qed + +lemma positive_integral_density: + "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ + integral\<^sup>P (density M f) g' = (\\<^sup>+ x. f x * g' x \M)" + by (subst (1 2) positive_integral_max_0[symmetric]) + (auto intro!: positive_integral_cong_AE + simp: measurable_If max_def ereal_zero_le_0_iff positive_integral_density') + +lemma emeasure_restricted: + assumes S: "S \ sets M" and X: "X \ sets M" + shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" +proof - + have "emeasure (density M (indicator S)) X = (\\<^sup>+x. indicator S x * indicator X x \M)" + using S X by (simp add: emeasure_density) + also have "\ = (\\<^sup>+x. indicator (S \ X) x \M)" + by (auto intro!: positive_integral_cong simp: indicator_def) + also have "\ = emeasure M (S \ X)" + using S X by (simp add: sets.Int) + finally show ?thesis . +qed + +lemma measure_restricted: + "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" + by (simp add: emeasure_restricted measure_def) + +lemma (in finite_measure) finite_measure_restricted: + "S \ sets M \ finite_measure (density M (indicator S))" + by default (simp add: emeasure_restricted) + +lemma emeasure_density_const: + "A \ sets M \ 0 \ c \ emeasure (density M (\_. c)) A = c * emeasure M A" + by (auto simp: positive_integral_cmult_indicator emeasure_density) + +lemma measure_density_const: + "A \ sets M \ 0 < c \ c \ \ \ measure (density M (\_. c)) A = real c * measure M A" + by (auto simp: emeasure_density_const measure_def) + +lemma density_density_eq: + "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. 0 \ f x \ + density (density M f) g = density M (\x. f x * g x)" + by (auto intro!: measure_eqI simp: emeasure_density positive_integral_density ac_simps) + +lemma distr_density_distr: + assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" + and inv: "\x\space M. T' (T x) = x" + assumes f: "f \ borel_measurable M'" + shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") +proof (rule measure_eqI) + fix A assume A: "A \ sets ?R" + { fix x assume "x \ space M" + with sets.sets_into_space[OF A] + have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ereal)" + using T inv by (auto simp: indicator_def measurable_space) } + with A T T' f show "emeasure ?R A = emeasure ?L A" + by (simp add: measurable_comp emeasure_density emeasure_distr + positive_integral_distr measurable_sets cong: positive_integral_cong) +qed simp + +lemma density_density_divide: + fixes f g :: "'a \ real" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + assumes ac: "AE x in M. f x = 0 \ g x = 0" + shows "density (density M f) (\x. g x / f x) = density M g" +proof - + have "density M g = density M (\x. f x * (g x / f x))" + using f g ac by (auto intro!: density_cong measurable_If) + then show ?thesis + using f g by (subst density_density_eq) auto +qed + +section {* Point measure *} + +definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where + "point_measure A f = density (count_space A) f" + +lemma + shows space_point_measure: "space (point_measure A f) = A" + and sets_point_measure: "sets (point_measure A f) = Pow A" + by (auto simp: point_measure_def) + +lemma measurable_point_measure_eq1[simp]: + "g \ measurable (point_measure A f) M \ g \ A \ space M" + unfolding point_measure_def by simp + +lemma measurable_point_measure_eq2_finite[simp]: + "finite A \ + g \ measurable M (point_measure A f) \ + (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" + unfolding point_measure_def by (simp add: measurable_count_space_eq2) + +lemma simple_function_point_measure[simp]: + "simple_function (point_measure A f) g \ finite (g ` A)" + by (simp add: point_measure_def) + +lemma emeasure_point_measure: + assumes A: "finite {a\X. 0 < f a}" "X \ A" + shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" +proof - + have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" + using `X \ A` by auto + with A show ?thesis + by (simp add: emeasure_density positive_integral_count_space ereal_zero_le_0_iff + point_measure_def indicator_def) +qed + +lemma emeasure_point_measure_finite: + "finite A \ (\i. i \ A \ 0 \ f i) \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" + by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) + +lemma emeasure_point_measure_finite2: + "X \ A \ finite X \ (\i. i \ X \ 0 \ f i) \ emeasure (point_measure A f) X = (\a\X. f a)" + by (subst emeasure_point_measure) + (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) + +lemma null_sets_point_measure_iff: + "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x \ 0)" + by (auto simp: AE_count_space null_sets_density_iff point_measure_def) + +lemma AE_point_measure: + "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" + unfolding point_measure_def + by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) + +lemma positive_integral_point_measure: + "finite {a\A. 0 < f a \ 0 < g a} \ + integral\<^sup>P (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" + unfolding point_measure_def + apply (subst density_max_0) + apply (subst positive_integral_density) + apply (simp_all add: AE_count_space positive_integral_density) + apply (subst positive_integral_count_space ) + apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff) + apply (rule finite_subset) + prefer 2 + apply assumption + apply auto + done + +lemma positive_integral_point_measure_finite: + "finite A \ (\a. a \ A \ 0 \ f a) \ (\a. a \ A \ 0 \ g a) \ + integral\<^sup>P (point_measure A f) g = (\a\A. f a * g a)" + by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le) + +section {* Uniform measure *} + +definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" + +lemma + shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M" + and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" + by (auto simp: uniform_measure_def) + +lemma emeasure_uniform_measure[simp]: + assumes A: "A \ sets M" and B: "B \ sets M" + shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" +proof - + from A B have "emeasure (uniform_measure M A) B = (\\<^sup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" + by (auto simp add: uniform_measure_def emeasure_density split: split_indicator + intro!: positive_integral_cong) + also have "\ = emeasure M (A \ B) / emeasure M A" + using A B + by (subst positive_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg) + finally show ?thesis . +qed + +lemma measure_uniform_measure[simp]: + assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" + shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" + using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A + by (cases "emeasure M A" "emeasure M (A \ B)" rule: ereal2_cases) (simp_all add: measure_def) + +section {* Uniform count measure *} + +definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" + +lemma + shows space_uniform_count_measure: "space (uniform_count_measure A) = A" + and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" + unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) + +lemma emeasure_uniform_count_measure: + "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" + by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) + +lemma measure_uniform_count_measure: + "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" + by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) + +end