hoelzl@56993: (* Title: HOL/Probability/Nonnegative_Lebesgue_Integration.thy hoelzl@42067: Author: Johannes Hölzl, TU München hoelzl@42067: Author: Armin Heller, TU München hoelzl@42067: *) hoelzl@38656: hoelzl@56993: header {* Lebesgue Integration for Nonnegative Functions *} hoelzl@35582: hoelzl@56993: theory Nonnegative_Lebesgue_Integration hoelzl@47694: imports Measure_Space Borel_Space hoelzl@35582: begin hoelzl@35582: hoelzl@56949: lemma indicator_less_ereal[simp]: hoelzl@56949: "indicator A x \ (indicator B x::ereal) \ (x \ A \ x \ B)" hoelzl@56949: by (simp add: indicator_def not_le) hoelzl@41981: hoelzl@56994: subsection "Simple function" hoelzl@35582: hoelzl@38656: text {* hoelzl@38656: hoelzl@56996: Our simple functions are not restricted to nonnegative real numbers. Instead hoelzl@38656: they are just functions with a finite range and are measurable when singleton hoelzl@38656: sets are measurable. hoelzl@35582: hoelzl@38656: *} hoelzl@38656: hoelzl@41689: definition "simple_function M g \ hoelzl@38656: finite (g ` space M) \ hoelzl@38656: (\x \ g ` space M. g -` {x} \ space M \ sets M)" hoelzl@36624: hoelzl@47694: lemma simple_functionD: hoelzl@41689: assumes "simple_function M g" hoelzl@40875: shows "finite (g ` space M)" and "g -` X \ space M \ sets M" hoelzl@40871: proof - hoelzl@40871: show "finite (g ` space M)" hoelzl@40871: using assms unfolding simple_function_def by auto hoelzl@40875: have "g -` X \ space M = g -` (X \ g`space M) \ space M" by auto hoelzl@40875: also have "\ = (\x\X \ g`space M. g-`{x} \ space M)" by auto hoelzl@40875: finally show "g -` X \ space M \ sets M" using assms hoelzl@50002: by (auto simp del: UN_simps simp: simple_function_def) hoelzl@40871: qed hoelzl@36624: hoelzl@56949: lemma measurable_simple_function[measurable_dest]: hoelzl@56949: "simple_function M f \ f \ measurable M (count_space UNIV)" hoelzl@56949: unfolding simple_function_def measurable_def hoelzl@56949: proof safe hoelzl@56949: fix A assume "finite (f ` space M)" "\x\f ` space M. f -` {x} \ space M \ sets M" hoelzl@56949: then have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) \ sets M" hoelzl@56949: by (intro sets.finite_UN) auto hoelzl@56949: also have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) = f -` A \ space M" hoelzl@56949: by (auto split: split_if_asm) hoelzl@56949: finally show "f -` A \ space M \ sets M" . hoelzl@56949: qed simp hoelzl@56949: hoelzl@56949: lemma borel_measurable_simple_function: hoelzl@56949: "simple_function M f \ f \ borel_measurable M" hoelzl@56949: by (auto dest!: measurable_simple_function simp: measurable_def) hoelzl@56949: hoelzl@47694: lemma simple_function_measurable2[intro]: hoelzl@41981: assumes "simple_function M f" "simple_function M g" hoelzl@41981: shows "f -` A \ g -` B \ space M \ sets M" hoelzl@41981: proof - hoelzl@41981: have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" hoelzl@41981: by auto hoelzl@41981: then show ?thesis using assms[THEN simple_functionD(2)] by auto hoelzl@41981: qed hoelzl@41981: hoelzl@47694: lemma simple_function_indicator_representation: hoelzl@43920: fixes f ::"'a \ ereal" hoelzl@41689: assumes f: "simple_function M f" and x: "x \ space M" hoelzl@38656: shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" hoelzl@38656: (is "?l = ?r") hoelzl@38656: proof - hoelzl@38705: have "?r = (\y \ f ` space M. hoelzl@38656: (if y = f x then y * indicator (f -` {y} \ space M) x else 0))" hoelzl@38656: by (auto intro!: setsum_cong2) hoelzl@38656: also have "... = f x * indicator (f -` {f x} \ space M) x" hoelzl@38656: using assms by (auto dest: simple_functionD simp: setsum_delta) hoelzl@38656: also have "... = f x" using x by (auto simp: indicator_def) hoelzl@38656: finally show ?thesis by auto hoelzl@38656: qed hoelzl@36624: hoelzl@47694: lemma simple_function_notspace: hoelzl@43920: "simple_function M (\x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h") hoelzl@35692: proof - hoelzl@38656: have "?h ` space M \ {0}" unfolding indicator_def by auto hoelzl@38656: hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) hoelzl@38656: have "?h -` {0} \ space M = space M" by auto hoelzl@38656: thus ?thesis unfolding simple_function_def by auto hoelzl@38656: qed hoelzl@38656: hoelzl@47694: lemma simple_function_cong: hoelzl@38656: assumes "\t. t \ space M \ f t = g t" hoelzl@41689: shows "simple_function M f \ simple_function M g" hoelzl@38656: proof - hoelzl@38656: have "f ` space M = g ` space M" hoelzl@38656: "\x. f -` {x} \ space M = g -` {x} \ space M" hoelzl@38656: using assms by (auto intro!: image_eqI) hoelzl@38656: thus ?thesis unfolding simple_function_def using assms by simp hoelzl@38656: qed hoelzl@38656: hoelzl@47694: lemma simple_function_cong_algebra: hoelzl@41689: assumes "sets N = sets M" "space N = space M" hoelzl@41689: shows "simple_function M f \ simple_function N f" hoelzl@41689: unfolding simple_function_def assms .. hoelzl@41689: hoelzl@47694: lemma simple_function_borel_measurable: hoelzl@41981: fixes f :: "'a \ 'x::{t2_space}" hoelzl@38656: assumes "f \ borel_measurable M" and "finite (f ` space M)" hoelzl@41689: shows "simple_function M f" hoelzl@38656: using assms unfolding simple_function_def hoelzl@38656: by (auto intro: borel_measurable_vimage) hoelzl@38656: hoelzl@56949: lemma simple_function_eq_measurable: hoelzl@43920: fixes f :: "'a \ ereal" hoelzl@56949: shows "simple_function M f \ finite (f`space M) \ f \ measurable M (count_space UNIV)" hoelzl@56949: using simple_function_borel_measurable[of f] measurable_simple_function[of M f] nipkow@44890: by (fastforce simp: simple_function_def) hoelzl@41981: hoelzl@47694: lemma simple_function_const[intro, simp]: hoelzl@41689: "simple_function M (\x. c)" hoelzl@38656: by (auto intro: finite_subset simp: simple_function_def) hoelzl@47694: lemma simple_function_compose[intro, simp]: hoelzl@41689: assumes "simple_function M f" hoelzl@41689: shows "simple_function M (g \ f)" hoelzl@38656: unfolding simple_function_def hoelzl@38656: proof safe hoelzl@38656: show "finite ((g \ f) ` space M)" haftmann@56154: using assms unfolding simple_function_def by (auto simp: image_comp [symmetric]) hoelzl@38656: next hoelzl@38656: fix x assume "x \ space M" hoelzl@38656: let ?G = "g -` {g (f x)} \ (f`space M)" hoelzl@38656: have *: "(g \ f) -` {(g \ f) x} \ space M = hoelzl@38656: (\x\?G. f -` {x} \ space M)" by auto hoelzl@38656: show "(g \ f) -` {(g \ f) x} \ space M \ sets M" hoelzl@38656: using assms unfolding simple_function_def * immler@50244: by (rule_tac sets.finite_UN) auto hoelzl@38656: qed hoelzl@38656: hoelzl@47694: lemma simple_function_indicator[intro, simp]: hoelzl@38656: assumes "A \ sets M" hoelzl@41689: shows "simple_function M (indicator A)" hoelzl@35692: proof - hoelzl@38656: have "indicator A ` space M \ {0, 1}" (is "?S \ _") hoelzl@38656: by (auto simp: indicator_def) hoelzl@38656: hence "finite ?S" by (rule finite_subset) simp hoelzl@38656: moreover have "- A \ space M = space M - A" by auto hoelzl@38656: ultimately show ?thesis unfolding simple_function_def wenzelm@46905: using assms by (auto simp: indicator_def [abs_def]) hoelzl@35692: qed hoelzl@35692: hoelzl@47694: lemma simple_function_Pair[intro, simp]: hoelzl@41689: assumes "simple_function M f" hoelzl@41689: assumes "simple_function M g" hoelzl@41689: shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") hoelzl@38656: unfolding simple_function_def hoelzl@38656: proof safe hoelzl@38656: show "finite (?p ` space M)" hoelzl@38656: using assms unfolding simple_function_def hoelzl@38656: by (rule_tac finite_subset[of _ "f`space M \ g`space M"]) auto hoelzl@38656: next hoelzl@38656: fix x assume "x \ space M" hoelzl@38656: have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = hoelzl@38656: (f -` {f x} \ space M) \ (g -` {g x} \ space M)" hoelzl@38656: by auto hoelzl@38656: with `x \ space M` show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" hoelzl@38656: using assms unfolding simple_function_def by auto hoelzl@38656: qed hoelzl@35692: hoelzl@47694: lemma simple_function_compose1: hoelzl@41689: assumes "simple_function M f" hoelzl@41689: shows "simple_function M (\x. g (f x))" hoelzl@38656: using simple_function_compose[OF assms, of g] hoelzl@38656: by (simp add: comp_def) hoelzl@35582: hoelzl@47694: lemma simple_function_compose2: hoelzl@41689: assumes "simple_function M f" and "simple_function M g" hoelzl@41689: shows "simple_function M (\x. h (f x) (g x))" hoelzl@38656: proof - hoelzl@41689: have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" hoelzl@38656: using assms by auto hoelzl@38656: thus ?thesis by (simp_all add: comp_def) hoelzl@38656: qed hoelzl@35582: hoelzl@47694: lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] hoelzl@38656: and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] hoelzl@38656: and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] hoelzl@38656: and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] hoelzl@38656: and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] hoelzl@38656: and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] hoelzl@41981: and simple_function_max[intro, simp] = simple_function_compose2[where h=max] hoelzl@38656: hoelzl@47694: lemma simple_function_setsum[intro, simp]: hoelzl@41689: assumes "\i. i \ P \ simple_function M (f i)" hoelzl@41689: shows "simple_function M (\x. \i\P. f i x)" hoelzl@38656: proof cases hoelzl@38656: assume "finite P" from this assms show ?thesis by induct auto hoelzl@38656: qed auto hoelzl@35582: hoelzl@56949: lemma simple_function_ereal[intro, simp]: hoelzl@41981: fixes f g :: "'a \ real" assumes sf: "simple_function M f" hoelzl@56949: shows "simple_function M (\x. ereal (f x))" hoelzl@41981: by (auto intro!: simple_function_compose1[OF sf]) hoelzl@41981: hoelzl@56949: lemma simple_function_real_of_nat[intro, simp]: hoelzl@41981: fixes f g :: "'a \ nat" assumes sf: "simple_function M f" hoelzl@56949: shows "simple_function M (\x. real (f x))" hoelzl@41981: by (auto intro!: simple_function_compose1[OF sf]) hoelzl@35582: hoelzl@47694: lemma borel_measurable_implies_simple_function_sequence: hoelzl@43920: fixes u :: "'a \ ereal" hoelzl@38656: assumes u: "u \ borel_measurable M" hoelzl@41981: shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ hoelzl@41981: (\x. (SUP i. f i x) = max 0 (u x)) \ (\i x. 0 \ f i x)" hoelzl@35582: proof - hoelzl@41981: def f \ "\x i. if real i \ u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)" hoelzl@41981: { fix x j have "f x j \ j * 2 ^ j" unfolding f_def hoelzl@41981: proof (split split_if, intro conjI impI) hoelzl@41981: assume "\ real j \ u x" hoelzl@41981: then have "natfloor (real (u x) * 2 ^ j) \ natfloor (j * 2 ^ j)" nipkow@56536: by (cases "u x") (auto intro!: natfloor_mono) hoelzl@41981: moreover have "real (natfloor (j * 2 ^ j)) \ j * 2^j" nipkow@56536: by (intro real_natfloor_le) auto hoelzl@41981: ultimately show "natfloor (real (u x) * 2 ^ j) \ j * 2 ^ j" hoelzl@41981: unfolding real_of_nat_le_iff by auto hoelzl@41981: qed auto } hoelzl@38656: note f_upper = this hoelzl@35582: hoelzl@41981: have real_f: hoelzl@41981: "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))" hoelzl@41981: unfolding f_def by auto hoelzl@35582: wenzelm@46731: let ?g = "\j x. real (f x j) / 2^j :: ereal" hoelzl@41981: show ?thesis hoelzl@41981: proof (intro exI[of _ ?g] conjI allI ballI) hoelzl@41981: fix i hoelzl@41981: have "simple_function M (\x. real (f x i))" hoelzl@41981: proof (intro simple_function_borel_measurable) hoelzl@41981: show "(\x. real (f x i)) \ borel_measurable M" hoelzl@50021: using u by (auto simp: real_f) hoelzl@41981: have "(\x. real (f x i))`space M \ real`{..i*2^i}" hoelzl@41981: using f_upper[of _ i] by auto hoelzl@41981: then show "finite ((\x. real (f x i))`space M)" hoelzl@41981: by (rule finite_subset) auto hoelzl@41981: qed hoelzl@41981: then show "simple_function M (?g i)" hoelzl@43920: by (auto intro: simple_function_ereal simple_function_div) hoelzl@41981: next hoelzl@41981: show "incseq ?g" hoelzl@43920: proof (intro incseq_ereal incseq_SucI le_funI) hoelzl@41981: fix x and i :: nat hoelzl@41981: have "f x i * 2 \ f x (Suc i)" unfolding f_def hoelzl@41981: proof ((split split_if)+, intro conjI impI) hoelzl@43920: assume "ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" hoelzl@41981: then show "i * 2 ^ i * 2 \ natfloor (real (u x) * 2 ^ Suc i)" hoelzl@41981: by (cases "u x") (auto intro!: le_natfloor) hoelzl@38656: next hoelzl@43920: assume "\ ereal (real i) \ u x" "ereal (real (Suc i)) \ u x" hoelzl@41981: then show "natfloor (real (u x) * 2 ^ i) * 2 \ Suc i * 2 ^ Suc i" hoelzl@41981: by (cases "u x") auto hoelzl@41981: next hoelzl@43920: assume "\ ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" hoelzl@41981: have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2" hoelzl@41981: by simp hoelzl@41981: also have "\ \ natfloor (real (u x) * 2 ^ i * 2)" hoelzl@41981: proof cases hoelzl@41981: assume "0 \ u x" then show ?thesis bulwahn@46671: by (intro le_mult_natfloor) hoelzl@41981: next hoelzl@41981: assume "\ 0 \ u x" then show ?thesis hoelzl@41981: by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg) hoelzl@38656: qed hoelzl@41981: also have "\ = natfloor (real (u x) * 2 ^ Suc i)" hoelzl@41981: by (simp add: ac_simps) hoelzl@41981: finally show "natfloor (real (u x) * 2 ^ i) * 2 \ natfloor (real (u x) * 2 ^ Suc i)" . hoelzl@41981: qed simp hoelzl@41981: then show "?g i x \ ?g (Suc i) x" hoelzl@41981: by (auto simp: field_simps) hoelzl@35582: qed hoelzl@38656: next hoelzl@41981: fix x show "(SUP i. ?g i x) = max 0 (u x)" hoelzl@51000: proof (rule SUP_eqI) hoelzl@41981: fix i show "?g i x \ max 0 (u x)" unfolding max_def f_def hoelzl@41981: by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg nipkow@56536: mult_nonpos_nonneg) hoelzl@41981: next hoelzl@41981: fix y assume *: "\i. i \ UNIV \ ?g i x \ y" hoelzl@56571: have "\i. 0 \ ?g i x" by auto hoelzl@41981: from order_trans[OF this *] have "0 \ y" by simp hoelzl@41981: show "max 0 (u x) \ y" hoelzl@41981: proof (cases y) hoelzl@41981: case (real r) hoelzl@41981: with * have *: "\i. f x i \ r * 2^i" by (auto simp: divide_le_eq) huffman@44666: from reals_Archimedean2[of r] * have "u x \ \" by (auto simp: f_def) (metis less_le_not_le) hoelzl@43920: then have "\p. max 0 (u x) = ereal p \ 0 \ p" by (cases "u x") (auto simp: max_def) hoelzl@41981: then guess p .. note ux = this huffman@44666: obtain m :: nat where m: "p < real m" using reals_Archimedean2 .. hoelzl@41981: have "p \ r" hoelzl@41981: proof (rule ccontr) hoelzl@41981: assume "\ p \ r" hoelzl@41981: with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] nipkow@56536: obtain N where "\n\N. r * 2^n < p * 2^n - 1" by (auto simp: field_simps) hoelzl@41981: then have "r * 2^max N m < p * 2^max N m - 1" by simp hoelzl@41981: moreover hoelzl@41981: have "real (natfloor (p * 2 ^ max N m)) \ r * 2 ^ max N m" hoelzl@41981: using *[of "max N m"] m unfolding real_f using ux nipkow@56536: by (cases "0 \ u x") (simp_all add: max_def split: split_if_asm) hoelzl@41981: then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" hoelzl@41981: by (metis real_natfloor_gt_diff_one less_le_trans) hoelzl@41981: ultimately show False by auto hoelzl@38656: qed hoelzl@41981: then show "max 0 (u x) \ y" using real ux by simp hoelzl@41981: qed (insert `0 \ y`, auto) hoelzl@41981: qed hoelzl@56571: qed auto hoelzl@41981: qed hoelzl@35582: hoelzl@47694: lemma borel_measurable_implies_simple_function_sequence': hoelzl@43920: fixes u :: "'a \ ereal" hoelzl@41981: assumes u: "u \ borel_measurable M" hoelzl@41981: obtains f where "\i. simple_function M (f i)" "incseq f" "\i. \ \ range (f i)" hoelzl@41981: "\x. (SUP i. f i x) = max 0 (u x)" "\i x. 0 \ f i x" hoelzl@41981: using borel_measurable_implies_simple_function_sequence[OF u] by auto hoelzl@41981: hoelzl@49796: lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]: hoelzl@49796: fixes u :: "'a \ ereal" hoelzl@49796: assumes u: "simple_function M u" hoelzl@49796: assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" hoelzl@49796: assumes set: "\A. A \ sets M \ P (indicator A)" hoelzl@49796: assumes mult: "\u c. P u \ P (\x. c * u x)" hoelzl@49796: assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" hoelzl@49796: shows "P u" hoelzl@49796: proof (rule cong) hoelzl@49796: from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" hoelzl@49796: proof eventually_elim hoelzl@49796: fix x assume x: "x \ space M" hoelzl@49796: from simple_function_indicator_representation[OF u x] hoelzl@49796: show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. hoelzl@49796: qed hoelzl@49796: next hoelzl@49796: from u have "finite (u ` space M)" hoelzl@49796: unfolding simple_function_def by auto hoelzl@49796: then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" hoelzl@49796: proof induct hoelzl@49796: case empty show ?case hoelzl@49796: using set[of "{}"] by (simp add: indicator_def[abs_def]) hoelzl@49796: qed (auto intro!: add mult set simple_functionD u) hoelzl@49796: next hoelzl@49796: show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" hoelzl@49796: apply (subst simple_function_cong) hoelzl@49796: apply (rule simple_function_indicator_representation[symmetric]) hoelzl@49796: apply (auto intro: u) hoelzl@49796: done hoelzl@49796: qed fact hoelzl@49796: hoelzl@49796: lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: hoelzl@49796: fixes u :: "'a \ ereal" hoelzl@49799: assumes u: "simple_function M u" and nn: "\x. 0 \ u x" hoelzl@49799: assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" hoelzl@49796: assumes set: "\A. A \ sets M \ P (indicator A)" hoelzl@49797: assumes mult: "\u c. 0 \ c \ simple_function M u \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" hoelzl@56993: 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)" hoelzl@49796: shows "P u" hoelzl@49796: proof - hoelzl@49796: show ?thesis hoelzl@49796: proof (rule cong) hoelzl@49799: fix x assume x: "x \ space M" hoelzl@49799: from simple_function_indicator_representation[OF u x] hoelzl@49799: show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. hoelzl@49796: next hoelzl@49799: show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" hoelzl@49796: apply (subst simple_function_cong) hoelzl@49796: apply (rule simple_function_indicator_representation[symmetric]) hoelzl@49799: apply (auto intro: u) hoelzl@49796: done hoelzl@49796: next hoelzl@56993: hoelzl@49799: from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" hoelzl@49796: unfolding simple_function_def by auto hoelzl@49799: then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" hoelzl@49796: proof induct hoelzl@49796: case empty show ?case hoelzl@49796: using set[of "{}"] by (simp add: indicator_def[abs_def]) hoelzl@56993: next hoelzl@56993: case (insert x S) hoelzl@56993: { fix z have "(\y\S. y * indicator (u -` {y} \ space M) z) = 0 \ hoelzl@56993: x * indicator (u -` {x} \ space M) z = 0" hoelzl@56993: using insert by (subst setsum_ereal_0) (auto simp: indicator_def) } hoelzl@56993: note disj = this hoelzl@56993: from insert show ?case hoelzl@56993: by (auto intro!: add mult set simple_functionD u setsum_nonneg simple_function_setsum disj) hoelzl@56993: qed hoelzl@49796: qed fact hoelzl@49796: qed hoelzl@49796: hoelzl@49796: lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: hoelzl@49796: fixes u :: "'a \ ereal" hoelzl@49799: assumes u: "u \ borel_measurable M" "\x. 0 \ u x" hoelzl@49799: assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" hoelzl@49796: assumes set: "\A. A \ sets M \ P (indicator A)" hoelzl@56993: 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)" hoelzl@56993: 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)" hoelzl@56993: 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)" hoelzl@49796: shows "P u" hoelzl@49796: using u hoelzl@49796: proof (induct rule: borel_measurable_implies_simple_function_sequence') hoelzl@49797: fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and hoelzl@49796: sup: "\x. (SUP i. U i x) = max 0 (u x)" and nn: "\i x. 0 \ U i x" hoelzl@49799: have u_eq: "u = (SUP i. U i)" hoelzl@49796: using nn u sup by (auto simp: max_def) hoelzl@56993: hoelzl@56993: have not_inf: "\x i. x \ space M \ U i x < \" hoelzl@56993: using U by (auto simp: image_iff eq_commute) hoelzl@49796: hoelzl@49797: from U have "\i. U i \ borel_measurable M" hoelzl@49797: by (simp add: borel_measurable_simple_function) hoelzl@49797: hoelzl@49799: show "P u" hoelzl@49796: unfolding u_eq hoelzl@49796: proof (rule seq) hoelzl@49796: fix i show "P (U i)" hoelzl@56993: using `simple_function M (U i)` nn[of i] not_inf[of _ i] hoelzl@56993: proof (induct rule: simple_function_induct_nn) hoelzl@56993: case (mult u c) hoelzl@56993: show ?case hoelzl@56993: proof cases hoelzl@56993: assume "c = 0 \ space M = {} \ (\x\space M. u x = 0)" hoelzl@56993: with mult(2) show ?thesis hoelzl@56993: by (intro cong[of "\x. c * u x" "indicator {}"] set) hoelzl@56993: (auto dest!: borel_measurable_simple_function) hoelzl@56993: next hoelzl@56993: assume "\ (c = 0 \ space M = {} \ (\x\space M. u x = 0))" hoelzl@56993: with mult obtain x where u_fin: "\x. x \ space M \ u x < \" hoelzl@56993: and x: "x \ space M" "u x \ 0" "c \ 0" hoelzl@56993: by auto hoelzl@56993: with mult have "P u" hoelzl@56993: by auto hoelzl@56993: from x mult(5)[OF `x \ space M`] mult(1) mult(3)[of x] have "c < \" hoelzl@56993: by auto hoelzl@56993: with u_fin mult hoelzl@56993: show ?thesis hoelzl@56993: by (intro mult') (auto dest!: borel_measurable_simple_function) hoelzl@56993: qed hoelzl@56993: qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) hoelzl@49797: qed fact+ hoelzl@49796: qed hoelzl@49796: hoelzl@47694: lemma simple_function_If_set: hoelzl@41981: assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" hoelzl@41981: shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") hoelzl@41981: proof - hoelzl@41981: def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" hoelzl@41981: show ?thesis unfolding simple_function_def hoelzl@41981: proof safe hoelzl@41981: have "?IF ` space M \ f ` space M \ g ` space M" by auto hoelzl@41981: from finite_subset[OF this] assms hoelzl@41981: show "finite (?IF ` space M)" unfolding simple_function_def by auto hoelzl@41981: next hoelzl@41981: fix x assume "x \ space M" hoelzl@41981: then have *: "?IF -` {?IF x} \ space M = (if x \ A hoelzl@41981: then ((F (f x) \ (A \ space M)) \ (G (f x) - (G (f x) \ (A \ space M)))) hoelzl@41981: else ((F (g x) \ (A \ space M)) \ (G (g x) - (G (g x) \ (A \ space M)))))" immler@50244: using sets.sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) hoelzl@41981: have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" hoelzl@41981: unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto hoelzl@41981: show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto hoelzl@35582: qed hoelzl@35582: qed hoelzl@35582: hoelzl@47694: lemma simple_function_If: hoelzl@41981: assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" hoelzl@41981: shows "simple_function M (\x. if P x then f x else g x)" hoelzl@35582: proof - hoelzl@41981: have "{x\space M. P x} = {x. P x} \ space M" by auto hoelzl@41981: with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp hoelzl@38656: qed hoelzl@38656: hoelzl@47694: lemma simple_function_subalgebra: hoelzl@41689: assumes "simple_function N f" hoelzl@41689: and N_subalgebra: "sets N \ sets M" "space N = space M" hoelzl@41689: shows "simple_function M f" hoelzl@41689: using assms unfolding simple_function_def by auto hoelzl@39092: hoelzl@47694: lemma simple_function_comp: hoelzl@47694: assumes T: "T \ measurable M M'" hoelzl@41689: and f: "simple_function M' f" hoelzl@41689: shows "simple_function M (\x. f (T x))" hoelzl@41661: proof (intro simple_function_def[THEN iffD2] conjI ballI) hoelzl@41661: have "(\x. f (T x)) ` space M \ f ` space M'" hoelzl@41661: using T unfolding measurable_def by auto hoelzl@41661: then show "finite ((\x. f (T x)) ` space M)" hoelzl@41689: using f unfolding simple_function_def by (auto intro: finite_subset) hoelzl@41661: fix i assume i: "i \ (\x. f (T x)) ` space M" hoelzl@41661: then have "i \ f ` space M'" hoelzl@41661: using T unfolding measurable_def by auto hoelzl@41661: then have "f -` {i} \ space M' \ sets M'" hoelzl@41689: using f unfolding simple_function_def by auto hoelzl@41661: then have "T -` (f -` {i} \ space M') \ space M \ sets M" hoelzl@41661: using T unfolding measurable_def by auto hoelzl@41661: also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" hoelzl@41661: using T unfolding measurable_def by auto hoelzl@41661: finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . hoelzl@40859: qed hoelzl@40859: hoelzl@56994: subsection "Simple integral" hoelzl@38656: wenzelm@53015: definition simple_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>S") where wenzelm@53015: "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" hoelzl@41689: hoelzl@41689: syntax wenzelm@53015: "_simple_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>S _. _ \_" [60,61] 110) hoelzl@41689: hoelzl@41689: translations wenzelm@53015: "\\<^sup>S x. f \M" == "CONST simple_integral M (%x. f)" hoelzl@35582: hoelzl@47694: lemma simple_integral_cong: hoelzl@38656: assumes "\t. t \ space M \ f t = g t" wenzelm@53015: shows "integral\<^sup>S M f = integral\<^sup>S M g" hoelzl@38656: proof - hoelzl@38656: have "f ` space M = g ` space M" hoelzl@38656: "\x. f -` {x} \ space M = g -` {x} \ space M" hoelzl@38656: using assms by (auto intro!: image_eqI) hoelzl@38656: thus ?thesis unfolding simple_integral_def by simp hoelzl@38656: qed hoelzl@38656: hoelzl@47694: lemma simple_integral_const[simp]: wenzelm@53015: "(\\<^sup>Sx. c \M) = c * (emeasure M) (space M)" hoelzl@38656: proof (cases "space M = {}") hoelzl@38656: case True thus ?thesis unfolding simple_integral_def by simp hoelzl@38656: next hoelzl@38656: case False hence "(\x. c) ` space M = {c}" by auto hoelzl@38656: thus ?thesis unfolding simple_integral_def by simp hoelzl@35582: qed hoelzl@35582: hoelzl@47694: lemma simple_function_partition: hoelzl@41981: assumes f: "simple_function M f" and g: "simple_function M g" hoelzl@56949: assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" hoelzl@56949: assumes v: "\x. x \ space M \ f x = v (g x)" hoelzl@56949: shows "integral\<^sup>S M f = (\y\g ` space M. v y * emeasure M {x\space M. g x = y})" hoelzl@56949: (is "_ = ?r") hoelzl@56949: proof - hoelzl@56949: from f g have [simp]: "finite (f`space M)" "finite (g`space M)" hoelzl@56949: by (auto simp: simple_function_def) hoelzl@56949: from f g have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" hoelzl@56949: by (auto intro: measurable_simple_function) hoelzl@35582: hoelzl@56949: { fix y assume "y \ space M" hoelzl@56949: then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" hoelzl@56949: by (auto cong: sub simp: v[symmetric]) } hoelzl@56949: note eq = this hoelzl@35582: hoelzl@56949: have "integral\<^sup>S M f = hoelzl@56949: (\y\f`space M. y * (\z\g`space M. hoelzl@56949: if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" hoelzl@56949: unfolding simple_integral_def hoelzl@56949: proof (safe intro!: setsum_cong ereal_left_mult_cong) hoelzl@56949: fix y assume y: "y \ space M" "f y \ 0" hoelzl@56949: have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = hoelzl@56949: {z. \x\space M. f y = f x \ z = g x}" hoelzl@56949: by auto hoelzl@56949: have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = hoelzl@56949: f -` {f y} \ space M" hoelzl@56949: by (auto simp: eq_commute cong: sub rev_conj_cong) hoelzl@56949: have "finite (g`space M)" by simp hoelzl@56949: then have "finite {z. \x\space M. f y = f x \ z = g x}" hoelzl@56949: by (rule rev_finite_subset) auto hoelzl@56949: then show "emeasure M (f -` {f y} \ space M) = hoelzl@56949: (\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)" hoelzl@56949: apply (simp add: setsum_cases) hoelzl@56949: apply (subst setsum_emeasure) hoelzl@56949: apply (auto simp: disjoint_family_on_def eq) hoelzl@56949: done hoelzl@38656: qed hoelzl@56949: also have "\ = (\y\f`space M. (\z\g`space M. hoelzl@56949: if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" hoelzl@56949: by (auto intro!: setsum_cong simp: setsum_ereal_right_distrib emeasure_nonneg) hoelzl@56949: also have "\ = ?r" hoelzl@56949: by (subst setsum_commute) hoelzl@56949: (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq) hoelzl@56949: finally show "integral\<^sup>S M f = ?r" . hoelzl@35582: qed hoelzl@35582: hoelzl@47694: lemma simple_integral_add[simp]: hoelzl@41981: assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" wenzelm@53015: shows "(\\<^sup>Sx. f x + g x \M) = integral\<^sup>S M f + integral\<^sup>S M g" hoelzl@35582: proof - hoelzl@56949: have "(\\<^sup>Sx. f x + g x \M) = hoelzl@56949: (\y\(\x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\space M. (f x, g x) = y})" hoelzl@56949: by (intro simple_function_partition) (auto intro: f g) hoelzl@56949: also have "\ = (\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) + hoelzl@56949: (\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y})" hoelzl@56949: using assms(2,4) by (auto intro!: setsum_cong ereal_left_distrib simp: setsum_addf[symmetric]) hoelzl@56949: 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)" hoelzl@56949: by (intro simple_function_partition[symmetric]) (auto intro: f g) hoelzl@56949: 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)" hoelzl@56949: by (intro simple_function_partition[symmetric]) (auto intro: f g) hoelzl@56949: finally show ?thesis . hoelzl@35582: qed hoelzl@35582: hoelzl@47694: lemma simple_integral_setsum[simp]: hoelzl@41981: assumes "\i x. i \ P \ 0 \ f i x" hoelzl@41689: assumes "\i. i \ P \ simple_function M (f i)" wenzelm@53015: shows "(\\<^sup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^sup>S M (f i))" hoelzl@38656: proof cases hoelzl@38656: assume "finite P" hoelzl@38656: from this assms show ?thesis hoelzl@41981: by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) hoelzl@38656: qed auto hoelzl@38656: hoelzl@47694: lemma simple_integral_mult[simp]: hoelzl@41981: assumes f: "simple_function M f" "\x. 0 \ f x" wenzelm@53015: shows "(\\<^sup>Sx. c * f x \M) = c * integral\<^sup>S M f" hoelzl@38656: proof - hoelzl@56949: have "(\\<^sup>Sx. c * f x \M) = (\y\f ` space M. (c * y) * emeasure M {x\space M. f x = y})" hoelzl@56949: using f by (intro simple_function_partition) auto hoelzl@56949: also have "\ = c * integral\<^sup>S M f" hoelzl@56949: using f unfolding simple_integral_def hoelzl@56949: by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult_assoc Int_def conj_commute) hoelzl@56949: finally show ?thesis . hoelzl@40871: qed hoelzl@40871: hoelzl@47694: lemma simple_integral_mono_AE: hoelzl@56949: assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" hoelzl@47694: and mono: "AE x in M. f x \ g x" wenzelm@53015: shows "integral\<^sup>S M f \ integral\<^sup>S M g" hoelzl@40859: proof - hoelzl@56949: let ?\ = "\P. emeasure M {x\space M. P x}" hoelzl@56949: have "integral\<^sup>S M f = (\y\(\x. (f x, g x))`space M. fst y * ?\ (\x. (f x, g x) = y))" hoelzl@56949: using f g by (intro simple_function_partition) auto hoelzl@56949: also have "\ \ (\y\(\x. (f x, g x))`space M. snd y * ?\ (\x. (f x, g x) = y))" hoelzl@56949: proof (clarsimp intro!: setsum_mono) hoelzl@40859: fix x assume "x \ space M" hoelzl@56949: let ?M = "?\ (\y. f y = f x \ g y = g x)" hoelzl@56949: show "f x * ?M \ g x * ?M" hoelzl@56949: proof cases hoelzl@56949: assume "?M \ 0" hoelzl@56949: then have "0 < ?M" hoelzl@56949: by (simp add: less_le emeasure_nonneg) hoelzl@56949: also have "\ \ ?\ (\y. f x \ g x)" hoelzl@56949: using mono by (intro emeasure_mono_AE) auto hoelzl@56949: finally have "\ \ f x \ g x" hoelzl@56949: by (intro notI) auto hoelzl@56949: then show ?thesis hoelzl@56949: by (intro ereal_mult_right_mono) auto hoelzl@56949: qed simp hoelzl@40859: qed hoelzl@56949: also have "\ = integral\<^sup>S M g" hoelzl@56949: using f g by (intro simple_function_partition[symmetric]) auto hoelzl@56949: finally show ?thesis . hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma simple_integral_mono: hoelzl@41689: assumes "simple_function M f" and "simple_function M g" hoelzl@38656: and mono: "\ x. x \ space M \ f x \ g x" wenzelm@53015: shows "integral\<^sup>S M f \ integral\<^sup>S M g" hoelzl@41705: using assms by (intro simple_integral_mono_AE) auto hoelzl@35582: hoelzl@47694: lemma simple_integral_cong_AE: hoelzl@41981: assumes "simple_function M f" and "simple_function M g" hoelzl@47694: and "AE x in M. f x = g x" wenzelm@53015: shows "integral\<^sup>S M f = integral\<^sup>S M g" hoelzl@40859: using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) hoelzl@40859: hoelzl@47694: lemma simple_integral_cong': hoelzl@41689: assumes sf: "simple_function M f" "simple_function M g" hoelzl@47694: and mea: "(emeasure M) {x\space M. f x \ g x} = 0" wenzelm@53015: shows "integral\<^sup>S M f = integral\<^sup>S M g" hoelzl@40859: proof (intro simple_integral_cong_AE sf AE_I) hoelzl@47694: show "(emeasure M) {x\space M. f x \ g x} = 0" by fact hoelzl@40859: show "{x \ space M. f x \ g x} \ sets M" hoelzl@40859: using sf[THEN borel_measurable_simple_function] by auto hoelzl@40859: qed simp hoelzl@40859: hoelzl@47694: lemma simple_integral_indicator: hoelzl@56949: assumes A: "A \ sets M" hoelzl@49796: assumes f: "simple_function M f" wenzelm@53015: shows "(\\<^sup>Sx. f x * indicator A x \M) = hoelzl@56949: (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" hoelzl@56949: proof - hoelzl@56949: have eq: "(\x. (f x, indicator A x)) ` space M \ {x. snd x = 1} = (\x. (f x, 1::ereal))`A" hoelzl@56949: using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: split_if_asm) hoelzl@56949: have eq2: "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" hoelzl@56949: by (auto simp: image_iff) hoelzl@56949: hoelzl@56949: have "(\\<^sup>Sx. f x * indicator A x \M) = hoelzl@56949: (\y\(\x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\space M. (f x, indicator A x) = y})" hoelzl@56949: using assms by (intro simple_function_partition) auto hoelzl@56949: also have "\ = (\y\(\x. (f x, indicator A x::ereal))`space M. hoelzl@56949: if snd y = 1 then fst y * emeasure M (f -` {fst y} \ space M \ A) else 0)" hoelzl@56949: by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum_cong) hoelzl@56949: also have "\ = (\y\(\x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \ space M \ A))" hoelzl@56949: using assms by (subst setsum_cases) (auto intro!: simple_functionD(1) simp: eq) hoelzl@56949: also have "\ = (\y\fst`(\x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \ space M \ A))" hoelzl@56949: by (subst setsum_reindex[where f=fst]) (auto simp: inj_on_def) hoelzl@56949: also have "\ = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" hoelzl@56949: using A[THEN sets.sets_into_space] hoelzl@56949: by (intro setsum_mono_zero_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) hoelzl@56949: finally show ?thesis . hoelzl@38656: qed hoelzl@35582: hoelzl@47694: lemma simple_integral_indicator_only[simp]: hoelzl@38656: assumes "A \ sets M" wenzelm@53015: shows "integral\<^sup>S M (indicator A) = emeasure M A" hoelzl@56949: using simple_integral_indicator[OF assms, of "\x. 1"] sets.sets_into_space[OF assms] hoelzl@56949: by (simp_all add: image_constant_conv Int_absorb1 split: split_if_asm) hoelzl@35582: hoelzl@47694: lemma simple_integral_null_set: hoelzl@47694: assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" wenzelm@53015: shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" hoelzl@38656: proof - hoelzl@47694: have "AE x in M. indicator N x = (0 :: ereal)" hoelzl@47694: using `N \ null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) wenzelm@53015: then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" hoelzl@41981: using assms apply (intro simple_integral_cong_AE) by auto hoelzl@40859: then show ?thesis by simp hoelzl@38656: qed hoelzl@35582: hoelzl@47694: lemma simple_integral_cong_AE_mult_indicator: hoelzl@47694: assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" wenzelm@53015: shows "integral\<^sup>S M f = (\\<^sup>Sx. f x * indicator S x \M)" hoelzl@41705: using assms by (intro simple_integral_cong_AE) auto hoelzl@35582: hoelzl@47694: lemma simple_integral_cmult_indicator: hoelzl@41981: assumes A: "A \ sets M" hoelzl@56949: shows "(\\<^sup>Sx. c * indicator A x \M) = c * emeasure M A" hoelzl@41981: using simple_integral_mult[OF simple_function_indicator[OF A]] hoelzl@41981: unfolding simple_integral_indicator_only[OF A] by simp hoelzl@41981: hoelzl@56996: lemma simple_integral_nonneg: hoelzl@47694: assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" wenzelm@53015: shows "0 \ integral\<^sup>S M f" hoelzl@41981: proof - wenzelm@53015: have "integral\<^sup>S M (\x. 0) \ integral\<^sup>S M f" hoelzl@41981: using simple_integral_mono_AE[OF _ f ae] by auto hoelzl@41981: then show ?thesis by simp hoelzl@41981: qed hoelzl@41981: hoelzl@56994: subsection {* Integral on nonnegative functions *} hoelzl@41689: hoelzl@56996: definition nn_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>N") where hoelzl@56996: "integral\<^sup>N M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^sup>S M g)" hoelzl@35692: hoelzl@41689: syntax hoelzl@56996: "_nn_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>+ _. _ \_" [60,61] 110) hoelzl@41689: hoelzl@41689: translations hoelzl@56996: "\\<^sup>+x. f \M" == "CONST nn_integral M (\x. f)" hoelzl@40872: hoelzl@56996: lemma nn_integral_nonneg: hoelzl@56996: "0 \ integral\<^sup>N M f" hoelzl@56996: by (auto intro!: SUP_upper2[of "\x. 0"] simp: nn_integral_def le_fun_def) hoelzl@40873: hoelzl@56996: lemma nn_integral_not_MInfty[simp]: "integral\<^sup>N M f \ -\" hoelzl@56996: using nn_integral_nonneg[of M f] by auto hoelzl@47694: hoelzl@56996: lemma nn_integral_def_finite: hoelzl@56996: "integral\<^sup>N M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f \ range g \ {0 ..< \}}. integral\<^sup>S M g)" haftmann@56218: (is "_ = SUPREMUM ?A ?f") hoelzl@56996: unfolding nn_integral_def hoelzl@44928: proof (safe intro!: antisym SUP_least) hoelzl@41981: fix g assume g: "simple_function M g" "g \ max 0 \ f" hoelzl@41981: let ?G = "{x \ space M. \ g x \ \}" hoelzl@41981: note gM = g(1)[THEN borel_measurable_simple_function] wenzelm@50252: have \_G_pos: "0 \ (emeasure M) ?G" using gM by auto wenzelm@46731: let ?g = "\y x. if g x = \ then y else max 0 (g x)" hoelzl@41981: from g gM have g_in_A: "\y. 0 \ y \ y \ \ \ ?g y \ ?A" hoelzl@41981: apply (safe intro!: simple_function_max simple_function_If) hoelzl@41981: apply (force simp: max_def le_fun_def split: split_if_asm)+ hoelzl@41981: done haftmann@56218: show "integral\<^sup>S M g \ SUPREMUM ?A ?f" hoelzl@41981: proof cases hoelzl@41981: have g0: "?g 0 \ ?A" by (intro g_in_A) auto hoelzl@47694: assume "(emeasure M) ?G = 0" hoelzl@47694: with gM have "AE x in M. x \ ?G" hoelzl@47694: by (auto simp add: AE_iff_null intro!: null_setsI) hoelzl@41981: with gM g show ?thesis hoelzl@44928: by (intro SUP_upper2[OF g0] simple_integral_mono_AE) hoelzl@41981: (auto simp: max_def intro!: simple_function_If) hoelzl@41981: next wenzelm@50252: assume \_G: "(emeasure M) ?G \ 0" haftmann@56218: have "SUPREMUM ?A (integral\<^sup>S M) = \" hoelzl@41981: proof (intro SUP_PInfty) hoelzl@41981: fix n :: nat hoelzl@47694: let ?y = "ereal (real n) / (if (emeasure M) ?G = \ then 1 else (emeasure M) ?G)" wenzelm@50252: have "0 \ ?y" "?y \ \" using \_G \_G_pos by (auto simp: ereal_divide_eq) hoelzl@41981: then have "?g ?y \ ?A" by (rule g_in_A) hoelzl@47694: have "real n \ ?y * (emeasure M) ?G" wenzelm@50252: using \_G \_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) wenzelm@53015: also have "\ = (\\<^sup>Sx. ?y * indicator ?G x \M)" hoelzl@41981: using `0 \ ?y` `?g ?y \ ?A` gM hoelzl@41981: by (subst simple_integral_cmult_indicator) auto wenzelm@53015: also have "\ \ integral\<^sup>S M (?g ?y)" using `?g ?y \ ?A` gM hoelzl@41981: by (intro simple_integral_mono) auto wenzelm@53015: finally show "\i\?A. real n \ integral\<^sup>S M i" hoelzl@41981: using `?g ?y \ ?A` by blast hoelzl@41981: qed hoelzl@41981: then show ?thesis by simp hoelzl@41981: qed hoelzl@44928: qed (auto intro: SUP_upper) hoelzl@40873: hoelzl@56996: lemma nn_integral_mono_AE: hoelzl@56996: assumes ae: "AE x in M. u x \ v x" shows "integral\<^sup>N M u \ integral\<^sup>N M v" hoelzl@56996: unfolding nn_integral_def hoelzl@41981: proof (safe intro!: SUP_mono) hoelzl@41981: fix n assume n: "simple_function M n" "n \ max 0 \ u" hoelzl@41981: from ae[THEN AE_E] guess N . note N = this hoelzl@47694: then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) wenzelm@46731: let ?n = "\x. n x * indicator (space M - N) x" hoelzl@47694: have "AE x in M. n x \ ?n x" "simple_function M ?n" hoelzl@41981: using n N ae_N by auto hoelzl@41981: moreover hoelzl@41981: { fix x have "?n x \ max 0 (v x)" hoelzl@41981: proof cases hoelzl@41981: assume x: "x \ space M - N" hoelzl@41981: with N have "u x \ v x" by auto hoelzl@41981: with n(2)[THEN le_funD, of x] x show ?thesis hoelzl@41981: by (auto simp: max_def split: split_if_asm) hoelzl@41981: qed simp } hoelzl@41981: then have "?n \ max 0 \ v" by (auto simp: le_funI) wenzelm@53015: moreover have "integral\<^sup>S M n \ integral\<^sup>S M ?n" hoelzl@41981: using ae_N N n by (auto intro!: simple_integral_mono_AE) wenzelm@53015: ultimately show "\m\{g. simple_function M g \ g \ max 0 \ v}. integral\<^sup>S M n \ integral\<^sup>S M m" hoelzl@41981: by force hoelzl@38656: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_mono: hoelzl@56996: "(\x. x \ space M \ u x \ v x) \ integral\<^sup>N M u \ integral\<^sup>N M v" hoelzl@56996: by (auto intro: nn_integral_mono_AE) hoelzl@40859: hoelzl@56996: lemma nn_integral_cong_AE: hoelzl@56996: "AE x in M. u x = v x \ integral\<^sup>N M u = integral\<^sup>N M v" hoelzl@56996: by (auto simp: eq_iff intro!: nn_integral_mono_AE) hoelzl@40859: hoelzl@56996: lemma nn_integral_cong: hoelzl@56996: "(\x. x \ space M \ u x = v x) \ integral\<^sup>N M u = integral\<^sup>N M v" hoelzl@56996: by (auto intro: nn_integral_cong_AE) hoelzl@40859: hoelzl@56996: lemma nn_integral_cong_strong: hoelzl@56996: "M = N \ (\x. x \ space M \ u x = v x) \ integral\<^sup>N M u = integral\<^sup>N N v" hoelzl@56996: by (auto intro: nn_integral_cong) hoelzl@56993: hoelzl@56996: lemma nn_integral_eq_simple_integral: hoelzl@56996: assumes f: "simple_function M f" "\x. 0 \ f x" shows "integral\<^sup>N M f = integral\<^sup>S M f" hoelzl@41981: proof - wenzelm@46731: let ?f = "\x. f x * indicator (space M) x" hoelzl@41981: have f': "simple_function M ?f" using f by auto hoelzl@41981: with f(2) have [simp]: "max 0 \ ?f = ?f" hoelzl@41981: by (auto simp: fun_eq_iff max_def split: split_indicator) hoelzl@56996: have "integral\<^sup>N M ?f \ integral\<^sup>S M ?f" using f' hoelzl@56996: by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def) hoelzl@56996: moreover have "integral\<^sup>S M ?f \ integral\<^sup>N M ?f" hoelzl@56996: unfolding nn_integral_def hoelzl@44928: using f' by (auto intro!: SUP_upper) hoelzl@41981: ultimately show ?thesis hoelzl@56996: by (simp cong: nn_integral_cong simple_integral_cong) hoelzl@41981: qed hoelzl@41981: hoelzl@56996: lemma nn_integral_eq_simple_integral_AE: hoelzl@56996: assumes f: "simple_function M f" "AE x in M. 0 \ f x" shows "integral\<^sup>N M f = integral\<^sup>S M f" hoelzl@41981: proof - hoelzl@47694: have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) hoelzl@56996: with f have "integral\<^sup>N M f = integral\<^sup>S M (\x. max 0 (f x))" hoelzl@56996: by (simp cong: nn_integral_cong_AE simple_integral_cong_AE hoelzl@56996: add: nn_integral_eq_simple_integral) hoelzl@41981: with assms show ?thesis hoelzl@41981: by (auto intro!: simple_integral_cong_AE split: split_max) hoelzl@41981: qed hoelzl@40873: hoelzl@56996: lemma nn_integral_SUP_approx: hoelzl@41981: assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" hoelzl@41981: and u: "simple_function M u" "u \ (SUP i. f i)" "u`space M \ {0..<\}" hoelzl@56996: shows "integral\<^sup>S M u \ (SUP i. integral\<^sup>N M (f i))" (is "_ \ ?S") hoelzl@43920: proof (rule ereal_le_mult_one_interval) hoelzl@56996: have "0 \ (SUP i. integral\<^sup>N M (f i))" hoelzl@56996: using f(3) by (auto intro!: SUP_upper2 nn_integral_nonneg) hoelzl@56996: then show "(SUP i. integral\<^sup>N M (f i)) \ -\" by auto hoelzl@41981: have u_range: "\x. x \ space M \ 0 \ u x \ u x \ \" hoelzl@41981: using u(3) by auto hoelzl@43920: fix a :: ereal assume "0 < a" "a < 1" hoelzl@38656: hence "a \ 0" by auto wenzelm@46731: let ?B = "\i. {x \ space M. a * u x \ f i x}" hoelzl@38656: have B: "\i. ?B i \ sets M" hoelzl@56949: using f `simple_function M u`[THEN borel_measurable_simple_function] by auto hoelzl@38656: wenzelm@46731: let ?uB = "\i x. u x * indicator (?B i) x" hoelzl@38656: hoelzl@38656: { fix i have "?B i \ ?B (Suc i)" hoelzl@38656: proof safe hoelzl@38656: fix i x assume "a * u x \ f i x" hoelzl@38656: also have "\ \ f (Suc i) x" hoelzl@41981: using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto hoelzl@38656: finally show "a * u x \ f (Suc i) x" . hoelzl@38656: qed } hoelzl@38656: note B_mono = this hoelzl@35582: immler@50244: note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B] hoelzl@38656: wenzelm@46731: let ?B' = "\i n. (u -` {i} \ space M) \ ?B n" hoelzl@47694: have measure_conv: "\i. (emeasure M) (u -` {i} \ space M) = (SUP n. (emeasure M) (?B' i n))" hoelzl@41981: proof - hoelzl@41981: fix i hoelzl@41981: have 1: "range (?B' i) \ sets M" using B_u by auto hoelzl@41981: have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI) hoelzl@41981: have "(\n. ?B' i n) = u -` {i} \ space M" hoelzl@41981: proof safe hoelzl@41981: fix x i assume x: "x \ space M" hoelzl@41981: show "x \ (\i. ?B' (u x) i)" hoelzl@41981: proof cases hoelzl@41981: assume "u x = 0" thus ?thesis using `x \ space M` f(3) by simp hoelzl@41981: next hoelzl@41981: assume "u x \ 0" hoelzl@41981: with `a < 1` u_range[OF `x \ space M`] hoelzl@41981: have "a * u x < 1 * u x" hoelzl@43920: by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) noschinl@46884: also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) hoelzl@44928: finally obtain i where "a * u x < f i x" unfolding SUP_def haftmann@56166: by (auto simp add: less_SUP_iff) hoelzl@41981: hence "a * u x \ f i x" by auto hoelzl@41981: thus ?thesis using `x \ space M` by auto hoelzl@41981: qed hoelzl@40859: qed hoelzl@47694: then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp hoelzl@41981: qed hoelzl@38656: wenzelm@53015: have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" hoelzl@41689: unfolding simple_integral_indicator[OF B `simple_function M u`] haftmann@56212: proof (subst SUP_ereal_setsum, safe) hoelzl@38656: fix x n assume "x \ space M" hoelzl@47694: with u_range show "incseq (\i. u x * (emeasure M) (?B' (u x) i))" "\i. 0 \ u x * (emeasure M) (?B' (u x) i)" hoelzl@47694: using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) hoelzl@38656: next wenzelm@53015: show "integral\<^sup>S M u = (\i\u ` space M. SUP n. i * (emeasure M) (?B' i n))" hoelzl@41981: using measure_conv u_range B_u unfolding simple_integral_def haftmann@56212: by (auto intro!: setsum_cong SUP_ereal_cmult [symmetric]) hoelzl@38656: qed hoelzl@38656: moreover wenzelm@53015: have "a * (SUP i. integral\<^sup>S M (?uB i)) \ ?S" haftmann@56212: apply (subst SUP_ereal_cmult [symmetric]) hoelzl@38705: proof (safe intro!: SUP_mono bexI) hoelzl@38656: fix i wenzelm@53015: have "a * integral\<^sup>S M (?uB i) = (\\<^sup>Sx. a * ?uB i x \M)" hoelzl@41981: using B `simple_function M u` u_range hoelzl@41981: by (subst simple_integral_mult) (auto split: split_indicator) hoelzl@56996: also have "\ \ integral\<^sup>N M (f i)" hoelzl@38656: proof - hoelzl@41981: have *: "simple_function M (\x. a * ?uB i x)" using B `0 < a` u(1) by auto hoelzl@41981: show ?thesis using f(3) * u_range `0 < a` hoelzl@56996: by (subst nn_integral_eq_simple_integral[symmetric]) hoelzl@56996: (auto intro!: nn_integral_mono split: split_indicator) hoelzl@38656: qed hoelzl@56996: finally show "a * integral\<^sup>S M (?uB i) \ integral\<^sup>N M (f i)" hoelzl@38656: by auto hoelzl@41981: next wenzelm@53015: fix i show "0 \ \\<^sup>S x. ?uB i x \M" using B `0 < a` u(1) u_range hoelzl@56996: by (intro simple_integral_nonneg) (auto split: split_indicator) hoelzl@41981: qed (insert `0 < a`, auto) wenzelm@53015: ultimately show "a * integral\<^sup>S M u \ ?S" by simp hoelzl@35582: qed hoelzl@35582: hoelzl@56996: lemma incseq_nn_integral: hoelzl@56996: assumes "incseq f" shows "incseq (\i. integral\<^sup>N M (f i))" hoelzl@41981: proof - hoelzl@41981: have "\i x. f i x \ f (Suc i) x" hoelzl@41981: using assms by (auto dest!: incseq_SucD simp: le_fun_def) hoelzl@41981: then show ?thesis hoelzl@56996: by (auto intro!: incseq_SucI nn_integral_mono) hoelzl@41981: qed hoelzl@41981: hoelzl@35582: text {* Beppo-Levi monotone convergence theorem *} hoelzl@56996: lemma nn_integral_monotone_convergence_SUP: hoelzl@41981: assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" hoelzl@56996: shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" hoelzl@41981: proof (rule antisym) hoelzl@56996: show "(SUP j. integral\<^sup>N M (f j)) \ (\\<^sup>+ x. (SUP i. f i x) \M)" hoelzl@56996: by (auto intro!: SUP_least SUP_upper nn_integral_mono) hoelzl@38656: next hoelzl@56996: show "(\\<^sup>+ x. (SUP i. f i x) \M) \ (SUP j. integral\<^sup>N M (f j))" hoelzl@56996: unfolding nn_integral_def_finite[of _ "\x. SUP i. f i x"] hoelzl@44928: proof (safe intro!: SUP_least) hoelzl@41981: fix g assume g: "simple_function M g" wenzelm@53374: and *: "g \ max 0 \ (\x. SUP i. f i x)" "range g \ {0..<\}" wenzelm@53374: then have "\x. 0 \ (SUP i. f i x)" and g': "g`space M \ {0..<\}" hoelzl@44928: using f by (auto intro!: SUP_upper2) hoelzl@56996: with * show "integral\<^sup>S M g \ (SUP j. integral\<^sup>N M (f j))" hoelzl@56996: by (intro nn_integral_SUP_approx[OF f g _ g']) noschinl@46884: (auto simp: le_fun_def max_def) hoelzl@35582: qed hoelzl@35582: qed hoelzl@35582: hoelzl@56996: lemma nn_integral_monotone_convergence_SUP_AE: hoelzl@47694: assumes f: "\i. AE x in M. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" hoelzl@56996: shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" hoelzl@40859: proof - hoelzl@47694: from f have "AE x in M. \i. f i x \ f (Suc i) x \ 0 \ f i x" hoelzl@41981: by (simp add: AE_all_countable) hoelzl@41981: from this[THEN AE_E] guess N . note N = this wenzelm@46731: let ?f = "\i x. if x \ space M - N then f i x else 0" hoelzl@47694: have f_eq: "AE x in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) wenzelm@53015: then have "(\\<^sup>+ x. (SUP i. f i x) \M) = (\\<^sup>+ x. (SUP i. ?f i x) \M)" hoelzl@56996: by (auto intro!: nn_integral_cong_AE) wenzelm@53015: also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \M))" hoelzl@56996: proof (rule nn_integral_monotone_convergence_SUP) hoelzl@41981: show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) hoelzl@41981: { fix i show "(\x. if x \ space M - N then f i x else 0) \ borel_measurable M" hoelzl@41981: using f N(3) by (intro measurable_If_set) auto hoelzl@41981: fix x show "0 \ ?f i x" hoelzl@41981: using N(1) by auto } hoelzl@40859: qed wenzelm@53015: also have "\ = (SUP i. (\\<^sup>+ x. f i x \M))" hoelzl@56996: using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] nn_integral_cong_AE ext) hoelzl@41981: finally show ?thesis . hoelzl@41981: qed hoelzl@41981: hoelzl@56996: lemma nn_integral_monotone_convergence_SUP_AE_incseq: hoelzl@47694: assumes f: "incseq f" "\i. AE x in M. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" hoelzl@56996: shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>N M (f i))" hoelzl@41981: using f[unfolded incseq_Suc_iff le_fun_def] hoelzl@56996: by (intro nn_integral_monotone_convergence_SUP_AE[OF _ borel]) hoelzl@41981: auto hoelzl@41981: hoelzl@56996: lemma nn_integral_monotone_convergence_simple: hoelzl@41981: assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" wenzelm@53015: shows "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" hoelzl@56996: using assms unfolding nn_integral_monotone_convergence_SUP[OF f(1) hoelzl@41981: f(3)[THEN borel_measurable_simple_function] f(2)] hoelzl@56996: by (auto intro!: nn_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPREMUM UNIV"] ext) hoelzl@41981: hoelzl@56996: lemma nn_integral_max_0: hoelzl@56996: "(\\<^sup>+x. max 0 (f x) \M) = integral\<^sup>N M f" hoelzl@56996: by (simp add: le_fun_def nn_integral_def) hoelzl@41981: hoelzl@56996: lemma nn_integral_cong_pos: hoelzl@41981: assumes "\x. x \ space M \ f x \ 0 \ g x \ 0 \ f x = g x" hoelzl@56996: shows "integral\<^sup>N M f = integral\<^sup>N M g" hoelzl@41981: proof - hoelzl@56996: have "integral\<^sup>N M (\x. max 0 (f x)) = integral\<^sup>N M (\x. max 0 (g x))" hoelzl@56996: proof (intro nn_integral_cong) hoelzl@41981: fix x assume "x \ space M" hoelzl@41981: from assms[OF this] show "max 0 (f x) = max 0 (g x)" hoelzl@41981: by (auto split: split_max) hoelzl@41981: qed hoelzl@56996: then show ?thesis by (simp add: nn_integral_max_0) hoelzl@40859: qed hoelzl@40859: hoelzl@47694: lemma SUP_simple_integral_sequences: hoelzl@41981: assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" hoelzl@41981: and g: "incseq g" "\i x. 0 \ g i x" "\i. simple_function M (g i)" hoelzl@47694: and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" wenzelm@53015: shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" haftmann@56218: (is "SUPREMUM _ ?F = SUPREMUM _ ?G") hoelzl@38656: proof - wenzelm@53015: have "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" hoelzl@56996: using f by (rule nn_integral_monotone_convergence_simple) wenzelm@53015: also have "\ = (\\<^sup>+x. (SUP i. g i x) \M)" hoelzl@56996: unfolding eq[THEN nn_integral_cong_AE] .. hoelzl@38656: also have "\ = (SUP i. ?G i)" hoelzl@56996: using g by (rule nn_integral_monotone_convergence_simple[symmetric]) hoelzl@41981: finally show ?thesis by simp hoelzl@38656: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_const[simp]: wenzelm@53015: "0 \ c \ (\\<^sup>+ x. c \M) = c * (emeasure M) (space M)" hoelzl@56996: by (subst nn_integral_eq_simple_integral) auto hoelzl@38656: hoelzl@56996: lemma nn_integral_linear: hoelzl@41981: assumes f: "f \ borel_measurable M" "\x. 0 \ f x" and "0 \ a" hoelzl@41981: and g: "g \ borel_measurable M" "\x. 0 \ g x" hoelzl@56996: shows "(\\<^sup>+ x. a * f x + g x \M) = a * integral\<^sup>N M f + integral\<^sup>N M g" hoelzl@56996: (is "integral\<^sup>N M ?L = _") hoelzl@35582: proof - hoelzl@41981: from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . hoelzl@56996: note u = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this hoelzl@41981: from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . hoelzl@56996: note v = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this wenzelm@46731: let ?L' = "\i x. a * u i x + v i x" hoelzl@38656: hoelzl@41981: have "?L \ borel_measurable M" using assms by auto hoelzl@38656: from borel_measurable_implies_simple_function_sequence'[OF this] guess l . hoelzl@56996: note l = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this hoelzl@41981: wenzelm@53015: have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" hoelzl@41981: using u v `0 \ a` hoelzl@41981: by (auto simp: incseq_Suc_iff le_fun_def hoelzl@43920: intro!: add_mono ereal_mult_left_mono simple_integral_mono) wenzelm@53015: 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)" hoelzl@56996: using u v `0 \ a` by (auto simp: simple_integral_nonneg) wenzelm@53015: { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \ -\" "integral\<^sup>S M (v i) \ -\" hoelzl@41981: by (auto split: split_if_asm) } hoelzl@41981: note not_MInf = this hoelzl@41981: wenzelm@53015: have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" hoelzl@41981: proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) hoelzl@41981: show "incseq ?L'" "\i x. 0 \ ?L' i x" "\i. simple_function M (?L' i)" hoelzl@41981: using u v `0 \ a` unfolding incseq_Suc_iff le_fun_def nipkow@56537: by (auto intro!: add_mono ereal_mult_left_mono) hoelzl@41981: { fix x hoelzl@41981: { 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] hoelzl@41981: by auto } hoelzl@41981: then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" hoelzl@41981: using `0 \ a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] haftmann@56212: by (subst SUP_ereal_cmult [symmetric, OF u(6) `0 \ a`]) haftmann@56212: (auto intro!: SUP_ereal_add nipkow@56537: simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } hoelzl@47694: then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" hoelzl@41981: unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) nipkow@56537: by (intro AE_I2) (auto split: split_max) hoelzl@38656: qed wenzelm@53015: also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" haftmann@56218: using u(2, 6) v(2, 6) `0 \ a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) wenzelm@53015: 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)" hoelzl@41981: unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] hoelzl@41981: unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] haftmann@56212: apply (subst SUP_ereal_cmult [symmetric, OF pos(1) `0 \ a`]) haftmann@56212: apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) . hoelzl@56996: then show ?thesis by (simp add: nn_integral_max_0) hoelzl@38656: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_cmult: hoelzl@49775: assumes f: "f \ borel_measurable M" "0 \ c" hoelzl@56996: shows "(\\<^sup>+ x. c * f x \M) = c * integral\<^sup>N M f" hoelzl@41981: proof - hoelzl@41981: have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` hoelzl@43920: by (auto split: split_max simp: ereal_zero_le_0_iff) wenzelm@53015: have "(\\<^sup>+ x. c * f x \M) = (\\<^sup>+ x. c * max 0 (f x) \M)" hoelzl@56996: by (simp add: nn_integral_max_0) hoelzl@41981: then show ?thesis hoelzl@56996: using nn_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" _ "\x. 0"] f hoelzl@56996: by (auto simp: nn_integral_max_0) hoelzl@41981: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_multc: hoelzl@49775: assumes "f \ borel_measurable M" "0 \ c" hoelzl@56996: shows "(\\<^sup>+ x. f x * c \M) = integral\<^sup>N M f * c" hoelzl@56996: unfolding mult_commute[of _ c] nn_integral_cmult[OF assms] by simp hoelzl@41096: hoelzl@56996: lemma nn_integral_indicator[simp]: wenzelm@53015: "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" hoelzl@56996: by (subst nn_integral_eq_simple_integral) hoelzl@49775: (auto simp: simple_integral_indicator) hoelzl@38656: hoelzl@56996: lemma nn_integral_cmult_indicator: wenzelm@53015: "0 \ c \ A \ sets M \ (\\<^sup>+ x. c * indicator A x \M) = c * (emeasure M) A" hoelzl@56996: by (subst nn_integral_eq_simple_integral) hoelzl@41544: (auto simp: simple_function_indicator simple_integral_indicator) hoelzl@38656: hoelzl@56996: lemma nn_integral_indicator': hoelzl@50097: assumes [measurable]: "A \ space M \ sets M" wenzelm@53015: shows "(\\<^sup>+ x. indicator A x \M) = emeasure M (A \ space M)" hoelzl@50097: proof - wenzelm@53015: have "(\\<^sup>+ x. indicator A x \M) = (\\<^sup>+ x. indicator (A \ space M) x \M)" hoelzl@56996: by (intro nn_integral_cong) (simp split: split_indicator) hoelzl@50097: also have "\ = emeasure M (A \ space M)" hoelzl@50097: by simp hoelzl@50097: finally show ?thesis . hoelzl@50097: qed hoelzl@50097: hoelzl@56996: lemma nn_integral_add: hoelzl@47694: assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" hoelzl@47694: and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" hoelzl@56996: shows "(\\<^sup>+ x. f x + g x \M) = integral\<^sup>N M f + integral\<^sup>N M g" hoelzl@41981: proof - hoelzl@47694: have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" nipkow@56537: using assms by (auto split: split_max) wenzelm@53015: have "(\\<^sup>+ x. f x + g x \M) = (\\<^sup>+ x. max 0 (f x + g x) \M)" hoelzl@56996: by (simp add: nn_integral_max_0) wenzelm@53015: also have "\ = (\\<^sup>+ x. max 0 (f x) + max 0 (g x) \M)" hoelzl@56996: unfolding ae[THEN nn_integral_cong_AE] .. wenzelm@53015: also have "\ = (\\<^sup>+ x. max 0 (f x) \M) + (\\<^sup>+ x. max 0 (g x) \M)" hoelzl@56996: using nn_integral_linear[of "\x. max 0 (f x)" _ 1 "\x. max 0 (g x)"] f g hoelzl@41981: by auto hoelzl@41981: finally show ?thesis hoelzl@56996: by (simp add: nn_integral_max_0) hoelzl@41981: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_setsum: hoelzl@47694: assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x in M. 0 \ f i x" hoelzl@56996: shows "(\\<^sup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^sup>N M (f i))" hoelzl@38656: proof cases hoelzl@41981: assume f: "finite P" hoelzl@47694: from assms have "AE x in M. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto hoelzl@41981: from f this assms(1) show ?thesis hoelzl@38656: proof induct hoelzl@38656: case (insert i P) hoelzl@47694: then have "f i \ borel_measurable M" "AE x in M. 0 \ f i x" hoelzl@47694: "(\x. \i\P. f i x) \ borel_measurable M" "AE x in M. 0 \ (\i\P. f i x)" hoelzl@50002: by (auto intro!: setsum_nonneg) hoelzl@56996: from nn_integral_add[OF this] hoelzl@38656: show ?case using insert by auto hoelzl@38656: qed simp hoelzl@38656: qed simp hoelzl@38656: hoelzl@56996: lemma nn_integral_Markov_inequality: hoelzl@49775: assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" and "A \ sets M" and c: "0 \ c" wenzelm@53015: shows "(emeasure M) ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^sup>+ x. u x * indicator A x \M)" hoelzl@47694: (is "(emeasure M) ?A \ _ * ?PI") hoelzl@41981: proof - hoelzl@41981: have "?A \ sets M" hoelzl@41981: using `A \ sets M` u by auto wenzelm@53015: hence "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" hoelzl@56996: using nn_integral_indicator by simp wenzelm@53015: also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u c hoelzl@56996: by (auto intro!: nn_integral_mono_AE hoelzl@43920: simp: indicator_def ereal_zero_le_0_iff) wenzelm@53015: also have "\ = c * (\\<^sup>+ x. u x * indicator A x \M)" hoelzl@41981: using assms hoelzl@56996: by (auto intro!: nn_integral_cmult simp: ereal_zero_le_0_iff) hoelzl@41981: finally show ?thesis . hoelzl@41981: qed hoelzl@41981: hoelzl@56996: lemma nn_integral_noteq_infinite: hoelzl@47694: assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" hoelzl@56996: and "integral\<^sup>N M g \ \" hoelzl@47694: shows "AE x in M. g x \ \" hoelzl@41981: proof (rule ccontr) hoelzl@47694: assume c: "\ (AE x in M. g x \ \)" hoelzl@47694: have "(emeasure M) {x\space M. g x = \} \ 0" hoelzl@47694: using c g by (auto simp add: AE_iff_null) hoelzl@47694: moreover have "0 \ (emeasure M) {x\space M. g x = \}" using g by (auto intro: measurable_sets) hoelzl@47694: ultimately have "0 < (emeasure M) {x\space M. g x = \}" by auto hoelzl@47694: then have "\ = \ * (emeasure M) {x\space M. g x = \}" by auto wenzelm@53015: also have "\ \ (\\<^sup>+x. \ * indicator {x\space M. g x = \} x \M)" hoelzl@56996: using g by (subst nn_integral_cmult_indicator) auto hoelzl@56996: also have "\ \ integral\<^sup>N M g" hoelzl@56996: using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) hoelzl@56996: finally show False using `integral\<^sup>N M g \ \` by auto hoelzl@41981: qed hoelzl@41981: hoelzl@56996: lemma nn_integral_PInf: hoelzl@56949: assumes f: "f \ borel_measurable M" hoelzl@56996: and not_Inf: "integral\<^sup>N M f \ \" hoelzl@56949: shows "(emeasure M) (f -` {\} \ space M) = 0" hoelzl@56949: proof - hoelzl@56949: have "\ * (emeasure M) (f -` {\} \ space M) = (\\<^sup>+ x. \ * indicator (f -` {\} \ space M) x \M)" hoelzl@56996: using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets) hoelzl@56996: also have "\ \ integral\<^sup>N M (\x. max 0 (f x))" hoelzl@56996: by (auto intro!: nn_integral_mono simp: indicator_def max_def) hoelzl@56996: finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^sup>N M f" hoelzl@56996: by (simp add: nn_integral_max_0) hoelzl@56949: moreover have "0 \ (emeasure M) (f -` {\} \ space M)" hoelzl@56949: by (rule emeasure_nonneg) hoelzl@56949: ultimately show ?thesis hoelzl@56949: using assms by (auto split: split_if_asm) hoelzl@56949: qed hoelzl@56949: hoelzl@56996: lemma nn_integral_PInf_AE: hoelzl@56996: assumes "f \ borel_measurable M" "integral\<^sup>N M f \ \" shows "AE x in M. f x \ \" hoelzl@56949: proof (rule AE_I) hoelzl@56949: show "(emeasure M) (f -` {\} \ space M) = 0" hoelzl@56996: by (rule nn_integral_PInf[OF assms]) hoelzl@56949: show "f -` {\} \ space M \ sets M" hoelzl@56949: using assms by (auto intro: borel_measurable_vimage) hoelzl@56949: qed auto hoelzl@56949: hoelzl@56949: lemma simple_integral_PInf: hoelzl@56949: assumes "simple_function M f" "\x. 0 \ f x" hoelzl@56949: and "integral\<^sup>S M f \ \" hoelzl@56949: shows "(emeasure M) (f -` {\} \ space M) = 0" hoelzl@56996: proof (rule nn_integral_PInf) hoelzl@56949: show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) hoelzl@56996: show "integral\<^sup>N M f \ \" hoelzl@56996: using assms by (simp add: nn_integral_eq_simple_integral) hoelzl@56949: qed hoelzl@56949: hoelzl@56996: lemma nn_integral_diff: hoelzl@41981: assumes f: "f \ borel_measurable M" hoelzl@47694: and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" hoelzl@56996: and fin: "integral\<^sup>N M g \ \" hoelzl@47694: and mono: "AE x in M. g x \ f x" hoelzl@56996: shows "(\\<^sup>+ x. f x - g x \M) = integral\<^sup>N M f - integral\<^sup>N M g" hoelzl@38656: proof - hoelzl@47694: have diff: "(\x. f x - g x) \ borel_measurable M" "AE x in M. 0 \ f x - g x" hoelzl@43920: using assms by (auto intro: ereal_diff_positive) hoelzl@47694: have pos_f: "AE x in M. 0 \ f x" using mono g by auto hoelzl@43920: { fix a b :: ereal assume "0 \ a" "a \ \" "0 \ b" "a \ b" then have "b - a + a = b" hoelzl@43920: by (cases rule: ereal2_cases[of a b]) auto } hoelzl@41981: note * = this hoelzl@47694: then have "AE x in M. f x = f x - g x + g x" hoelzl@56996: using mono nn_integral_noteq_infinite[OF g fin] assms by auto hoelzl@56996: then have **: "integral\<^sup>N M f = (\\<^sup>+x. f x - g x \M) + integral\<^sup>N M g" hoelzl@56996: unfolding nn_integral_add[OF diff g, symmetric] hoelzl@56996: by (rule nn_integral_cong_AE) hoelzl@41981: show ?thesis unfolding ** hoelzl@56996: using fin nn_integral_nonneg[of M g] hoelzl@56996: by (cases rule: ereal2_cases[of "\\<^sup>+ x. f x - g x \M" "integral\<^sup>N M g"]) auto hoelzl@38656: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_suminf: hoelzl@47694: assumes f: "\i. f i \ borel_measurable M" "\i. AE x in M. 0 \ f i x" hoelzl@56996: shows "(\\<^sup>+ x. (\i. f i x) \M) = (\i. integral\<^sup>N M (f i))" hoelzl@38656: proof - hoelzl@47694: have all_pos: "AE x in M. \i. 0 \ f i x" hoelzl@41981: using assms by (auto simp: AE_all_countable) hoelzl@56996: have "(\i. integral\<^sup>N M (f i)) = (SUP n. \iN M (f i))" hoelzl@56996: using nn_integral_nonneg by (rule suminf_ereal_eq_SUP) wenzelm@53015: also have "\ = (SUP n. \\<^sup>+x. (\iM)" hoelzl@56996: unfolding nn_integral_setsum[OF f] .. wenzelm@53015: also have "\ = \\<^sup>+x. (SUP n. \iM" using f all_pos hoelzl@56996: by (intro nn_integral_monotone_convergence_SUP_AE[symmetric]) hoelzl@41981: (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3) wenzelm@53015: also have "\ = \\<^sup>+x. (\i. f i x) \M" using all_pos hoelzl@56996: by (intro nn_integral_cong_AE) (auto simp: suminf_ereal_eq_SUP) hoelzl@41981: finally show ?thesis by simp hoelzl@38656: qed hoelzl@38656: hoelzl@56996: lemma nn_integral_mult_bounded_inf: hoelzl@56993: assumes f: "f \ borel_measurable M" "(\\<^sup>+x. f x \M) < \" hoelzl@56993: and c: "0 \ c" "c \ \" and ae: "AE x in M. g x \ c * f x" hoelzl@56993: shows "(\\<^sup>+x. g x \M) < \" hoelzl@56993: proof - hoelzl@56993: have "(\\<^sup>+x. g x \M) \ (\\<^sup>+x. c * f x \M)" hoelzl@56996: by (intro nn_integral_mono_AE ae) hoelzl@56993: also have "(\\<^sup>+x. c * f x \M) < \" hoelzl@56996: using c f by (subst nn_integral_cmult) auto hoelzl@56993: finally show ?thesis . hoelzl@56993: qed hoelzl@56993: hoelzl@38656: text {* Fatou's lemma: convergence theorem on limes inferior *} hoelzl@56993: hoelzl@56996: lemma nn_integral_liminf: hoelzl@43920: fixes u :: "nat \ 'a \ ereal" hoelzl@47694: assumes u: "\i. u i \ borel_measurable M" "\i. AE x in M. 0 \ u i x" hoelzl@56996: shows "(\\<^sup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>N M (u n))" hoelzl@38656: proof - hoelzl@47694: have pos: "AE x in M. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) wenzelm@53015: have "(\\<^sup>+ x. liminf (\n. u n x) \M) = wenzelm@53015: (SUP n. \\<^sup>+ x. (INF i:{n..}. u i x) \M)" haftmann@56212: unfolding liminf_SUP_INF using pos u hoelzl@56996: by (intro nn_integral_monotone_convergence_SUP_AE) hoelzl@44937: (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) hoelzl@56996: also have "\ \ liminf (\n. integral\<^sup>N M (u n))" haftmann@56212: unfolding liminf_SUP_INF hoelzl@56996: by (auto intro!: SUP_mono exI INF_greatest nn_integral_mono INF_lower) hoelzl@38656: finally show ?thesis . hoelzl@35582: qed hoelzl@35582: hoelzl@56993: lemma le_Limsup: hoelzl@56993: "F \ bot \ eventually (\x. c \ g x) F \ c \ Limsup F g" hoelzl@56993: using Limsup_mono[of "\_. c" g F] by (simp add: Limsup_const) hoelzl@56993: hoelzl@56993: lemma Limsup_le: hoelzl@56993: "F \ bot \ eventually (\x. f x \ c) F \ Limsup F f \ c" hoelzl@56993: using Limsup_mono[of f "\_. c" F] by (simp add: Limsup_const) hoelzl@56993: hoelzl@56993: lemma ereal_mono_minus_cancel: hoelzl@56993: fixes a b c :: ereal hoelzl@56993: shows "c - a \ c - b \ 0 \ c \ c < \ \ b \ a" hoelzl@56993: by (cases a b c rule: ereal3_cases) auto hoelzl@56993: hoelzl@56996: lemma nn_integral_limsup: hoelzl@56993: fixes u :: "nat \ 'a \ ereal" hoelzl@56993: assumes [measurable]: "\i. u i \ borel_measurable M" "w \ borel_measurable M" hoelzl@56993: 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) < \" hoelzl@56996: shows "limsup (\n. integral\<^sup>N M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" hoelzl@56993: proof - hoelzl@56993: have bnd: "AE x in M. \i. 0 \ u i x \ u i x \ w x" hoelzl@56993: using bounds by (auto simp: AE_all_countable) hoelzl@56993: hoelzl@56993: from bounds[of 0] have w_nonneg: "AE x in M. 0 \ w x" hoelzl@56993: by auto hoelzl@56993: hoelzl@56993: have "(\\<^sup>+x. w x \M) - (\\<^sup>+x. limsup (\n. u n x) \M) = (\\<^sup>+x. w x - limsup (\n. u n x) \M)" hoelzl@56996: proof (intro nn_integral_diff[symmetric]) hoelzl@56993: show "AE x in M. 0 \ limsup (\n. u n x)" hoelzl@56993: using bnd by (auto intro!: le_Limsup) hoelzl@56993: show "AE x in M. limsup (\n. u n x) \ w x" hoelzl@56993: using bnd by (auto intro!: Limsup_le) hoelzl@56993: then have "(\\<^sup>+x. limsup (\n. u n x) \M) < \" hoelzl@56996: by (intro nn_integral_mult_bounded_inf[OF _ w, of 1]) auto hoelzl@56993: then show "(\\<^sup>+x. limsup (\n. u n x) \M) \ \" hoelzl@56993: by simp hoelzl@56993: qed auto hoelzl@56993: also have "\ = (\\<^sup>+x. liminf (\n. w x - u n x) \M)" hoelzl@56993: using w_nonneg hoelzl@56996: by (intro nn_integral_cong_AE, eventually_elim) hoelzl@56993: (auto intro!: liminf_ereal_cminus[symmetric]) hoelzl@56993: also have "\ \ liminf (\n. \\<^sup>+x. w x - u n x \M)" hoelzl@56996: proof (rule nn_integral_liminf) hoelzl@56993: fix i show "AE x in M. 0 \ w x - u i x" hoelzl@56993: using bounds[of i] by eventually_elim (auto intro: ereal_diff_positive) hoelzl@56993: qed simp hoelzl@56993: also have "(\n. \\<^sup>+x. w x - u n x \M) = (\n. (\\<^sup>+x. w x \M) - (\\<^sup>+x. u n x \M))" hoelzl@56996: proof (intro ext nn_integral_diff) hoelzl@56993: fix i have "(\\<^sup>+x. u i x \M) < \" hoelzl@56996: using bounds by (intro nn_integral_mult_bounded_inf[OF _ w, of 1]) auto hoelzl@56993: then show "(\\<^sup>+x. u i x \M) \ \" by simp hoelzl@56993: qed (insert bounds, auto) hoelzl@56993: 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)" hoelzl@56993: using w by (intro liminf_ereal_cminus) auto hoelzl@56993: finally show ?thesis hoelzl@56996: by (rule ereal_mono_minus_cancel) (intro w nn_integral_nonneg)+ hoelzl@56993: qed hoelzl@56993: hoelzl@56996: lemma nn_integral_dominated_convergence: hoelzl@56993: assumes [measurable]: hoelzl@56993: "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" hoelzl@56993: and bound: "\j. AE x in M. 0 \ u j x" "\j. AE x in M. u j x \ w x" hoelzl@56993: and w: "(\\<^sup>+x. w x \M) < \" hoelzl@56993: and u': "AE x in M. (\i. u i x) ----> u' x" hoelzl@56993: shows "(\i. (\\<^sup>+x. u i x \M)) ----> (\\<^sup>+x. u' x \M)" hoelzl@56993: proof - hoelzl@56996: have "limsup (\n. integral\<^sup>N M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" hoelzl@56996: by (intro nn_integral_limsup[OF _ _ bound w]) auto hoelzl@56993: moreover have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" hoelzl@56996: using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) hoelzl@56993: moreover have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" hoelzl@56996: using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) hoelzl@56996: moreover have "(\\<^sup>+x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>N M (u n))" hoelzl@56996: by (intro nn_integral_liminf[OF _ bound(1)]) auto hoelzl@56996: moreover have "liminf (\n. integral\<^sup>N M (u n)) \ limsup (\n. integral\<^sup>N M (u n))" hoelzl@56993: by (intro Liminf_le_Limsup sequentially_bot) hoelzl@56993: ultimately show ?thesis hoelzl@56993: by (intro Liminf_eq_Limsup) auto hoelzl@56993: qed hoelzl@56993: hoelzl@56996: lemma nn_integral_null_set: wenzelm@53015: assumes "N \ null_sets M" shows "(\\<^sup>+ x. u x * indicator N x \M) = 0" hoelzl@38656: proof - wenzelm@53015: have "(\\<^sup>+ x. u x * indicator N x \M) = (\\<^sup>+ x. 0 \M)" hoelzl@56996: proof (intro nn_integral_cong_AE AE_I) hoelzl@40859: show "{x \ space M. u x * indicator N x \ 0} \ N" hoelzl@40859: by (auto simp: indicator_def) hoelzl@47694: show "(emeasure M) N = 0" "N \ sets M" hoelzl@40859: using assms by auto hoelzl@35582: qed hoelzl@40859: then show ?thesis by simp hoelzl@38656: qed hoelzl@35582: hoelzl@56996: lemma nn_integral_0_iff: hoelzl@47694: assumes u: "u \ borel_measurable M" and pos: "AE x in M. 0 \ u x" hoelzl@56996: shows "integral\<^sup>N M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" hoelzl@47694: (is "_ \ (emeasure M) ?A = 0") hoelzl@35582: proof - hoelzl@56996: have u_eq: "(\\<^sup>+ x. u x * indicator ?A x \M) = integral\<^sup>N M u" hoelzl@56996: by (auto intro!: nn_integral_cong simp: indicator_def) hoelzl@38656: show ?thesis hoelzl@38656: proof hoelzl@47694: assume "(emeasure M) ?A = 0" hoelzl@56996: with nn_integral_null_set[of ?A M u] u hoelzl@56996: show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def) hoelzl@38656: next hoelzl@43920: { fix r :: ereal and n :: nat assume gt_1: "1 \ real n * r" hoelzl@43920: then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) hoelzl@43920: then have "0 \ r" by (auto simp add: ereal_zero_less_0_iff) } hoelzl@41981: note gt_1 = this hoelzl@56996: assume *: "integral\<^sup>N M u = 0" wenzelm@46731: let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" hoelzl@47694: have "0 = (SUP n. (emeasure M) (?M n \ ?A))" hoelzl@38656: proof - hoelzl@41981: { fix n :: nat hoelzl@56996: from nn_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] hoelzl@47694: have "(emeasure M) (?M n \ ?A) \ 0" unfolding u_eq * using u by simp hoelzl@47694: moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto hoelzl@47694: ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } hoelzl@38656: thus ?thesis by simp hoelzl@35582: qed hoelzl@47694: also have "\ = (emeasure M) (\n. ?M n \ ?A)" hoelzl@47694: proof (safe intro!: SUP_emeasure_incseq) hoelzl@38656: fix n show "?M n \ ?A \ sets M" immler@50244: using u by (auto intro!: sets.Int) hoelzl@38656: next hoelzl@41981: show "incseq (\n. {x \ space M. 1 \ real n * u x} \ {x \ space M. u x \ 0})" hoelzl@41981: proof (safe intro!: incseq_SucI) hoelzl@41981: fix n :: nat and x hoelzl@41981: assume *: "1 \ real n * u x" wenzelm@53374: also from gt_1[OF *] have "real n * u x \ real (Suc n) * u x" hoelzl@43920: using `0 \ u x` by (auto intro!: ereal_mult_right_mono) hoelzl@41981: finally show "1 \ real (Suc n) * u x" by auto hoelzl@41981: qed hoelzl@38656: qed hoelzl@47694: also have "\ = (emeasure M) {x\space M. 0 < u x}" hoelzl@47694: proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1) hoelzl@41981: fix x assume "0 < u x" and [simp, intro]: "x \ space M" hoelzl@38656: show "x \ (\n. ?M n \ ?A)" hoelzl@38656: proof (cases "u x") hoelzl@41981: case (real r) with `0 < u x` have "0 < r" by auto hoelzl@41981: obtain j :: nat where "1 / r \ real j" using real_arch_simple .. hoelzl@41981: hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using `0 < r` by auto hoelzl@41981: hence "1 \ real j * r" using real `0 < r` by auto hoelzl@43920: thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) hoelzl@41981: qed (insert `0 < u x`, auto) hoelzl@41981: qed auto hoelzl@47694: finally have "(emeasure M) {x\space M. 0 < u x} = 0" by simp hoelzl@41981: moreover hoelzl@47694: from pos have "AE x in M. \ (u x < 0)" by auto hoelzl@47694: then have "(emeasure M) {x\space M. u x < 0} = 0" hoelzl@47694: using AE_iff_null[of M] u by auto hoelzl@47694: 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}" hoelzl@47694: using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) hoelzl@47694: ultimately show "(emeasure M) ?A = 0" by simp hoelzl@35582: qed hoelzl@35582: qed hoelzl@35582: hoelzl@56996: lemma nn_integral_0_iff_AE: hoelzl@41705: assumes u: "u \ borel_measurable M" hoelzl@56996: shows "integral\<^sup>N M u = 0 \ (AE x in M. u x \ 0)" hoelzl@41705: proof - hoelzl@41981: have sets: "{x\space M. max 0 (u x) \ 0} \ sets M" hoelzl@41705: using u by auto hoelzl@56996: from nn_integral_0_iff[of "\x. max 0 (u x)"] hoelzl@56996: have "integral\<^sup>N M u = 0 \ (AE x in M. max 0 (u x) = 0)" hoelzl@56996: unfolding nn_integral_max_0 hoelzl@47694: using AE_iff_null[OF sets] u by auto hoelzl@47694: also have "\ \ (AE x in M. u x \ 0)" by (auto split: split_max) hoelzl@41981: finally show ?thesis . hoelzl@41705: qed hoelzl@41705: hoelzl@56996: lemma AE_iff_nn_integral: hoelzl@56996: "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>N M (indicator {x. \ P x}) = 0" hoelzl@56996: by (subst nn_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def immler@50244: sets.sets_Collect_neg indicator_def[abs_def] measurable_If) hoelzl@50001: hoelzl@56996: lemma nn_integral_const_If: wenzelm@53015: "(\\<^sup>+x. a \M) = (if 0 \ a then a * (emeasure M) (space M) else 0)" hoelzl@56996: by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) hoelzl@42991: hoelzl@56996: lemma nn_integral_subalgebra: hoelzl@49799: assumes f: "f \ borel_measurable N" "\x. 0 \ f x" hoelzl@47694: and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" hoelzl@56996: shows "integral\<^sup>N N f = integral\<^sup>N M f" hoelzl@39092: proof - hoelzl@49799: have [simp]: "\f :: 'a \ ereal. f \ borel_measurable N \ f \ borel_measurable M" hoelzl@49799: using N by (auto simp: measurable_def) hoelzl@49799: have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" hoelzl@49799: using N by (auto simp add: eventually_ae_filter null_sets_def) hoelzl@49799: have [simp]: "\A. A \ sets N \ A \ sets M" hoelzl@49799: using N by auto hoelzl@49799: from f show ?thesis hoelzl@49799: apply induct hoelzl@56996: apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N) hoelzl@56996: apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric]) hoelzl@49799: done hoelzl@39092: qed hoelzl@39092: hoelzl@56996: lemma nn_integral_nat_function: hoelzl@50097: fixes f :: "'a \ nat" hoelzl@50097: assumes "f \ measurable M (count_space UNIV)" wenzelm@53015: shows "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\t. emeasure M {x\space M. t < f x})" hoelzl@50097: proof - hoelzl@50097: def F \ "\i. {x\space M. i < f x}" hoelzl@50097: with assms have [measurable]: "\i. F i \ sets M" hoelzl@50097: by auto hoelzl@50097: hoelzl@50097: { fix x assume "x \ space M" hoelzl@50097: have "(\i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" hoelzl@50097: using sums_If_finite[of "\i. i < f x" "\_. 1::real"] by simp hoelzl@50097: then have "(\i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))" hoelzl@50097: unfolding sums_ereal . hoelzl@50097: moreover have "\i. ereal (if i < f x then 1 else 0) = indicator (F i) x" hoelzl@50097: using `x \ space M` by (simp add: one_ereal_def F_def) hoelzl@50097: ultimately have "ereal(of_nat(f x)) = (\i. indicator (F i) x)" hoelzl@50097: by (simp add: sums_iff) } wenzelm@53015: then have "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" hoelzl@56996: by (simp cong: nn_integral_cong) hoelzl@50097: also have "\ = (\i. emeasure M (F i))" hoelzl@56996: by (simp add: nn_integral_suminf) hoelzl@50097: finally show ?thesis hoelzl@50097: by (simp add: F_def) hoelzl@50097: qed hoelzl@50097: hoelzl@56994: subsection {* Integral under concrete measures *} hoelzl@56994: hoelzl@56994: subsubsection {* Distributions *} hoelzl@47694: hoelzl@56996: lemma nn_integral_distr': hoelzl@49797: assumes T: "T \ measurable M M'" hoelzl@49799: and f: "f \ borel_measurable (distr M M' T)" "\x. 0 \ f x" hoelzl@56996: shows "integral\<^sup>N (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" hoelzl@49797: using f hoelzl@49797: proof induct hoelzl@49797: case (cong f g) hoelzl@49799: with T show ?case hoelzl@56996: apply (subst nn_integral_cong[of _ f g]) hoelzl@49799: apply simp hoelzl@56996: apply (subst nn_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) hoelzl@49799: apply (simp add: measurable_def Pi_iff) hoelzl@49799: apply simp hoelzl@49797: done hoelzl@49797: next hoelzl@49797: case (set A) hoelzl@49797: then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" hoelzl@49797: by (auto simp: indicator_def) hoelzl@49797: from set T show ?case hoelzl@56996: by (subst nn_integral_cong[OF eq]) hoelzl@56996: (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets) hoelzl@56996: qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add hoelzl@56996: nn_integral_monotone_convergence_SUP le_fun_def incseq_def) hoelzl@47694: hoelzl@56996: lemma nn_integral_distr: hoelzl@56996: "T \ measurable M M' \ f \ borel_measurable M' \ integral\<^sup>N (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" hoelzl@56996: by (subst (1 2) nn_integral_max_0[symmetric]) hoelzl@56996: (simp add: nn_integral_distr') hoelzl@35692: hoelzl@56994: subsubsection {* Counting space *} hoelzl@47694: hoelzl@47694: lemma simple_function_count_space[simp]: hoelzl@47694: "simple_function (count_space A) f \ finite (f ` A)" hoelzl@47694: unfolding simple_function_def by simp hoelzl@47694: hoelzl@56996: lemma nn_integral_count_space: hoelzl@47694: assumes A: "finite {a\A. 0 < f a}" hoelzl@56996: shows "integral\<^sup>N (count_space A) f = (\a|a\A \ 0 < f a. f a)" hoelzl@35582: proof - wenzelm@53015: have *: "(\\<^sup>+x. max 0 (f x) \count_space A) = wenzelm@53015: (\\<^sup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" hoelzl@56996: by (auto intro!: nn_integral_cong hoelzl@47694: simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less) wenzelm@53015: also have "\ = (\a|a\A \ 0 < f a. \\<^sup>+ x. f a * indicator {a} x \count_space A)" hoelzl@56996: by (subst nn_integral_setsum) hoelzl@47694: (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) hoelzl@47694: also have "\ = (\a|a\A \ 0 < f a. f a)" hoelzl@56996: by (auto intro!: setsum_cong simp: nn_integral_cmult_indicator one_ereal_def[symmetric]) hoelzl@56996: finally show ?thesis by (simp add: nn_integral_max_0) hoelzl@47694: qed hoelzl@47694: hoelzl@56996: lemma nn_integral_count_space_finite: wenzelm@53015: "finite A \ (\\<^sup>+x. f x \count_space A) = (\a\A. max 0 (f a))" hoelzl@56996: by (subst nn_integral_max_0[symmetric]) hoelzl@56996: (auto intro!: setsum_mono_zero_left simp: nn_integral_count_space less_le) hoelzl@47694: hoelzl@54418: lemma emeasure_UN_countable: hoelzl@54418: assumes sets: "\i. i \ I \ X i \ sets M" and I: "countable I" hoelzl@54418: assumes disj: "disjoint_family_on X I" hoelzl@54418: shows "emeasure M (UNION I X) = (\\<^sup>+i. emeasure M (X i) \count_space I)" hoelzl@54418: proof cases hoelzl@54418: assume "finite I" with sets disj show ?thesis hoelzl@54418: by (subst setsum_emeasure[symmetric]) hoelzl@56996: (auto intro!: setsum_cong simp add: max_def subset_eq nn_integral_count_space_finite emeasure_nonneg) hoelzl@54418: next hoelzl@54418: assume f: "\ finite I" hoelzl@54418: then have [intro]: "I \ {}" by auto hoelzl@54418: from from_nat_into_inj_infinite[OF I f] from_nat_into[OF this] disj hoelzl@54418: have disj2: "disjoint_family (\i. X (from_nat_into I i))" hoelzl@54418: unfolding disjoint_family_on_def by metis hoelzl@54418: hoelzl@54418: from f have "bij_betw (from_nat_into I) UNIV I" hoelzl@54418: using bij_betw_from_nat_into[OF I] by simp hoelzl@54418: then have "(\i\I. X i) = (\i. (X \ from_nat_into I) i)" haftmann@56154: unfolding SUP_def image_comp [symmetric] by (simp add: bij_betw_def) hoelzl@54418: then have "emeasure M (UNION I X) = emeasure M (\i. X (from_nat_into I i))" hoelzl@54418: by simp hoelzl@54418: also have "\ = (\i. emeasure M (X (from_nat_into I i)))" hoelzl@54418: by (intro suminf_emeasure[symmetric] disj disj2) (auto intro!: sets from_nat_into[OF `I \ {}`]) hoelzl@54418: also have "\ = (\n. \\<^sup>+i. emeasure M (X i) * indicator {from_nat_into I n} i \count_space I)" hoelzl@54418: proof (intro arg_cong[where f=suminf] ext) hoelzl@54418: fix i hoelzl@54418: have eq: "{a \ I. 0 < emeasure M (X a) * indicator {from_nat_into I i} a} hoelzl@54418: = (if 0 < emeasure M (X (from_nat_into I i)) then {from_nat_into I i} else {})" hoelzl@54418: using ereal_0_less_1 hoelzl@54418: by (auto simp: ereal_zero_less_0_iff indicator_def from_nat_into `I \ {}` simp del: ereal_0_less_1) hoelzl@54418: have "(\\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I) = hoelzl@54418: (if 0 < emeasure M (X (from_nat_into I i)) then emeasure M (X (from_nat_into I i)) else 0)" hoelzl@56996: by (subst nn_integral_count_space) (simp_all add: eq) hoelzl@54418: also have "\ = emeasure M (X (from_nat_into I i))" hoelzl@54418: by (simp add: less_le emeasure_nonneg) hoelzl@54418: finally show "emeasure M (X (from_nat_into I i)) = hoelzl@54418: \\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I" .. hoelzl@54418: qed hoelzl@54418: also have "\ = (\\<^sup>+i. emeasure M (X i) \count_space I)" hoelzl@56996: apply (subst nn_integral_suminf[symmetric]) hoelzl@56996: apply (auto simp: emeasure_nonneg intro!: nn_integral_cong) hoelzl@54418: proof - hoelzl@54418: fix x assume "x \ I" hoelzl@54418: 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)" hoelzl@54418: by (intro suminf_finite) (auto simp: indicator_def I f) hoelzl@54418: also have "\ = emeasure M (X x)" hoelzl@54418: by (simp add: I f `x\I`) hoelzl@54418: finally show "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = emeasure M (X x)" . hoelzl@54418: qed hoelzl@54418: finally show ?thesis . hoelzl@54418: qed hoelzl@54418: hoelzl@56994: subsubsection {* Measures with Restricted Space *} hoelzl@54417: hoelzl@56996: lemma nn_integral_restrict_space: hoelzl@54417: assumes \: "\ \ sets M" and f: "f \ borel_measurable M" "\x. 0 \ f x" "\x. x \ space M - \ \ f x = 0" hoelzl@56996: shows "nn_integral (restrict_space M \) f = nn_integral M f" hoelzl@54417: using f proof (induct rule: borel_measurable_induct) hoelzl@54417: case (cong f g) then show ?case hoelzl@56996: using nn_integral_cong[of M f g] nn_integral_cong[of "restrict_space M \" f g] hoelzl@54417: sets.sets_into_space[OF `\ \ sets M`] hoelzl@54417: by (simp add: subset_eq space_restrict_space) hoelzl@54417: next hoelzl@54417: case (set A) hoelzl@54417: then have "A \ \" hoelzl@54417: unfolding indicator_eq_0_iff by (auto dest: sets.sets_into_space) hoelzl@54417: with set `\ \ sets M` sets.sets_into_space[OF `\ \ sets M`] show ?case hoelzl@56996: by (subst nn_integral_indicator') hoelzl@54417: (auto simp add: sets_restrict_space_iff space_restrict_space hoelzl@54417: emeasure_restrict_space Int_absorb2 hoelzl@54417: dest: sets.sets_into_space) hoelzl@54417: next hoelzl@54417: case (mult f c) then show ?case hoelzl@56996: by (cases "c = 0") (simp_all add: measurable_restrict_space1 \ nn_integral_cmult) hoelzl@54417: next hoelzl@54417: case (add f g) then show ?case hoelzl@56996: by (simp add: measurable_restrict_space1 \ nn_integral_add ereal_add_nonneg_eq_0_iff) hoelzl@54417: next hoelzl@54417: case (seq F) then show ?case hoelzl@56996: by (auto simp add: SUP_eq_iff measurable_restrict_space1 \ nn_integral_monotone_convergence_SUP) hoelzl@54417: qed hoelzl@54417: hoelzl@56994: subsubsection {* Measure spaces with an associated density *} hoelzl@47694: hoelzl@47694: definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where wenzelm@53015: "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" hoelzl@35582: hoelzl@47694: lemma hoelzl@47694: shows sets_density[simp]: "sets (density M f) = sets M" hoelzl@47694: and space_density[simp]: "space (density M f) = space M" hoelzl@47694: by (auto simp: density_def) hoelzl@47694: hoelzl@50003: (* FIXME: add conversion to simplify space, sets and measurable *) hoelzl@50003: lemma space_density_imp[measurable_dest]: hoelzl@50003: "\x M f. x \ space (density M f) \ x \ space M" by auto hoelzl@50003: hoelzl@47694: lemma hoelzl@47694: shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" hoelzl@47694: and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" hoelzl@47694: and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" hoelzl@47694: unfolding measurable_def simple_function_def by simp_all hoelzl@47694: hoelzl@47694: lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ hoelzl@47694: (AE x in M. f x = f' x) \ density M f = density M f'" hoelzl@56996: unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed) hoelzl@47694: hoelzl@47694: lemma density_max_0: "density M f = density M (\x. max 0 (f x))" hoelzl@47694: proof - hoelzl@47694: have "\x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" hoelzl@47694: by (auto simp: indicator_def) hoelzl@47694: then show ?thesis hoelzl@56996: unfolding density_def by (simp add: nn_integral_max_0) hoelzl@47694: qed hoelzl@47694: hoelzl@47694: lemma density_ereal_max_0: "density M (\x. ereal (f x)) = density M (\x. ereal (max 0 (f x)))" hoelzl@47694: by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) hoelzl@38656: hoelzl@47694: lemma emeasure_density: hoelzl@50002: assumes f[measurable]: "f \ borel_measurable M" and A[measurable]: "A \ sets M" wenzelm@53015: shows "emeasure (density M f) A = (\\<^sup>+ x. f x * indicator A x \M)" hoelzl@47694: (is "_ = ?\ A") hoelzl@47694: unfolding density_def hoelzl@47694: proof (rule emeasure_measure_of_sigma) hoelzl@47694: show "sigma_algebra (space M) (sets M)" .. hoelzl@47694: show "positive (sets M) ?\" hoelzl@56996: using f by (auto simp: positive_def intro!: nn_integral_nonneg) wenzelm@53015: have \_eq: "?\ = (\A. \\<^sup>+ x. max 0 (f x) * indicator A x \M)" (is "?\ = ?\'") hoelzl@56996: apply (subst nn_integral_max_0[symmetric]) hoelzl@56996: apply (intro ext nn_integral_cong_AE AE_I2) hoelzl@47694: apply (auto simp: indicator_def) hoelzl@47694: done hoelzl@47694: show "countably_additive (sets M) ?\" hoelzl@47694: unfolding \_eq hoelzl@47694: proof (intro countably_additiveI) hoelzl@47694: fix A :: "nat \ 'a set" assume "range A \ sets M" hoelzl@50002: then have "\i. A i \ sets M" by auto hoelzl@47694: then have *: "\i. (\x. max 0 (f x) * indicator (A i) x) \ borel_measurable M" hoelzl@50002: by (auto simp: set_eq_iff) hoelzl@47694: assume disj: "disjoint_family A" wenzelm@53015: have "(\n. ?\' (A n)) = (\\<^sup>+ x. (\n. max 0 (f x) * indicator (A n) x) \M)" hoelzl@56996: using f * by (simp add: nn_integral_suminf) wenzelm@53015: also have "\ = (\\<^sup>+ x. max 0 (f x) * (\n. indicator (A n) x) \M)" using f hoelzl@56996: by (auto intro!: suminf_cmult_ereal nn_integral_cong_AE) wenzelm@53015: also have "\ = (\\<^sup>+ x. max 0 (f x) * indicator (\n. A n) x \M)" hoelzl@47694: unfolding suminf_indicator[OF disj] .. hoelzl@47694: finally show "(\n. ?\' (A n)) = ?\' (\x. A x)" by simp hoelzl@47694: qed hoelzl@47694: qed fact hoelzl@38656: hoelzl@47694: lemma null_sets_density_iff: hoelzl@47694: assumes f: "f \ borel_measurable M" hoelzl@47694: shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x \ 0)" hoelzl@47694: proof - hoelzl@47694: { assume "A \ sets M" wenzelm@53015: have eq: "(\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. max 0 (f x) * indicator A x \M)" hoelzl@56996: apply (subst nn_integral_max_0[symmetric]) hoelzl@56996: apply (intro nn_integral_cong) hoelzl@47694: apply (auto simp: indicator_def) hoelzl@47694: done wenzelm@53015: have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ hoelzl@47694: emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" hoelzl@47694: unfolding eq hoelzl@47694: using f `A \ sets M` hoelzl@56996: by (intro nn_integral_0_iff) auto hoelzl@47694: also have "\ \ (AE x in M. max 0 (f x) * indicator A x = 0)" hoelzl@47694: using f `A \ sets M` hoelzl@50002: by (intro AE_iff_measurable[OF _ refl, symmetric]) auto hoelzl@47694: also have "(AE x in M. max 0 (f x) * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" hoelzl@47694: by (auto simp add: indicator_def max_def split: split_if_asm) wenzelm@53015: finally have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } hoelzl@47694: with f show ?thesis hoelzl@47694: by (simp add: null_sets_def emeasure_density cong: conj_cong) hoelzl@47694: qed hoelzl@47694: hoelzl@47694: lemma AE_density: hoelzl@47694: assumes f: "f \ borel_measurable M" hoelzl@47694: shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" hoelzl@47694: proof hoelzl@47694: assume "AE x in density M f. P x" hoelzl@47694: 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" hoelzl@47694: by (auto simp: eventually_ae_filter null_sets_density_iff) hoelzl@47694: then have "AE x in M. x \ N \ P x" by auto hoelzl@47694: with ae show "AE x in M. 0 < f x \ P x" hoelzl@47694: by (rule eventually_elim2) auto hoelzl@47694: next hoelzl@47694: fix N assume ae: "AE x in M. 0 < f x \ P x" hoelzl@47694: then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" hoelzl@47694: by (auto simp: eventually_ae_filter) hoelzl@47694: then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. \ 0 < f x}" hoelzl@47694: "N \ {x\space M. \ 0 < f x} \ sets M" and ae2: "AE x in M. x \ N" immler@50244: using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in) hoelzl@47694: show "AE x in density M f. P x" hoelzl@47694: using ae2 hoelzl@47694: unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] hoelzl@47694: by (intro exI[of _ "N \ {x\space M. \ 0 < f x}"] conjI *) hoelzl@47694: (auto elim: eventually_elim2) hoelzl@35582: qed hoelzl@35582: hoelzl@56996: lemma nn_integral_density': hoelzl@47694: assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" hoelzl@49799: assumes g: "g \ borel_measurable M" "\x. 0 \ g x" hoelzl@56996: shows "integral\<^sup>N (density M f) g = (\\<^sup>+ x. f x * g x \M)" hoelzl@49798: using g proof induct hoelzl@49798: case (cong u v) hoelzl@49799: then show ?case hoelzl@56996: apply (subst nn_integral_cong[OF cong(3)]) hoelzl@56996: apply (simp_all cong: nn_integral_cong) hoelzl@49798: done hoelzl@49798: next hoelzl@49798: case (set A) then show ?case hoelzl@49798: by (simp add: emeasure_density f) hoelzl@49798: next hoelzl@49798: case (mult u c) hoelzl@49798: moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) hoelzl@49798: ultimately show ?case hoelzl@56996: using f by (simp add: nn_integral_cmult) hoelzl@49798: next hoelzl@49798: case (add u v) wenzelm@53374: then have "\x. f x * (v x + u x) = f x * v x + f x * u x" hoelzl@49798: by (simp add: ereal_right_distrib) wenzelm@53374: with add f show ?case hoelzl@56996: by (auto simp add: nn_integral_add ereal_zero_le_0_iff intro!: nn_integral_add[symmetric]) hoelzl@49798: next hoelzl@49798: case (seq U) hoelzl@49798: from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" haftmann@56212: by eventually_elim (simp add: SUP_ereal_cmult seq) hoelzl@49798: from seq f show ?case hoelzl@56996: apply (simp add: nn_integral_monotone_convergence_SUP) hoelzl@56996: apply (subst nn_integral_cong_AE[OF eq]) hoelzl@56996: apply (subst nn_integral_monotone_convergence_SUP_AE) hoelzl@49798: apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) hoelzl@49798: done hoelzl@47694: qed hoelzl@38705: hoelzl@56996: lemma nn_integral_density: hoelzl@49798: "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ hoelzl@56996: integral\<^sup>N (density M f) g' = (\\<^sup>+ x. f x * g' x \M)" hoelzl@56996: by (subst (1 2) nn_integral_max_0[symmetric]) hoelzl@56996: (auto intro!: nn_integral_cong_AE hoelzl@56996: simp: measurable_If max_def ereal_zero_le_0_iff nn_integral_density') hoelzl@49798: hoelzl@47694: lemma emeasure_restricted: hoelzl@47694: assumes S: "S \ sets M" and X: "X \ sets M" hoelzl@47694: shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" hoelzl@38705: proof - wenzelm@53015: have "emeasure (density M (indicator S)) X = (\\<^sup>+x. indicator S x * indicator X x \M)" hoelzl@47694: using S X by (simp add: emeasure_density) wenzelm@53015: also have "\ = (\\<^sup>+x. indicator (S \ X) x \M)" hoelzl@56996: by (auto intro!: nn_integral_cong simp: indicator_def) hoelzl@47694: also have "\ = emeasure M (S \ X)" immler@50244: using S X by (simp add: sets.Int) hoelzl@47694: finally show ?thesis . hoelzl@47694: qed hoelzl@47694: hoelzl@47694: lemma measure_restricted: hoelzl@47694: "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" hoelzl@47694: by (simp add: emeasure_restricted measure_def) hoelzl@47694: hoelzl@47694: lemma (in finite_measure) finite_measure_restricted: hoelzl@47694: "S \ sets M \ finite_measure (density M (indicator S))" hoelzl@47694: by default (simp add: emeasure_restricted) hoelzl@47694: hoelzl@47694: lemma emeasure_density_const: hoelzl@47694: "A \ sets M \ 0 \ c \ emeasure (density M (\_. c)) A = c * emeasure M A" hoelzl@56996: by (auto simp: nn_integral_cmult_indicator emeasure_density) hoelzl@47694: hoelzl@47694: lemma measure_density_const: hoelzl@47694: "A \ sets M \ 0 < c \ c \ \ \ measure (density M (\_. c)) A = real c * measure M A" hoelzl@47694: by (auto simp: emeasure_density_const measure_def) hoelzl@47694: hoelzl@47694: lemma density_density_eq: hoelzl@47694: "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. 0 \ f x \ hoelzl@47694: density (density M f) g = density M (\x. f x * g x)" hoelzl@56996: by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps) hoelzl@47694: hoelzl@47694: lemma distr_density_distr: hoelzl@47694: assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" hoelzl@47694: and inv: "\x\space M. T' (T x) = x" hoelzl@47694: assumes f: "f \ borel_measurable M'" hoelzl@47694: shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") hoelzl@47694: proof (rule measure_eqI) hoelzl@47694: fix A assume A: "A \ sets ?R" hoelzl@47694: { fix x assume "x \ space M" immler@50244: with sets.sets_into_space[OF A] hoelzl@47694: have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ereal)" hoelzl@47694: using T inv by (auto simp: indicator_def measurable_space) } hoelzl@47694: with A T T' f show "emeasure ?R A = emeasure ?L A" hoelzl@47694: by (simp add: measurable_comp emeasure_density emeasure_distr hoelzl@56996: nn_integral_distr measurable_sets cong: nn_integral_cong) hoelzl@47694: qed simp hoelzl@47694: hoelzl@47694: lemma density_density_divide: hoelzl@47694: fixes f g :: "'a \ real" hoelzl@47694: assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" hoelzl@47694: assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" hoelzl@47694: assumes ac: "AE x in M. f x = 0 \ g x = 0" hoelzl@47694: shows "density (density M f) (\x. g x / f x) = density M g" hoelzl@47694: proof - hoelzl@47694: have "density M g = density M (\x. f x * (g x / f x))" hoelzl@47694: using f g ac by (auto intro!: density_cong measurable_If) hoelzl@47694: then show ?thesis hoelzl@47694: using f g by (subst density_density_eq) auto hoelzl@38705: qed hoelzl@38705: hoelzl@56994: subsubsection {* Point measure *} hoelzl@47694: hoelzl@47694: definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where hoelzl@47694: "point_measure A f = density (count_space A) f" hoelzl@47694: hoelzl@47694: lemma hoelzl@47694: shows space_point_measure: "space (point_measure A f) = A" hoelzl@47694: and sets_point_measure: "sets (point_measure A f) = Pow A" hoelzl@47694: by (auto simp: point_measure_def) hoelzl@47694: hoelzl@47694: lemma measurable_point_measure_eq1[simp]: hoelzl@47694: "g \ measurable (point_measure A f) M \ g \ A \ space M" hoelzl@47694: unfolding point_measure_def by simp hoelzl@47694: hoelzl@47694: lemma measurable_point_measure_eq2_finite[simp]: hoelzl@47694: "finite A \ hoelzl@47694: g \ measurable M (point_measure A f) \ hoelzl@47694: (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" hoelzl@50002: unfolding point_measure_def by (simp add: measurable_count_space_eq2) hoelzl@47694: hoelzl@47694: lemma simple_function_point_measure[simp]: hoelzl@47694: "simple_function (point_measure A f) g \ finite (g ` A)" hoelzl@47694: by (simp add: point_measure_def) hoelzl@47694: hoelzl@47694: lemma emeasure_point_measure: hoelzl@47694: assumes A: "finite {a\X. 0 < f a}" "X \ A" hoelzl@47694: shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" hoelzl@35977: proof - hoelzl@47694: have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" hoelzl@47694: using `X \ A` by auto hoelzl@47694: with A show ?thesis hoelzl@56996: by (simp add: emeasure_density nn_integral_count_space ereal_zero_le_0_iff hoelzl@47694: point_measure_def indicator_def) hoelzl@35977: qed hoelzl@35977: hoelzl@47694: lemma emeasure_point_measure_finite: hoelzl@49795: "finite A \ (\i. i \ A \ 0 \ f i) \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" hoelzl@47694: by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) hoelzl@47694: hoelzl@49795: lemma emeasure_point_measure_finite2: hoelzl@49795: "X \ A \ finite X \ (\i. i \ X \ 0 \ f i) \ emeasure (point_measure A f) X = (\a\X. f a)" hoelzl@49795: by (subst emeasure_point_measure) hoelzl@49795: (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) hoelzl@49795: hoelzl@47694: lemma null_sets_point_measure_iff: hoelzl@47694: "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x \ 0)" hoelzl@47694: by (auto simp: AE_count_space null_sets_density_iff point_measure_def) hoelzl@47694: hoelzl@47694: lemma AE_point_measure: hoelzl@47694: "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" hoelzl@47694: unfolding point_measure_def hoelzl@47694: by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) hoelzl@47694: hoelzl@56996: lemma nn_integral_point_measure: hoelzl@47694: "finite {a\A. 0 < f a \ 0 < g a} \ hoelzl@56996: integral\<^sup>N (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" hoelzl@47694: unfolding point_measure_def hoelzl@47694: apply (subst density_max_0) hoelzl@56996: apply (subst nn_integral_density) hoelzl@56996: apply (simp_all add: AE_count_space nn_integral_density) hoelzl@56996: apply (subst nn_integral_count_space ) hoelzl@47694: apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff) hoelzl@47694: apply (rule finite_subset) hoelzl@47694: prefer 2 hoelzl@47694: apply assumption hoelzl@47694: apply auto hoelzl@47694: done hoelzl@47694: hoelzl@56996: lemma nn_integral_point_measure_finite: hoelzl@47694: "finite A \ (\a. a \ A \ 0 \ f a) \ (\a. a \ A \ 0 \ g a) \ hoelzl@56996: integral\<^sup>N (point_measure A f) g = (\a\A. f a * g a)" hoelzl@56996: by (subst nn_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le) hoelzl@47694: hoelzl@56994: subsubsection {* Uniform measure *} hoelzl@47694: hoelzl@47694: definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" hoelzl@47694: hoelzl@47694: lemma hoelzl@47694: shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M" hoelzl@47694: and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" hoelzl@47694: by (auto simp: uniform_measure_def) hoelzl@47694: hoelzl@47694: lemma emeasure_uniform_measure[simp]: hoelzl@47694: assumes A: "A \ sets M" and B: "B \ sets M" hoelzl@47694: shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" hoelzl@47694: proof - wenzelm@53015: from A B have "emeasure (uniform_measure M A) B = (\\<^sup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" hoelzl@47694: by (auto simp add: uniform_measure_def emeasure_density split: split_indicator hoelzl@56996: intro!: nn_integral_cong) hoelzl@47694: also have "\ = emeasure M (A \ B) / emeasure M A" hoelzl@47694: using A B hoelzl@56996: by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg) hoelzl@47694: finally show ?thesis . hoelzl@47694: qed hoelzl@47694: hoelzl@47694: lemma measure_uniform_measure[simp]: hoelzl@47694: assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" hoelzl@47694: shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" hoelzl@47694: using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A hoelzl@47694: by (cases "emeasure M A" "emeasure M (A \ B)" rule: ereal2_cases) (simp_all add: measure_def) hoelzl@47694: hoelzl@56994: subsubsection {* Uniform count measure *} hoelzl@47694: hoelzl@47694: definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" hoelzl@47694: hoelzl@47694: lemma hoelzl@47694: shows space_uniform_count_measure: "space (uniform_count_measure A) = A" hoelzl@47694: and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" hoelzl@47694: unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) hoelzl@47694: hoelzl@47694: lemma emeasure_uniform_count_measure: hoelzl@47694: "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" hoelzl@47694: by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) hoelzl@47694: hoelzl@47694: lemma measure_uniform_count_measure: hoelzl@47694: "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" hoelzl@47694: by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) hoelzl@47694: hoelzl@35748: end