diff -r 5001ed24e129 -r d5d342611edb src/HOL/Probability/Lebesgue.thy --- a/src/HOL/Probability/Lebesgue.thy Mon Aug 23 17:46:13 2010 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1733 +0,0 @@ -header {*Lebesgue Integration*} - -theory Lebesgue -imports Measure Borel -begin - -text{*From the HOL4 Hurd/Coble Lebesgue integration, translated by Armin Heller and Johannes Hoelzl.*} - -definition - "pos_part f = (\x. max 0 (f x))" - -definition - "neg_part f = (\x. - min 0 (f x))" - -definition - "nonneg f = (\x. 0 \ f x)" - -lemma nonneg_pos_part[intro!]: - fixes f :: "'c \ 'd\{linorder,zero}" - shows "nonneg (pos_part f)" - unfolding nonneg_def pos_part_def by auto - -lemma nonneg_neg_part[intro!]: - fixes f :: "'c \ 'd\{linorder,ordered_ab_group_add}" - shows "nonneg (neg_part f)" - unfolding nonneg_def neg_part_def min_def by auto - -lemma pos_neg_part_abs: - fixes f :: "'a \ real" - shows "pos_part f x + neg_part f x = \f x\" -unfolding abs_if pos_part_def neg_part_def by auto - -lemma pos_part_abs: - fixes f :: "'a \ real" - shows "pos_part (\ x. \f x\) y = \f y\" -unfolding pos_part_def abs_if by auto - -lemma neg_part_abs: - fixes f :: "'a \ real" - shows "neg_part (\ x. \f x\) y = 0" -unfolding neg_part_def abs_if by auto - -lemma (in measure_space) - assumes "f \ borel_measurable M" - shows pos_part_borel_measurable: "pos_part f \ borel_measurable M" - and neg_part_borel_measurable: "neg_part f \ borel_measurable M" -using assms -proof - - { fix a :: real - { assume asm: "0 \ a" - from asm have pp: "\ w. (pos_part f w \ a) = (f w \ a)" unfolding pos_part_def by auto - have "{w | w. w \ space M \ f w \ a} \ sets M" - unfolding pos_part_def using assms borel_measurable_le_iff by auto - hence "{w . w \ space M \ pos_part f w \ a} \ sets M" using pp by auto } - moreover have "a < 0 \ {w \ space M. pos_part f w \ a} \ sets M" - unfolding pos_part_def using empty_sets by auto - ultimately have "{w . w \ space M \ pos_part f w \ a} \ sets M" - using le_less_linear by auto - } hence pos: "pos_part f \ borel_measurable M" using borel_measurable_le_iff by auto - { fix a :: real - { assume asm: "0 \ a" - from asm have pp: "\ w. (neg_part f w \ a) = (f w \ - a)" unfolding neg_part_def by auto - have "{w | w. w \ space M \ f w \ - a} \ sets M" - unfolding neg_part_def using assms borel_measurable_ge_iff by auto - hence "{w . w \ space M \ neg_part f w \ a} \ sets M" using pp by auto } - moreover have "a < 0 \ {w \ space M. neg_part f w \ a} = {}" unfolding neg_part_def by auto - moreover hence "a < 0 \ {w \ space M. neg_part f w \ a} \ sets M" by (simp only: empty_sets) - ultimately have "{w . w \ space M \ neg_part f w \ a} \ sets M" - using le_less_linear by auto - } hence neg: "neg_part f \ borel_measurable M" using borel_measurable_le_iff by auto - from pos neg show "pos_part f \ borel_measurable M" and "neg_part f \ borel_measurable M" by auto -qed - -lemma (in measure_space) pos_part_neg_part_borel_measurable_iff: - "f \ borel_measurable M \ - pos_part f \ borel_measurable M \ neg_part f \ borel_measurable M" -proof - - { fix x - have "pos_part f x - neg_part f x = f x" - unfolding pos_part_def neg_part_def unfolding max_def min_def - by auto } - hence "(\ x. pos_part f x - neg_part f x) = (\x. f x)" by auto - hence f: "(\ x. pos_part f x - neg_part f x) = f" by blast - from pos_part_borel_measurable[of f] neg_part_borel_measurable[of f] - borel_measurable_diff_borel_measurable[of "pos_part f" "neg_part f"] - show ?thesis unfolding f by fast -qed - -context measure_space -begin - -section "Simple discrete step function" - -definition - "pos_simple f = - { (s :: nat set, a, x). - finite s \ nonneg f \ nonneg x \ a ` s \ sets M \ - (\t \ space M. (\!i\s. t\a i) \ (\i\s. t \ a i \ f t = x i)) }" - -definition - "pos_simple_integral \ \(s, a, x). \ i \ s. x i * measure M (a i)" - -definition - "psfis f = pos_simple_integral ` (pos_simple f)" - -lemma pos_simpleE[consumes 1]: - assumes ps: "(s, a, x) \ pos_simple f" - obtains "finite s" and "nonneg f" and "nonneg x" - and "a ` s \ sets M" and "(\i\s. a i) = space M" - and "disjoint_family_on a s" - and "\t. t \ space M \ (\!i. i \ s \ t \ a i)" - and "\t i. \ t \ space M ; i \ s ; t \ a i \ \ f t = x i" -proof - show "finite s" and "nonneg f" and "nonneg x" - and as_in_M: "a ` s \ sets M" - and *: "\t. t \ space M \ (\!i. i \ s \ t \ a i)" - and **: "\t i. \ t \ space M ; i \ s ; t \ a i \ \ f t = x i" - using ps unfolding pos_simple_def by auto - - thus t: "(\i\s. a i) = space M" - proof safe - fix x assume "x \ space M" - from *[OF this] show "x \ (\i\s. a i)" by auto - next - fix t i assume "i \ s" - hence "a i \ sets M" using as_in_M by auto - moreover assume "t \ a i" - ultimately show "t \ space M" using sets_into_space by blast - qed - - show "disjoint_family_on a s" unfolding disjoint_family_on_def - proof safe - fix i j and t assume "i \ s" "t \ a i" and "j \ s" "t \ a j" and "i \ j" - with t * show "t \ {}" by auto - qed -qed - -lemma pos_simple_cong: - assumes "nonneg f" and "nonneg g" and "\t. t \ space M \ f t = g t" - shows "pos_simple f = pos_simple g" - unfolding pos_simple_def using assms by auto - -lemma psfis_cong: - assumes "nonneg f" and "nonneg g" and "\t. t \ space M \ f t = g t" - shows "psfis f = psfis g" - unfolding psfis_def using pos_simple_cong[OF assms] by simp - -lemma psfis_0: "0 \ psfis (\x. 0)" - unfolding psfis_def pos_simple_def pos_simple_integral_def - by (auto simp: nonneg_def - intro: image_eqI[where x="({0}, (\i. space M), (\i. 0))"]) - -lemma pos_simple_setsum_indicator_fn: - assumes ps: "(s, a, x) \ pos_simple f" - and "t \ space M" - shows "(\i\s. x i * indicator_fn (a i) t) = f t" -proof - - from assms obtain i where *: "i \ s" "t \ a i" - and "finite s" and xi: "x i = f t" by (auto elim!: pos_simpleE) - - have **: "(\i\s. x i * indicator_fn (a i) t) = - (\j\s. if j \ {i} then x i else 0)" - unfolding indicator_fn_def using assms * - by (auto intro!: setsum_cong elim!: pos_simpleE) - show ?thesis unfolding ** setsum_cases[OF `finite s`] xi - using `i \ s` by simp -qed - -lemma pos_simple_common_partition: - assumes psf: "(s, a, x) \ pos_simple f" - assumes psg: "(s', b, y) \ pos_simple g" - obtains z z' c k where "(k, c, z) \ pos_simple f" "(k, c, z') \ pos_simple g" -proof - - (* definitions *) - - def k == "{0 ..< card (s \ s')}" - have fs: "finite s" "finite s'" "finite k" unfolding k_def - using psf psg unfolding pos_simple_def by auto - hence "finite (s \ s')" by simp - then obtain p where p: "p ` k = s \ s'" "inj_on p k" - using ex_bij_betw_nat_finite[of "s \ s'"] unfolding bij_betw_def k_def by blast - def c == "\ i. a (fst (p i)) \ b (snd (p i))" - def z == "\ i. x (fst (p i))" - def z' == "\ i. y (snd (p i))" - - have "finite k" unfolding k_def by simp - - have "nonneg z" and "nonneg z'" - using psf psg unfolding z_def z'_def pos_simple_def nonneg_def by auto - - have ck_subset_M: "c ` k \ sets M" - proof - fix x assume "x \ c ` k" - then obtain j where "j \ k" and "x = c j" by auto - hence "p j \ s \ s'" using p(1) by auto - hence "a (fst (p j)) \ sets M" (is "?a \ _") - and "b (snd (p j)) \ sets M" (is "?b \ _") - using psf psg unfolding pos_simple_def by auto - thus "x \ sets M" unfolding `x = c j` c_def using Int by simp - qed - - { fix t assume "t \ space M" - hence ex1s: "\!i\s. t \ a i" and ex1s': "\!i\s'. t \ b i" - using psf psg unfolding pos_simple_def by auto - then obtain j j' where j: "j \ s" "t \ a j" and j': "j' \ s'" "t \ b j'" - by auto - then obtain i :: nat where i: "(j,j') = p i" and "i \ k" using p by auto - have "\!i\k. t \ c i" - proof (rule ex1I[of _ i]) - show "\x. x \ k \ t \ c x \ x = i" - proof - - fix l assume "l \ k" "t \ c l" - hence "p l \ s \ s'" and t_in: "t \ a (fst (p l))" "t \ b (snd (p l))" - using p unfolding c_def by auto - hence "fst (p l) \ s" and "snd (p l) \ s'" by auto - with t_in j j' ex1s ex1s' have "p l = (j, j')" by (cases "p l", auto) - thus "l = i" - using `(j, j') = p i` p(2)[THEN inj_onD] `l \ k` `i \ k` by auto - qed - - show "i \ k \ t \ c i" - using `i \ k` `t \ a j` `t \ b j'` c_def i[symmetric] by auto - qed auto - } note ex1 = this - - show thesis - proof (rule that) - { fix t i assume "t \ space M" and "i \ k" - hence "p i \ s \ s'" using p(1) by auto - hence "fst (p i) \ s" by auto - moreover - assume "t \ c i" - hence "t \ a (fst (p i))" unfolding c_def by auto - ultimately have "f t = z i" using psf `t \ space M` - unfolding z_def pos_simple_def by auto - } - thus "(k, c, z) \ pos_simple f" - using psf `finite k` `nonneg z` ck_subset_M ex1 - unfolding pos_simple_def by auto - - { fix t i assume "t \ space M" and "i \ k" - hence "p i \ s \ s'" using p(1) by auto - hence "snd (p i) \ s'" by auto - moreover - assume "t \ c i" - hence "t \ b (snd (p i))" unfolding c_def by auto - ultimately have "g t = z' i" using psg `t \ space M` - unfolding z'_def pos_simple_def by auto - } - thus "(k, c, z') \ pos_simple g" - using psg `finite k` `nonneg z'` ck_subset_M ex1 - unfolding pos_simple_def by auto - qed -qed - -lemma pos_simple_integral_equal: - assumes psx: "(s, a, x) \ pos_simple f" - assumes psy: "(s', b, y) \ pos_simple f" - shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)" - unfolding pos_simple_integral_def -proof simp - have "(\i\s. x i * measure M (a i)) = - (\i\s. (\j \ s'. x i * measure M (a i \ b j)))" (is "?left = _") - using psy psx unfolding setsum_right_distrib[symmetric] - by (auto intro!: setsum_cong measure_setsum_split elim!: pos_simpleE) - also have "... = (\i\s. (\j \ s'. y j * measure M (a i \ b j)))" - proof (rule setsum_cong, simp, rule setsum_cong, simp_all) - fix i j assume i: "i \ s" and j: "j \ s'" - hence "a i \ sets M" using psx by (auto elim!: pos_simpleE) - - show "measure M (a i \ b j) = 0 \ x i = y j" - proof (cases "a i \ b j = {}") - case True thus ?thesis using empty_measure by simp - next - case False then obtain t where t: "t \ a i" "t \ b j" by auto - hence "t \ space M" using `a i \ sets M` sets_into_space by auto - with psx psy t i j have "x i = f t" and "y j = f t" - unfolding pos_simple_def by auto - thus ?thesis by simp - qed - qed - also have "... = (\j\s'. (\i\s. y j * measure M (a i \ b j)))" - by (subst setsum_commute) simp - also have "... = (\i\s'. y i * measure M (b i))" (is "?sum_sum = ?right") - proof (rule setsum_cong) - fix j assume "j \ s'" - have "y j * measure M (b j) = (\i\s. y j * measure M (b j \ a i))" - using psx psy `j \ s'` unfolding setsum_right_distrib[symmetric] - by (auto intro!: measure_setsum_split elim!: pos_simpleE) - thus "(\i\s. y j * measure M (a i \ b j)) = y j * measure M (b j)" - by (auto intro!: setsum_cong arg_cong[where f="measure M"]) - qed simp - finally show "?left = ?right" . -qed - -lemma psfis_present: - assumes "A \ psfis f" - assumes "B \ psfis g" - obtains z z' c k where - "A = pos_simple_integral (k, c, z)" - "B = pos_simple_integral (k, c, z')" - "(k, c, z) \ pos_simple f" - "(k, c, z') \ pos_simple g" -using assms -proof - - from assms obtain s a x s' b y where - ps: "(s, a, x) \ pos_simple f" "(s', b, y) \ pos_simple g" and - A: "A = pos_simple_integral (s, a, x)" and - B: "B = pos_simple_integral (s', b, y)" - unfolding psfis_def pos_simple_integral_def by auto - - guess z z' c k using pos_simple_common_partition[OF ps] . note part = this - show thesis - proof (rule that[of k c z z', OF _ _ part]) - show "A = pos_simple_integral (k, c, z)" - unfolding A by (rule pos_simple_integral_equal[OF ps(1) part(1)]) - show "B = pos_simple_integral (k, c, z')" - unfolding B by (rule pos_simple_integral_equal[OF ps(2) part(2)]) - qed -qed - -lemma pos_simple_integral_add: - assumes "(s, a, x) \ pos_simple f" - assumes "(s', b, y) \ pos_simple g" - obtains s'' c z where - "(s'', c, z) \ pos_simple (\x. f x + g x)" - "(pos_simple_integral (s, a, x) + - pos_simple_integral (s', b, y) = - pos_simple_integral (s'', c, z))" -using assms -proof - - guess z z' c k - by (rule pos_simple_common_partition[OF `(s, a, x) \ pos_simple f ` `(s', b, y) \ pos_simple g`]) - note kczz' = this - have x: "pos_simple_integral (s, a, x) = pos_simple_integral (k, c, z)" - by (rule pos_simple_integral_equal) (fact, fact) - have y: "pos_simple_integral (s', b, y) = pos_simple_integral (k, c, z')" - by (rule pos_simple_integral_equal) (fact, fact) - - have "pos_simple_integral (k, c, (\ x. z x + z' x)) - = (\ x \ k. (z x + z' x) * measure M (c x))" - unfolding pos_simple_integral_def by auto - also have "\ = (\ x \ k. z x * measure M (c x) + z' x * measure M (c x))" - by (simp add: left_distrib) - also have "\ = (\ x \ k. z x * measure M (c x)) + (\ x \ k. z' x * measure M (c x))" using setsum_addf by auto - also have "\ = pos_simple_integral (k, c, z) + pos_simple_integral (k, c, z')" unfolding pos_simple_integral_def by auto - finally have ths: "pos_simple_integral (s, a, x) + pos_simple_integral (s', b, y) = - pos_simple_integral (k, c, (\ x. z x + z' x))" using x y by auto - show ?thesis - apply (rule that[of k c "(\ x. z x + z' x)", OF _ ths]) - using kczz' unfolding pos_simple_def nonneg_def by (auto intro!: add_nonneg_nonneg) -qed - -lemma psfis_add: - assumes "a \ psfis f" "b \ psfis g" - shows "a + b \ psfis (\x. f x + g x)" -using assms pos_simple_integral_add -unfolding psfis_def by auto blast - -lemma pos_simple_integral_mono_on_mspace: - assumes f: "(s, a, x) \ pos_simple f" - assumes g: "(s', b, y) \ pos_simple g" - assumes mono: "\ x. x \ space M \ f x \ g x" - shows "pos_simple_integral (s, a, x) \ pos_simple_integral (s', b, y)" -using assms -proof - - guess z z' c k by (rule pos_simple_common_partition[OF f g]) - note kczz' = this - (* w = z and w' = z' except where c _ = {} or undef *) - def w == "\ i. if i \ k \ c i = {} then 0 else z i" - def w' == "\ i. if i \ k \ c i = {} then 0 else z' i" - { fix i - have "w i \ w' i" - proof (cases "i \ k \ c i = {}") - case False hence "i \ k" "c i \ {}" by auto - then obtain v where v: "v \ c i" and "c i \ sets M" - using kczz'(1) unfolding pos_simple_def by blast - hence "v \ space M" using sets_into_space by blast - with mono[OF `v \ space M`] kczz' `i \ k` `v \ c i` - have "z i \ z' i" unfolding pos_simple_def by simp - thus ?thesis unfolding w_def w'_def using False by auto - next - case True thus ?thesis unfolding w_def w'_def by auto - qed - } note w_mn = this - - (* some technical stuff for the calculation*) - have "\ i. i \ k \ c i \ sets M" using kczz' unfolding pos_simple_def by auto - hence "\ i. i \ k \ measure M (c i) \ 0" using positive by auto - hence w_mono: "\ i. i \ k \ w i * measure M (c i) \ w' i * measure M (c i)" - using mult_right_mono w_mn by blast - - { fix i have "\i \ k ; z i \ w i\ \ measure M (c i) = 0" - unfolding w_def by (cases "c i = {}") auto } - hence zw: "\ i. i \ k \ z i * measure M (c i) = w i * measure M (c i)" by auto - - { fix i have "i \ k \ z i * measure M (c i) = w i * measure M (c i)" - unfolding w_def by (cases "c i = {}") simp_all } - note zw = this - - { fix i have "i \ k \ z' i * measure M (c i) = w' i * measure M (c i)" - unfolding w'_def by (cases "c i = {}") simp_all } - note z'w' = this - - (* the calculation *) - have "pos_simple_integral (s, a, x) = pos_simple_integral (k, c, z)" - using f kczz'(1) by (rule pos_simple_integral_equal) - also have "\ = (\ i \ k. z i * measure M (c i))" - unfolding pos_simple_integral_def by auto - also have "\ = (\ i \ k. w i * measure M (c i))" - using setsum_cong2[of k, OF zw] by auto - also have "\ \ (\ i \ k. w' i * measure M (c i))" - using setsum_mono[OF w_mono, of k] by auto - also have "\ = (\ i \ k. z' i * measure M (c i))" - using setsum_cong2[of k, OF z'w'] by auto - also have "\ = pos_simple_integral (k, c, z')" - unfolding pos_simple_integral_def by auto - also have "\ = pos_simple_integral (s', b, y)" - using kczz'(2) g by (rule pos_simple_integral_equal) - finally show "pos_simple_integral (s, a, x) \ pos_simple_integral (s', b, y)" - by simp -qed - -lemma pos_simple_integral_mono: - assumes a: "(s, a, x) \ pos_simple f" "(s', b, y) \ pos_simple g" - assumes "\ z. f z \ g z" - shows "pos_simple_integral (s, a, x) \ pos_simple_integral (s', b, y)" -using assms pos_simple_integral_mono_on_mspace by auto - -lemma psfis_mono: - assumes "a \ psfis f" "b \ psfis g" - assumes "\ x. x \ space M \ f x \ g x" - shows "a \ b" -using assms pos_simple_integral_mono_on_mspace unfolding psfis_def by auto - -lemma pos_simple_fn_integral_unique: - assumes "(s, a, x) \ pos_simple f" "(s', b, y) \ pos_simple f" - shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)" -using assms by (rule pos_simple_integral_equal) (* FIXME: redundant lemma *) - -lemma psfis_unique: - assumes "a \ psfis f" "b \ psfis f" - shows "a = b" -using assms by (intro order_antisym psfis_mono [OF _ _ order_refl]) - -lemma pos_simple_integral_indicator: - assumes "A \ sets M" - obtains s a x where - "(s, a, x) \ pos_simple (indicator_fn A)" - "measure M A = pos_simple_integral (s, a, x)" -using assms -proof - - def s == "{0, 1} :: nat set" - def a == "\ i :: nat. if i = 0 then A else space M - A" - def x == "\ i :: nat. if i = 0 then 1 else (0 :: real)" - have eq: "pos_simple_integral (s, a, x) = measure M A" - unfolding s_def a_def x_def pos_simple_integral_def by auto - have "(s, a, x) \ pos_simple (indicator_fn A)" - unfolding pos_simple_def indicator_fn_def s_def a_def x_def nonneg_def - using assms sets_into_space by auto - from that[OF this] eq show thesis by auto -qed - -lemma psfis_indicator: - assumes "A \ sets M" - shows "measure M A \ psfis (indicator_fn A)" -using pos_simple_integral_indicator[OF assms] - unfolding psfis_def image_def by auto - -lemma pos_simple_integral_mult: - assumes f: "(s, a, x) \ pos_simple f" - assumes "0 \ z" - obtains s' b y where - "(s', b, y) \ pos_simple (\x. z * f x)" - "pos_simple_integral (s', b, y) = z * pos_simple_integral (s, a, x)" - using assms that[of s a "\i. z * x i"] - by (simp add: pos_simple_def pos_simple_integral_def setsum_right_distrib algebra_simps nonneg_def mult_nonneg_nonneg) - -lemma psfis_mult: - assumes "r \ psfis f" - assumes "0 \ z" - shows "z * r \ psfis (\x. z * f x)" -using assms -proof - - from assms obtain s a x - where sax: "(s, a, x) \ pos_simple f" - and r: "r = pos_simple_integral (s, a, x)" - unfolding psfis_def image_def by auto - obtain s' b y where - "(s', b, y) \ pos_simple (\x. z * f x)" - "z * pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)" - using pos_simple_integral_mult[OF sax `0 \ z`, of thesis] by auto - thus ?thesis using r unfolding psfis_def image_def by auto -qed - -lemma psfis_setsum_image: - assumes "finite P" - assumes "\i. i \ P \ a i \ psfis (f i)" - shows "(setsum a P) \ psfis (\t. \i \ P. f i t)" -using assms -proof (induct P) - case empty - let ?s = "{0 :: nat}" - let ?a = "\ i. if i = (0 :: nat) then space M else {}" - let ?x = "\ (i :: nat). (0 :: real)" - have "(?s, ?a, ?x) \ pos_simple (\ t. (0 :: real))" - unfolding pos_simple_def image_def nonneg_def by auto - moreover have "(\ i \ ?s. ?x i * measure M (?a i)) = 0" by auto - ultimately have "0 \ psfis (\ t. 0)" - unfolding psfis_def image_def pos_simple_integral_def nonneg_def - by (auto intro!:bexI[of _ "(?s, ?a, ?x)"]) - thus ?case by auto -next - case (insert x P) note asms = this - have "finite P" by fact - have "x \ P" by fact - have "(\i. i \ P \ a i \ psfis (f i)) \ - setsum a P \ psfis (\t. \i\P. f i t)" by fact - have "setsum a (insert x P) = a x + setsum a P" using `finite P` `x \ P` by auto - also have "\ \ psfis (\ t. f x t + (\ i \ P. f i t))" - using asms psfis_add by auto - also have "\ = psfis (\ t. \ i \ insert x P. f i t)" - using `x \ P` `finite P` by auto - finally show ?case by simp -qed - -lemma psfis_intro: - assumes "a ` P \ sets M" and "nonneg x" and "finite P" - shows "(\i\P. x i * measure M (a i)) \ psfis (\t. \i\P. x i * indicator_fn (a i) t)" -using assms psfis_mult psfis_indicator -unfolding image_def nonneg_def -by (auto intro!:psfis_setsum_image) - -lemma psfis_nonneg: "a \ psfis f \ nonneg f" -unfolding psfis_def pos_simple_def by auto - -lemma pos_psfis: "r \ psfis f \ 0 \ r" -unfolding psfis_def pos_simple_integral_def image_def pos_simple_def nonneg_def -using positive[unfolded positive_def] by (auto intro!:setsum_nonneg mult_nonneg_nonneg) - -lemma psfis_borel_measurable: - assumes "a \ psfis f" - shows "f \ borel_measurable M" -using assms -proof - - from assms obtain s a' x where sa'x: "(s, a', x) \ pos_simple f" and sa'xa: "pos_simple_integral (s, a', x) = a" - and fs: "finite s" - unfolding psfis_def pos_simple_integral_def image_def - by (auto elim:pos_simpleE) - { fix w assume "w \ space M" - hence "(f w \ a) = ((\ i \ s. x i * indicator_fn (a' i) w) \ a)" - using pos_simple_setsum_indicator_fn[OF sa'x, of w] by simp - } hence "\ w. (w \ space M \ f w \ a) = (w \ space M \ (\ i \ s. x i * indicator_fn (a' i) w) \ a)" - by auto - { fix i assume "i \ s" - hence "indicator_fn (a' i) \ borel_measurable M" - using borel_measurable_indicator using sa'x[unfolded pos_simple_def] by auto - hence "(\ w. x i * indicator_fn (a' i) w) \ borel_measurable M" - using affine_borel_measurable[of "\ w. indicator_fn (a' i) w" 0 "x i"] - by (simp add: mult_commute) } - from borel_measurable_setsum_borel_measurable[OF fs this] affine_borel_measurable - have "(\ w. (\ i \ s. x i * indicator_fn (a' i) w)) \ borel_measurable M" by auto - from borel_measurable_cong[OF pos_simple_setsum_indicator_fn[OF sa'x]] this - show ?thesis by simp -qed - -lemma psfis_mono_conv_mono: - assumes mono: "mono_convergent u f (space M)" - and ps_u: "\n. x n \ psfis (u n)" - and "x ----> y" - and "r \ psfis s" - and f_upper_bound: "\t. t \ space M \ s t \ f t" - shows "r <= y" -proof (rule field_le_mult_one_interval) - fix z :: real assume "0 < z" and "z < 1" - hence "0 \ z" by auto - let "?B' n" = "{w \ space M. z * s w <= u n w}" - - have "incseq x" unfolding incseq_def - proof safe - fix m n :: nat assume "m \ n" - show "x m \ x n" - proof (rule psfis_mono[OF `x m \ psfis (u m)` `x n \ psfis (u n)`]) - fix t assume "t \ space M" - with mono_convergentD[OF mono this] `m \ n` show "u m t \ u n t" - unfolding incseq_def by auto - qed - qed - - from `r \ psfis s` - obtain s' a x' where r: "r = pos_simple_integral (s', a, x')" - and ps_s: "(s', a, x') \ pos_simple s" - unfolding psfis_def by auto - - { fix t i assume "i \ s'" "t \ a i" - hence "t \ space M" using ps_s by (auto elim!: pos_simpleE) } - note t_in_space = this - - { fix n - from psfis_borel_measurable[OF `r \ psfis s`] - psfis_borel_measurable[OF ps_u] - have "?B' n \ sets M" - by (auto intro!: - borel_measurable_leq_borel_measurable - borel_measurable_times_borel_measurable - borel_measurable_const) } - note B'_in_M = this - - { fix n have "(\i. a i \ ?B' n) ` s' \ sets M" using B'_in_M ps_s - by (auto intro!: Int elim!: pos_simpleE) } - note B'_inter_a_in_M = this - - let "?sum n" = "(\i\s'. x' i * measure M (a i \ ?B' n))" - { fix n - have "z * ?sum n \ x n" - proof (rule psfis_mono[OF _ ps_u]) - have *: "\i t. indicator_fn (?B' n) t * (x' i * indicator_fn (a i) t) = - x' i * indicator_fn (a i \ ?B' n) t" unfolding indicator_fn_def by auto - have ps': "?sum n \ psfis (\t. indicator_fn (?B' n) t * (\i\s'. x' i * indicator_fn (a i) t))" - unfolding setsum_right_distrib * using B'_in_M ps_s - by (auto intro!: psfis_intro Int elim!: pos_simpleE) - also have "... = psfis (\t. indicator_fn (?B' n) t * s t)" (is "psfis ?l = psfis ?r") - proof (rule psfis_cong) - show "nonneg ?l" using psfis_nonneg[OF ps'] . - show "nonneg ?r" using psfis_nonneg[OF `r \ psfis s`] unfolding nonneg_def indicator_fn_def by auto - fix t assume "t \ space M" - show "?l t = ?r t" unfolding pos_simple_setsum_indicator_fn[OF ps_s `t \ space M`] .. - qed - finally show "z * ?sum n \ psfis (\t. z * ?r t)" using psfis_mult[OF _ `0 \ z`] by simp - next - fix t assume "t \ space M" - show "z * (indicator_fn (?B' n) t * s t) \ u n t" - using psfis_nonneg[OF ps_u] unfolding nonneg_def indicator_fn_def by auto - qed } - hence *: "\N. \n\N. z * ?sum n \ x n" by (auto intro!: exI[of _ 0]) - - show "z * r \ y" unfolding r pos_simple_integral_def - proof (rule LIMSEQ_le[OF mult_right.LIMSEQ `x ----> y` *], - simp, rule LIMSEQ_setsum, rule mult_right.LIMSEQ) - fix i assume "i \ s'" - from psfis_nonneg[OF `r \ psfis s`, unfolded nonneg_def] - have "\t. 0 \ s t" by simp - - have *: "(\j. a i \ ?B' j) = a i" - proof (safe, simp, safe) - fix t assume "t \ a i" - thus "t \ space M" using t_in_space[OF `i \ s'`] by auto - show "\j. z * s t \ u j t" - proof (cases "s t = 0") - case True thus ?thesis - using psfis_nonneg[OF ps_u] unfolding nonneg_def by auto - next - case False with `0 \ s t` - have "0 < s t" by auto - hence "z * s t < 1 * s t" using `0 < z` `z < 1` - by (auto intro!: mult_strict_right_mono simp del: mult_1_left) - also have "... \ f t" using f_upper_bound `t \ space M` by auto - finally obtain b where "\j. b \ j \ z * s t < u j t" using `t \ space M` - using mono_conv_outgrow[of "\n. u n t" "f t" "z * s t"] - using mono_convergentD[OF mono] by auto - from this[of b] show ?thesis by (auto intro!: exI[of _ b]) - qed - qed - - show "(\n. measure M (a i \ ?B' n)) ----> measure M (a i)" - proof (safe intro!: - monotone_convergence[of "\n. a i \ ?B' n", unfolded comp_def *]) - fix n show "a i \ ?B' n \ sets M" - using B'_inter_a_in_M[of n] `i \ s'` by auto - next - fix j t assume "t \ space M" and "z * s t \ u j t" - thus "z * s t \ u (Suc j) t" - using mono_convergentD(1)[OF mono, unfolded incseq_def, - rule_format, of t j "Suc j"] - by auto - qed - qed -qed - -section "Continuous posititve integration" - -definition - "nnfis f = { y. \u x. mono_convergent u f (space M) \ - (\n. x n \ psfis (u n)) \ x ----> y }" - -lemma psfis_nnfis: - "a \ psfis f \ a \ nnfis f" - unfolding nnfis_def mono_convergent_def incseq_def - by (auto intro!: exI[of _ "\n. f"] exI[of _ "\n. a"] LIMSEQ_const) - -lemma nnfis_0: "0 \ nnfis (\ x. 0)" - by (rule psfis_nnfis[OF psfis_0]) - -lemma nnfis_times: - assumes "a \ nnfis f" and "0 \ z" - shows "z * a \ nnfis (\t. z * f t)" -proof - - obtain u x where "mono_convergent u f (space M)" and - "\n. x n \ psfis (u n)" "x ----> a" - using `a \ nnfis f` unfolding nnfis_def by auto - with `0 \ z`show ?thesis unfolding nnfis_def mono_convergent_def incseq_def - by (auto intro!: exI[of _ "\n w. z * u n w"] exI[of _ "\n. z * x n"] - LIMSEQ_mult LIMSEQ_const psfis_mult mult_left_mono) -qed - -lemma nnfis_add: - assumes "a \ nnfis f" and "b \ nnfis g" - shows "a + b \ nnfis (\t. f t + g t)" -proof - - obtain u x w y - where "mono_convergent u f (space M)" and - "\n. x n \ psfis (u n)" "x ----> a" and - "mono_convergent w g (space M)" and - "\n. y n \ psfis (w n)" "y ----> b" - using `a \ nnfis f` `b \ nnfis g` unfolding nnfis_def by auto - thus ?thesis unfolding nnfis_def mono_convergent_def incseq_def - by (auto intro!: exI[of _ "\n a. u n a + w n a"] exI[of _ "\n. x n + y n"] - LIMSEQ_add LIMSEQ_const psfis_add add_mono) -qed - -lemma nnfis_mono: - assumes nnfis: "a \ nnfis f" "b \ nnfis g" - and mono: "\t. t \ space M \ f t \ g t" - shows "a \ b" -proof - - obtain u x w y where - mc: "mono_convergent u f (space M)" "mono_convergent w g (space M)" and - psfis: "\n. x n \ psfis (u n)" "\n. y n \ psfis (w n)" and - limseq: "x ----> a" "y ----> b" using nnfis unfolding nnfis_def by auto - show ?thesis - proof (rule LIMSEQ_le_const2[OF limseq(1)], rule exI[of _ 0], safe) - fix n - show "x n \ b" - proof (rule psfis_mono_conv_mono[OF mc(2) psfis(2) limseq(2) psfis(1)]) - fix t assume "t \ space M" - from mono_convergent_le[OF mc(1) this] mono[OF this] - show "u n t \ g t" by (rule order_trans) - qed - qed -qed - -lemma nnfis_unique: - assumes a: "a \ nnfis f" and b: "b \ nnfis f" - shows "a = b" - using nnfis_mono[OF a b] nnfis_mono[OF b a] - by (auto intro!: order_antisym[of a b]) - -lemma psfis_equiv: - assumes "a \ psfis f" and "nonneg g" - and "\t. t \ space M \ f t = g t" - shows "a \ psfis g" - using assms unfolding psfis_def pos_simple_def by auto - -lemma psfis_mon_upclose: - assumes "\m. a m \ psfis (u m)" - shows "\c. c \ psfis (mon_upclose n u)" -proof (induct n) - case 0 thus ?case unfolding mon_upclose.simps using assms .. -next - case (Suc n) - then obtain sn an xn where ps: "(sn, an, xn) \ pos_simple (mon_upclose n u)" - unfolding psfis_def by auto - obtain ss as xs where ps': "(ss, as, xs) \ pos_simple (u (Suc n))" - using assms[of "Suc n"] unfolding psfis_def by auto - from pos_simple_common_partition[OF ps ps'] guess x x' a s . - hence "(s, a, upclose x x') \ pos_simple (mon_upclose (Suc n) u)" - by (simp add: upclose_def pos_simple_def nonneg_def max_def) - thus ?case unfolding psfis_def by auto -qed - -text {* Beppo-Levi monotone convergence theorem *} -lemma nnfis_mon_conv: - assumes mc: "mono_convergent f h (space M)" - and nnfis: "\n. x n \ nnfis (f n)" - and "x ----> z" - shows "z \ nnfis h" -proof - - have "\n. \u y. mono_convergent u (f n) (space M) \ (\m. y m \ psfis (u m)) \ - y ----> x n" - using nnfis unfolding nnfis_def by auto - from choice[OF this] guess u .. - from choice[OF this] guess y .. - hence mc_u: "\n. mono_convergent (u n) (f n) (space M)" - and psfis: "\n m. y n m \ psfis (u n m)" and "\n. y n ----> x n" - by auto - - let "?upclose n" = "mon_upclose n (\m. u m n)" - - have "\c. \n. c n \ psfis (?upclose n)" - by (safe intro!: choice psfis_mon_upclose) (rule psfis) - then guess c .. note c = this[rule_format] - - show ?thesis unfolding nnfis_def - proof (safe intro!: exI) - show mc_upclose: "mono_convergent ?upclose h (space M)" - by (rule mon_upclose_mono_convergent[OF mc_u mc]) - show psfis_upclose: "\n. c n \ psfis (?upclose n)" - using c . - - { fix n m :: nat assume "n \ m" - hence "c n \ c m" - using psfis_mono[OF c c] - using mono_convergentD(1)[OF mc_upclose, unfolded incseq_def] - by auto } - hence "incseq c" unfolding incseq_def by auto - - { fix n - have c_nnfis: "c n \ nnfis (?upclose n)" using c by (rule psfis_nnfis) - from nnfis_mono[OF c_nnfis nnfis] - mon_upclose_le_mono_convergent[OF mc_u] - mono_convergentD(1)[OF mc] - have "c n \ x n" by fast } - note c_less_x = this - - { fix n - note c_less_x[of n] - also have "x n \ z" - proof (rule incseq_le) - show "x ----> z" by fact - from mono_convergentD(1)[OF mc] - show "incseq x" unfolding incseq_def - by (auto intro!: nnfis_mono[OF nnfis nnfis]) - qed - finally have "c n \ z" . } - note c_less_z = this - - have "convergent c" - proof (rule Bseq_mono_convergent[unfolded incseq_def[symmetric]]) - show "Bseq c" - using pos_psfis[OF c] c_less_z - by (auto intro!: BseqI'[of _ z]) - show "incseq c" by fact - qed - then obtain l where l: "c ----> l" unfolding convergent_def by auto - - have "l \ z" using c_less_x l - by (auto intro!: LIMSEQ_le[OF _ `x ----> z`]) - moreover have "z \ l" - proof (rule LIMSEQ_le_const2[OF `x ----> z`], safe intro!: exI[of _ 0]) - fix n - have "l \ nnfis h" - unfolding nnfis_def using l mc_upclose psfis_upclose by auto - from nnfis this mono_convergent_le[OF mc] - show "x n \ l" by (rule nnfis_mono) - qed - ultimately have "l = z" by (rule order_antisym) - thus "c ----> z" using `c ----> l` by simp - qed -qed - -lemma nnfis_pos_on_mspace: - assumes "a \ nnfis f" and "x \space M" - shows "0 \ f x" -using assms -proof - - from assms[unfolded nnfis_def] guess u y by auto note uy = this - hence "\ n. 0 \ u n x" - unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def - by auto - thus "0 \ f x" using uy[rule_format] - unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def - using incseq_le[of "\ n. u n x" "f x"] order_trans - by fast -qed - -lemma nnfis_borel_measurable: - assumes "a \ nnfis f" - shows "f \ borel_measurable M" -using assms unfolding nnfis_def -apply auto -apply (rule mono_convergent_borel_measurable) -using psfis_borel_measurable -by auto - -lemma borel_measurable_mon_conv_psfis: - assumes f_borel: "f \ borel_measurable M" - and nonneg: "\t. t \ space M \ 0 \ f t" - shows"\u x. mono_convergent u f (space M) \ (\n. x n \ psfis (u n))" -proof (safe intro!: exI) - let "?I n" = "{0<.. space M" - have "?u n t = ?w n t" (is "_ = (if _ then real ?i / _ else _)") - proof (cases "f t < real n") - case True - with nonneg t - have i: "?i < n * 2^n" - by (auto simp: real_of_nat_power[symmetric] - intro!: less_natfloor mult_nonneg_nonneg) - - hence t_in_A: "t \ ?A n ?i" - unfolding divide_le_eq less_divide_eq - using nonneg t True - by (auto intro!: real_natfloor_le - real_natfloor_gt_diff_one[unfolded diff_less_eq] - simp: real_of_nat_Suc zero_le_mult_iff) - - hence *: "real ?i / 2^n \ f t" - "f t < real (?i + 1) / 2^n" by auto - { fix j assume "t \ ?A n j" - hence "real j / 2^n \ f t" - and "f t < real (j + 1) / 2^n" by auto - with * have "j \ {?i}" unfolding divide_le_eq less_divide_eq - by auto } - hence *: "\j. t \ ?A n j \ j \ {?i}" using t_in_A by auto - - have "?u n t = real ?i / 2^n" - unfolding indicator_fn_def if_distrib * - setsum_cases[OF finite_greaterThanLessThan] - using i by (cases "?i = 0") simp_all - thus ?thesis using True by auto - next - case False - have "?u n t = (\i \ {0 <..< n*2^n}. 0)" - proof (rule setsum_cong) - fix i assume "i \ {0 <..< n*2^n}" - hence "i + 1 \ n * 2^n" by simp - hence "real (i + 1) \ real n * 2^n" - unfolding real_of_nat_le_iff[symmetric] - by (auto simp: real_of_nat_power[symmetric]) - also have "... \ f t * 2^n" - using False by (auto intro!: mult_nonneg_nonneg) - finally have "t \ ?A n i" - by (auto simp: divide_le_eq less_divide_eq) - thus "real i / 2^n * indicator_fn (?A n i) t = 0" - unfolding indicator_fn_def by auto - qed simp - thus ?thesis using False by auto - qed } - note u_at_t = this - - show "mono_convergent ?u f (space M)" unfolding mono_convergent_def - proof safe - fix t assume t: "t \ space M" - { fix m n :: nat assume "m \ n" - hence *: "(2::real)^n = 2^m * 2^(n - m)" unfolding power_add[symmetric] by auto - have "real (natfloor (f t * 2^m) * natfloor (2^(n-m))) \ real (natfloor (f t * 2 ^ n))" - apply (subst *) - apply (subst mult_assoc[symmetric]) - apply (subst real_of_nat_le_iff) - apply (rule le_mult_natfloor) - using nonneg[OF t] by (auto intro!: mult_nonneg_nonneg) - hence "real (natfloor (f t * 2^m)) * 2^n \ real (natfloor (f t * 2^n)) * 2^m" - apply (subst *) - apply (subst (3) mult_commute) - apply (subst mult_assoc[symmetric]) - by (auto intro: mult_right_mono simp: natfloor_power real_of_nat_power[symmetric]) } - thus "incseq (\n. ?u n t)" unfolding u_at_t[OF t] unfolding incseq_def - by (auto simp add: le_divide_eq divide_le_eq less_divide_eq) - - show "(\i. ?u i t) ----> f t" unfolding u_at_t[OF t] - proof (rule LIMSEQ_I, safe intro!: exI) - fix r :: real and n :: nat - let ?N = "natfloor (1/r) + 1" - assume "0 < r" and N: "max ?N (natfloor (f t) + 1) \ n" - hence "?N \ n" by auto - - have "1 / r < real (natfloor (1/r) + 1)" using real_natfloor_add_one_gt - by (simp add: real_of_nat_Suc) - also have "... < 2^?N" by (rule two_realpow_gt) - finally have less_r: "1 / 2^?N < r" using `0 < r` - by (auto simp: less_divide_eq divide_less_eq algebra_simps) - - have "f t < real (natfloor (f t) + 1)" using real_natfloor_add_one_gt[of "f t"] by auto - also have "... \ real n" unfolding real_of_nat_le_iff using N by auto - finally have "f t < real n" . - - have "real (natfloor (f t * 2^n)) \ f t * 2^n" - using nonneg[OF t] by (auto intro!: real_natfloor_le mult_nonneg_nonneg) - hence less: "real (natfloor (f t * 2^n)) / 2^n \ f t" unfolding divide_le_eq by auto - - have "f t * 2 ^ n - 1 < real (natfloor (f t * 2^n))" using real_natfloor_gt_diff_one . - hence "f t - real (natfloor (f t * 2^n)) / 2^n < 1 / 2^n" - by (auto simp: less_divide_eq divide_less_eq algebra_simps) - also have "... \ 1 / 2^?N" using `?N \ n` - by (auto intro!: divide_left_mono mult_pos_pos simp del: power_Suc) - also have "... < r" using less_r . - finally show "norm (?w n t - f t) < r" using `f t < real n` less by auto - qed - qed - - fix n - show "?x n \ psfis (?u n)" - proof (rule psfis_intro) - show "?A n ` ?I n \ sets M" - proof safe - fix i :: nat - from Int[OF - f_borel[unfolded borel_measurable_less_iff, rule_format, of "real (i+1) / 2^n"] - f_borel[unfolded borel_measurable_ge_iff, rule_format, of "real i / 2^n"]] - show "?A n i \ sets M" - by (metis Collect_conj_eq Int_commute Int_left_absorb Int_left_commute) - qed - show "nonneg (\i :: nat. real i / 2 ^ n)" - unfolding nonneg_def by (auto intro!: divide_nonneg_pos) - qed auto -qed - -lemma nnfis_dom_conv: - assumes borel: "f \ borel_measurable M" - and nnfis: "b \ nnfis g" - and ord: "\t. t \ space M \ f t \ g t" - and nonneg: "\t. t \ space M \ 0 \ f t" - shows "\a. a \ nnfis f \ a \ b" -proof - - obtain u x where mc_f: "mono_convergent u f (space M)" and - psfis: "\n. x n \ psfis (u n)" - using borel_measurable_mon_conv_psfis[OF borel nonneg] by auto - - { fix n - { fix t assume t: "t \ space M" - note mono_convergent_le[OF mc_f this, of n] - also note ord[OF t] - finally have "u n t \ g t" . } - from nnfis_mono[OF psfis_nnfis[OF psfis] nnfis this] - have "x n \ b" by simp } - note x_less_b = this - - have "convergent x" - proof (safe intro!: Bseq_mono_convergent) - from x_less_b pos_psfis[OF psfis] - show "Bseq x" by (auto intro!: BseqI'[of _ b]) - - fix n m :: nat assume *: "n \ m" - show "x n \ x m" - proof (rule psfis_mono[OF `x n \ psfis (u n)` `x m \ psfis (u m)`]) - fix t assume "t \ space M" - from mc_f[THEN mono_convergentD(1), unfolded incseq_def, OF this] - show "u n t \ u m t" using * by auto - qed - qed - then obtain a where "x ----> a" unfolding convergent_def by auto - with LIMSEQ_le_const2[OF `x ----> a`] x_less_b mc_f psfis - show ?thesis unfolding nnfis_def by auto -qed - -lemma the_nnfis[simp]: "a \ nnfis f \ (THE a. a \ nnfis f) = a" - by (auto intro: the_equality nnfis_unique) - -lemma nnfis_cong: - assumes cong: "\x. x \ space M \ f x = g x" - shows "nnfis f = nnfis g" -proof - - { fix f g :: "'a \ real" assume cong: "\x. x \ space M \ f x = g x" - fix x assume "x \ nnfis f" - then guess u y unfolding nnfis_def by safe note x = this - hence "mono_convergent u g (space M)" - unfolding mono_convergent_def using cong by auto - with x(2,3) have "x \ nnfis g" unfolding nnfis_def by auto } - from this[OF cong] this[OF cong[symmetric]] - show ?thesis by safe -qed - -section "Lebesgue Integral" - -definition - "integrable f \ (\x. x \ nnfis (pos_part f)) \ (\y. y \ nnfis (neg_part f))" - -definition - "integral f = (THE i :: real. i \ nnfis (pos_part f)) - (THE j. j \ nnfis (neg_part f))" - -lemma integral_cong: - assumes cong: "\x. x \ space M \ f x = g x" - shows "integral f = integral g" -proof - - have "nnfis (pos_part f) = nnfis (pos_part g)" - using cong by (auto simp: pos_part_def intro!: nnfis_cong) - moreover - have "nnfis (neg_part f) = nnfis (neg_part g)" - using cong by (auto simp: neg_part_def intro!: nnfis_cong) - ultimately show ?thesis - unfolding integral_def by auto -qed - -lemma nnfis_integral: - assumes "a \ nnfis f" - shows "integrable f" and "integral f = a" -proof - - have a: "a \ nnfis (pos_part f)" - using assms nnfis_pos_on_mspace[OF assms] - by (auto intro!: nnfis_mon_conv[of "\i. f" _ "\i. a"] - LIMSEQ_const simp: mono_convergent_def pos_part_def incseq_def max_def) - - have "\t. t \ space M \ neg_part f t = 0" - unfolding neg_part_def min_def - using nnfis_pos_on_mspace[OF assms] by auto - hence 0: "0 \ nnfis (neg_part f)" - by (auto simp: nnfis_def mono_convergent_def psfis_0 incseq_def - intro!: LIMSEQ_const exI[of _ "\ x n. 0"] exI[of _ "\ n. 0"]) - - from 0 a show "integrable f" "integral f = a" - unfolding integrable_def integral_def by auto -qed - -lemma nnfis_minus_nnfis_integral: - assumes "a \ nnfis f" and "b \ nnfis g" - shows "integrable (\t. f t - g t)" and "integral (\t. f t - g t) = a - b" -proof - - have borel: "(\t. f t - g t) \ borel_measurable M" using assms - by (blast intro: - borel_measurable_diff_borel_measurable nnfis_borel_measurable) - - have "\x. x \ nnfis (pos_part (\t. f t - g t)) \ x \ a + b" - (is "\x. x \ nnfis ?pp \ _") - proof (rule nnfis_dom_conv) - show "?pp \ borel_measurable M" - using borel by (rule pos_part_borel_measurable neg_part_borel_measurable) - show "a + b \ nnfis (\t. f t + g t)" using assms by (rule nnfis_add) - fix t assume "t \ space M" - with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]] - show "?pp t \ f t + g t" unfolding pos_part_def by auto - show "0 \ ?pp t" using nonneg_pos_part[of "\t. f t - g t"] - unfolding nonneg_def by auto - qed - then obtain x where x: "x \ nnfis ?pp" by auto - moreover - have "\x. x \ nnfis (neg_part (\t. f t - g t)) \ x \ a + b" - (is "\x. x \ nnfis ?np \ _") - proof (rule nnfis_dom_conv) - show "?np \ borel_measurable M" - using borel by (rule pos_part_borel_measurable neg_part_borel_measurable) - show "a + b \ nnfis (\t. f t + g t)" using assms by (rule nnfis_add) - fix t assume "t \ space M" - with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]] - show "?np t \ f t + g t" unfolding neg_part_def by auto - show "0 \ ?np t" using nonneg_neg_part[of "\t. f t - g t"] - unfolding nonneg_def by auto - qed - then obtain y where y: "y \ nnfis ?np" by auto - ultimately show "integrable (\t. f t - g t)" - unfolding integrable_def by auto - - from x and y - have "a + y \ nnfis (\t. f t + ?np t)" - and "b + x \ nnfis (\t. g t + ?pp t)" using assms by (auto intro: nnfis_add) - moreover - have "\t. f t + ?np t = g t + ?pp t" - unfolding pos_part_def neg_part_def by auto - ultimately have "a - b = x - y" - using nnfis_unique by (auto simp: algebra_simps) - thus "integral (\t. f t - g t) = a - b" - unfolding integral_def - using the_nnfis[OF x] the_nnfis[OF y] by simp -qed - -lemma integral_borel_measurable: - "integrable f \ f \ borel_measurable M" - unfolding integrable_def - by (subst pos_part_neg_part_borel_measurable_iff) - (auto intro: nnfis_borel_measurable) - -lemma integral_indicator_fn: - assumes "a \ sets M" - shows "integral (indicator_fn a) = measure M a" - and "integrable (indicator_fn a)" - using psfis_indicator[OF assms, THEN psfis_nnfis] - by (auto intro!: nnfis_integral) - -lemma integral_add: - assumes "integrable f" and "integrable g" - shows "integrable (\t. f t + g t)" - and "integral (\t. f t + g t) = integral f + integral g" -proof - - { fix t - have "pos_part f t + pos_part g t - (neg_part f t + neg_part g t) = - f t + g t" - unfolding pos_part_def neg_part_def by auto } - note part_sum = this - - from assms obtain a b c d where - a: "a \ nnfis (pos_part f)" and b: "b \ nnfis (neg_part f)" and - c: "c \ nnfis (pos_part g)" and d: "d \ nnfis (neg_part g)" - unfolding integrable_def by auto - note sums = nnfis_add[OF a c] nnfis_add[OF b d] - note int = nnfis_minus_nnfis_integral[OF sums, unfolded part_sum] - - show "integrable (\t. f t + g t)" using int(1) . - - show "integral (\t. f t + g t) = integral f + integral g" - using int(2) sums a b c d by (simp add: the_nnfis integral_def) -qed - -lemma integral_mono: - assumes "integrable f" and "integrable g" - and mono: "\t. t \ space M \ f t \ g t" - shows "integral f \ integral g" -proof - - from assms obtain a b c d where - a: "a \ nnfis (pos_part f)" and b: "b \ nnfis (neg_part f)" and - c: "c \ nnfis (pos_part g)" and d: "d \ nnfis (neg_part g)" - unfolding integrable_def by auto - - have "a \ c" - proof (rule nnfis_mono[OF a c]) - fix t assume "t \ space M" - from mono[OF this] show "pos_part f t \ pos_part g t" - unfolding pos_part_def by auto - qed - moreover have "d \ b" - proof (rule nnfis_mono[OF d b]) - fix t assume "t \ space M" - from mono[OF this] show "neg_part g t \ neg_part f t" - unfolding neg_part_def by auto - qed - ultimately have "a - b \ c - d" by auto - thus ?thesis unfolding integral_def - using a b c d by (simp add: the_nnfis) -qed - -lemma integral_uminus: - assumes "integrable f" - shows "integrable (\t. - f t)" - and "integral (\t. - f t) = - integral f" -proof - - have "pos_part f = neg_part (\t.-f t)" and "neg_part f = pos_part (\t.-f t)" - unfolding pos_part_def neg_part_def by (auto intro!: ext) - with assms show "integrable (\t.-f t)" and "integral (\t.-f t) = -integral f" - unfolding integrable_def integral_def by simp_all -qed - -lemma integral_times_const: - assumes "integrable f" - shows "integrable (\t. a * f t)" (is "?P a") - and "integral (\t. a * f t) = a * integral f" (is "?I a") -proof - - { fix a :: real assume "0 \ a" - hence "pos_part (\t. a * f t) = (\t. a * pos_part f t)" - and "neg_part (\t. a * f t) = (\t. a * neg_part f t)" - unfolding pos_part_def neg_part_def max_def min_def - by (auto intro!: ext simp: zero_le_mult_iff) - moreover - obtain x y where x: "x \ nnfis (pos_part f)" and y: "y \ nnfis (neg_part f)" - using assms unfolding integrable_def by auto - ultimately - have "a * x \ nnfis (pos_part (\t. a * f t))" and - "a * y \ nnfis (neg_part (\t. a * f t))" - using nnfis_times[OF _ `0 \ a`] by auto - with x y have "?P a \ ?I a" - unfolding integrable_def integral_def by (auto simp: algebra_simps) } - note int = this - - have "?P a \ ?I a" - proof (cases "0 \ a") - case True from int[OF this] show ?thesis . - next - case False with int[of "- a"] integral_uminus[of "\t. - a * f t"] - show ?thesis by auto - qed - thus "integrable (\t. a * f t)" - and "integral (\t. a * f t) = a * integral f" by simp_all -qed - -lemma integral_cmul_indicator: - assumes "s \ sets M" - shows "integral (\x. c * indicator_fn s x) = c * (measure M s)" - and "integrable (\x. c * indicator_fn s x)" -using assms integral_times_const integral_indicator_fn by auto - -lemma integral_zero: - shows "integral (\x. 0) = 0" - and "integrable (\x. 0)" - using integral_cmul_indicator[OF empty_sets, of 0] - unfolding indicator_fn_def by auto - -lemma integral_setsum: - assumes "finite S" - assumes "\n. n \ S \ integrable (f n)" - shows "integral (\x. \ i \ S. f i x) = (\ i \ S. integral (f i))" (is "?int S") - and "integrable (\x. \ i \ S. f i x)" (is "?I s") -proof - - from assms have "?int S \ ?I S" - proof (induct S) - case empty thus ?case by (simp add: integral_zero) - next - case (insert i S) - thus ?case - apply simp - apply (subst integral_add) - using assms apply auto - apply (subst integral_add) - using assms by auto - qed - thus "?int S" and "?I S" by auto -qed - -lemma (in measure_space) integrable_abs: - assumes "integrable f" - shows "integrable (\ x. \f x\)" -using assms -proof - - from assms obtain p q where pq: "p \ nnfis (pos_part f)" "q \ nnfis (neg_part f)" - unfolding integrable_def by auto - hence "p + q \ nnfis (\ x. pos_part f x + neg_part f x)" - using nnfis_add by auto - hence "p + q \ nnfis (\ x. \f x\)" using pos_neg_part_abs[of f] by simp - thus ?thesis unfolding integrable_def - using ext[OF pos_part_abs[of f], of "\ y. y"] - ext[OF neg_part_abs[of f], of "\ y. y"] - using nnfis_0 by auto -qed - -lemma markov_ineq: - assumes "integrable f" "0 < a" "integrable (\x. \f x\^n)" - shows "measure M (f -` {a ..} \ space M) \ integral (\x. \f x\^n) / a^n" -using assms -proof - - from assms have "0 < a ^ n" using real_root_pow_pos by auto - from assms have "f \ borel_measurable M" - using integral_borel_measurable[OF `integrable f`] by auto - hence w: "{w . w \ space M \ a \ f w} \ sets M" - using borel_measurable_ge_iff by auto - have i: "integrable (indicator_fn {w . w \ space M \ a \ f w})" - using integral_indicator_fn[OF w] by simp - have v1: "\ t. a ^ n * (indicator_fn {w . w \ space M \ a \ f w}) t - \ (f t) ^ n * (indicator_fn {w . w \ space M \ a \ f w}) t" - unfolding indicator_fn_def - using `0 < a` power_mono[of a] assms by auto - have v2: "\ t. (f t) ^ n * (indicator_fn {w . w \ space M \ a \ f w}) t \ \f t\ ^ n" - unfolding indicator_fn_def - using power_mono[of a _ n] abs_ge_self `a > 0` - by auto - have "{w \ space M. a \ f w} \ space M = {w . a \ f w} \ space M" - using Collect_eq by auto - from Int_absorb2[OF sets_into_space[OF w]] `0 < a ^ n` sets_into_space[OF w] w this - have "(a ^ n) * (measure M ((f -` {y . a \ y}) \ space M)) = - (a ^ n) * measure M {w . w \ space M \ a \ f w}" - unfolding vimage_Collect_eq[of f] by simp - also have "\ = integral (\ t. a ^ n * (indicator_fn {w . w \ space M \ a \ f w}) t)" - using integral_cmul_indicator[OF w] i by auto - also have "\ \ integral (\ t. \ f t \ ^ n)" - apply (rule integral_mono) - using integral_cmul_indicator[OF w] - `integrable (\ x. \f x\ ^ n)` order_trans[OF v1 v2] by auto - finally show "measure M (f -` {a ..} \ space M) \ integral (\x. \f x\^n) / a^n" - unfolding atLeast_def - by (auto intro!: mult_imp_le_div_pos[OF `0 < a ^ n`], simp add: mult_commute) -qed - -lemma (in measure_space) integral_0: - fixes f :: "'a \ real" - assumes "integrable f" "integral f = 0" "nonneg f" and borel: "f \ borel_measurable M" - shows "measure M ({x. f x \ 0} \ space M) = 0" -proof - - have "{x. f x \ 0} = {x. \f x\ > 0}" by auto - moreover - { fix y assume "y \ {x. \ f x \ > 0}" - hence "\ f y \ > 0" by auto - hence "\ n. \f y\ \ inverse (real (Suc n))" - using ex_inverse_of_nat_Suc_less[of "\f y\"] less_imp_le unfolding real_of_nat_def by auto - hence "y \ (\ n. {x. \f x\ \ inverse (real (Suc n))})" - by auto } - moreover - { fix y assume "y \ (\ n. {x. \f x\ \ inverse (real (Suc n))})" - then obtain n where n: "y \ {x. \f x\ \ inverse (real (Suc n))}" by auto - hence "\f y\ \ inverse (real (Suc n))" by auto - hence "\f y\ > 0" - using real_of_nat_Suc_gt_zero - positive_imp_inverse_positive[of "real_of_nat (Suc n)"] by fastsimp - hence "y \ {x. \f x\ > 0}" by auto } - ultimately have fneq0_UN: "{x. f x \ 0} = (\ n. {x. \f x\ \ inverse (real (Suc n))})" - by blast - { fix n - have int_one: "integrable (\ x. \f x\ ^ 1)" using integrable_abs assms by auto - have "measure M (f -` {inverse (real (Suc n))..} \ space M) - \ integral (\ x. \f x\ ^ 1) / (inverse (real (Suc n)) ^ 1)" - using markov_ineq[OF `integrable f` _ int_one] real_of_nat_Suc_gt_zero by auto - hence le0: "measure M (f -` {inverse (real (Suc n))..} \ space M) \ 0" - using assms unfolding nonneg_def by auto - have "{x. f x \ inverse (real (Suc n))} \ space M \ sets M" - apply (subst Int_commute) unfolding Int_def - using borel[unfolded borel_measurable_ge_iff] by simp - hence m0: "measure M ({x. f x \ inverse (real (Suc n))} \ space M) = 0 \ - {x. f x \ inverse (real (Suc n))} \ space M \ sets M" - using positive le0 unfolding atLeast_def by fastsimp } - moreover hence "range (\ n. {x. f x \ inverse (real (Suc n))} \ space M) \ sets M" - by auto - moreover - { fix n - have "inverse (real (Suc n)) \ inverse (real (Suc (Suc n)))" - using less_imp_inverse_less real_of_nat_Suc_gt_zero[of n] by fastsimp - hence "\ x. f x \ inverse (real (Suc n)) \ f x \ inverse (real (Suc (Suc n)))" by (rule order_trans) - hence "{x. f x \ inverse (real (Suc n))} \ space M - \ {x. f x \ inverse (real (Suc (Suc n)))} \ space M" by auto } - ultimately have "(\ x. 0) ----> measure M (\ n. {x. f x \ inverse (real (Suc n))} \ space M)" - using monotone_convergence[of "\ n. {x. f x \ inverse (real (Suc n))} \ space M"] - unfolding o_def by (simp del: of_nat_Suc) - hence "measure M (\ n. {x. f x \ inverse (real (Suc n))} \ space M) = 0" - using LIMSEQ_const[of 0] LIMSEQ_unique by simp - hence "measure M ((\ n. {x. \f x\ \ inverse (real (Suc n))}) \ space M) = 0" - using assms unfolding nonneg_def by auto - thus "measure M ({x. f x \ 0} \ space M) = 0" using fneq0_UN by simp -qed - -section "Lebesgue integration on countable spaces" - -lemma nnfis_on_countable: - assumes borel: "f \ borel_measurable M" - and bij: "bij_betw enum S (f ` space M - {0})" - and enum_zero: "enum ` (-S) \ {0}" - and nn_enum: "\n. 0 \ enum n" - and sums: "(\r. enum r * measure M (f -` {enum r} \ space M)) sums x" (is "?sum sums x") - shows "x \ nnfis f" -proof - - have inj_enum: "inj_on enum S" - and range_enum: "enum ` S = f ` space M - {0}" - using bij by (auto simp: bij_betw_def) - - let "?x n z" = "\i = 0.. space M) z" - - show ?thesis - proof (rule nnfis_mon_conv) - show "(\n. \i = 0.. x" using sums unfolding sums_def . - next - fix n - show "(\i = 0.. nnfis (?x n)" - proof (induct n) - case 0 thus ?case by (simp add: nnfis_0) - next - case (Suc n) thus ?case using nn_enum - by (auto intro!: nnfis_add nnfis_times psfis_nnfis[OF psfis_indicator] borel_measurable_vimage[OF borel]) - qed - next - show "mono_convergent ?x f (space M)" - proof (rule mono_convergentI) - fix x - show "incseq (\n. ?x n x)" - by (rule incseq_SucI, auto simp: indicator_fn_def nn_enum) - - have fin: "\n. finite (enum ` ({0.. S))" by auto - - assume "x \ space M" - hence "f x \ enum ` S \ f x = 0" using range_enum by auto - thus "(\n. ?x n x) ----> f x" - proof (rule disjE) - assume "f x \ enum ` S" - then obtain i where "i \ S" and "f x = enum i" by auto - - { fix n - have sum_ranges: - "i < n \ enum`({0.. S) \ {z. enum i = z \ x\space M} = {enum i}" - "\ i < n \ enum`({0.. S) \ {z. enum i = z \ x\space M} = {}" - using `x \ space M` `i \ S` inj_enum[THEN inj_on_iff] by auto - have "?x n x = - (\i \ {0.. S. enum i * indicator_fn (f -` {enum i} \ space M) x)" - using enum_zero by (auto intro!: setsum_mono_zero_cong_right) - also have "... = - (\z \ enum`({0.. S). z * indicator_fn (f -` {z} \ space M) x)" - using inj_enum[THEN subset_inj_on] by (auto simp: setsum_reindex) - also have "... = (if i < n then f x else 0)" - unfolding indicator_fn_def if_distrib[where x=1 and y=0] - setsum_cases[OF fin] - using sum_ranges `f x = enum i` - by auto - finally have "?x n x = (if i < n then f x else 0)" . } - note sum_equals_if = this - - show ?thesis unfolding sum_equals_if - by (rule LIMSEQ_offset[where k="i + 1"]) (auto intro!: LIMSEQ_const) - next - assume "f x = 0" - { fix n have "?x n x = 0" - unfolding indicator_fn_def if_distrib[where x=1 and y=0] - setsum_cases[OF finite_atLeastLessThan] - using `f x = 0` `x \ space M` - by (auto split: split_if) } - thus ?thesis using `f x = 0` by (auto intro!: LIMSEQ_const) - qed - qed - qed -qed - -lemma integral_on_countable: - fixes enum :: "nat \ real" - assumes borel: "f \ borel_measurable M" - and bij: "bij_betw enum S (f ` space M)" - and enum_zero: "enum ` (-S) \ {0}" - and abs_summable: "summable (\r. \enum r * measure M (f -` {enum r} \ space M)\)" - shows "integrable f" - and "integral f = (\r. enum r * measure M (f -` {enum r} \ space M))" (is "_ = suminf (?sum f enum)") -proof - - { fix f enum - assume borel: "f \ borel_measurable M" - and bij: "bij_betw enum S (f ` space M)" - and enum_zero: "enum ` (-S) \ {0}" - and abs_summable: "summable (\r. \enum r * measure M (f -` {enum r} \ space M)\)" - have inj_enum: "inj_on enum S" and range_enum: "f ` space M = enum ` S" - using bij unfolding bij_betw_def by auto - - have [simp, intro]: "\X. 0 \ measure M (f -` {X} \ space M)" - by (rule positive, rule borel_measurable_vimage[OF borel]) - - have "(\r. ?sum (pos_part f) (pos_part enum) r) \ nnfis (pos_part f) \ - summable (\r. ?sum (pos_part f) (pos_part enum) r)" - proof (rule conjI, rule nnfis_on_countable) - have pos_f_image: "pos_part f ` space M - {0} = f ` space M \ {0<..}" - unfolding pos_part_def max_def by auto - - show "bij_betw (pos_part enum) {x \ S. 0 < enum x} (pos_part f ` space M - {0})" - unfolding bij_betw_def pos_f_image - unfolding pos_part_def max_def range_enum - by (auto intro!: inj_onI simp: inj_enum[THEN inj_on_eq_iff]) - - show "\n. 0 \ pos_part enum n" unfolding pos_part_def max_def by auto - - show "pos_part f \ borel_measurable M" - by (rule pos_part_borel_measurable[OF borel]) - - show "pos_part enum ` (- {x \ S. 0 < enum x}) \ {0}" - unfolding pos_part_def max_def using enum_zero by auto - - show "summable (\r. ?sum (pos_part f) (pos_part enum) r)" - proof (rule summable_comparison_test[OF _ abs_summable], safe intro!: exI[of _ 0]) - fix n :: nat - have "pos_part enum n \ 0 \ (pos_part f -` {enum n} \ space M) = - (if 0 < enum n then (f -` {enum n} \ space M) else {})" - unfolding pos_part_def max_def by (auto split: split_if_asm) - thus "norm (?sum (pos_part f) (pos_part enum) n) \ \?sum f enum n \" - by (cases "pos_part enum n = 0", - auto simp: pos_part_def max_def abs_mult not_le split: split_if_asm intro!: mult_nonpos_nonneg) - qed - thus "(\r. ?sum (pos_part f) (pos_part enum) r) sums (\r. ?sum (pos_part f) (pos_part enum) r)" - by (rule summable_sums) - qed } - note pos = this - - note pos_part = pos[OF assms(1-4)] - moreover - have neg_part_to_pos_part: - "\f :: _ \ real. neg_part f = pos_part (uminus \ f)" - by (auto simp: pos_part_def neg_part_def min_def max_def expand_fun_eq) - - have neg_part: "(\r. ?sum (neg_part f) (neg_part enum) r) \ nnfis (neg_part f) \ - summable (\r. ?sum (neg_part f) (neg_part enum) r)" - unfolding neg_part_to_pos_part - proof (rule pos) - show "uminus \ f \ borel_measurable M" unfolding comp_def - by (rule borel_measurable_uminus_borel_measurable[OF borel]) - - show "bij_betw (uminus \ enum) S ((uminus \ f) ` space M)" - using bij unfolding bij_betw_def - by (auto intro!: comp_inj_on simp: image_compose) - - show "(uminus \ enum) ` (- S) \ {0}" - using enum_zero by auto - - have minus_image: "\r. (uminus \ f) -` {(uminus \ enum) r} \ space M = f -` {enum r} \ space M" - by auto - show "summable (\r. \(uminus \ enum) r * measure_space.measure M ((uminus \ f) -` {(uminus \ enum) r} \ space M)\)" - unfolding minus_image using abs_summable by simp - qed - ultimately show "integrable f" unfolding integrable_def by auto - - { fix r - have "?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r = ?sum f enum r" - proof (cases rule: linorder_cases) - assume "0 < enum r" - hence "pos_part f -` {enum r} \ space M = f -` {enum r} \ space M" - unfolding pos_part_def max_def by (auto split: split_if_asm) - with `0 < enum r` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq - by auto - next - assume "enum r < 0" - hence "neg_part f -` {- enum r} \ space M = f -` {enum r} \ space M" - unfolding neg_part_def min_def by (auto split: split_if_asm) - with `enum r < 0` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq - by auto - qed (simp add: neg_part_def pos_part_def) } - note sum_diff_eq_sum = this - - have "(\r. ?sum (pos_part f) (pos_part enum) r) - (\r. ?sum (neg_part f) (neg_part enum) r) - = (\r. ?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r)" - using neg_part pos_part by (auto intro: suminf_diff) - also have "... = (\r. ?sum f enum r)" unfolding sum_diff_eq_sum .. - finally show "integral f = suminf (?sum f enum)" - unfolding integral_def using pos_part neg_part - by (auto dest: the_nnfis) -qed - -section "Lebesgue integration on finite space" - -lemma integral_finite_on_sets: - assumes "f \ borel_measurable M" and "finite (space M)" and "a \ sets M" - shows "integral (\x. f x * indicator_fn a x) = - (\ r \ f`a. r * measure M (f -` {r} \ a))" (is "integral ?f = _") -proof - - { fix x assume "x \ a" - with assms have "f -` {f x} \ space M \ sets M" - by (subst Int_commute) - (auto simp: vimage_def Int_def - intro!: borel_measurable_const - borel_measurable_eq_borel_measurable) - from Int[OF this assms(3)] - sets_into_space[OF assms(3), THEN Int_absorb1] - have "f -` {f x} \ a \ sets M" by (simp add: Int_assoc) } - note vimage_f = this - - have "finite a" - using assms(2,3) sets_into_space - by (auto intro: finite_subset) - - have "integral (\x. f x * indicator_fn a x) = - integral (\x. \i\f ` a. i * indicator_fn (f -` {i} \ a) x)" - (is "_ = integral (\x. setsum (?f x) _)") - unfolding indicator_fn_def if_distrib - using `finite a` by (auto simp: setsum_cases intro!: integral_cong) - also have "\ = (\i\f`a. integral (\x. ?f x i))" - proof (rule integral_setsum, safe) - fix n x assume "x \ a" - thus "integrable (\y. ?f y (f x))" - using integral_indicator_fn(2)[OF vimage_f] - by (auto intro!: integral_times_const) - qed (simp add: `finite a`) - also have "\ = (\i\f`a. i * measure M (f -` {i} \ a))" - using integral_cmul_indicator[OF vimage_f] - by (auto intro!: setsum_cong) - finally show ?thesis . -qed - -lemma integral_finite: - assumes "f \ borel_measurable M" and "finite (space M)" - shows "integral f = (\ r \ f ` space M. r * measure M (f -` {r} \ space M))" - using integral_finite_on_sets[OF assms top] - integral_cong[of "\x. f x * indicator_fn (space M) x" f] - by (auto simp add: indicator_fn_def) - -section "Radon–Nikodym derivative" - -definition - "RN_deriv v \ SOME f. measure_space (M\measure := v\) \ - f \ borel_measurable M \ - (\a \ sets M. (integral (\x. f x * indicator_fn a x) = v a))" - -end - -lemma sigma_algebra_cong: - fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme" - assumes *: "sigma_algebra M" - and cong: "space M = space M'" "sets M = sets M'" - shows "sigma_algebra M'" -using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong . - -lemma finite_Pow_additivity_sufficient: - assumes "finite (space M)" and "sets M = Pow (space M)" - and "positive M (measure M)" and "additive M (measure M)" - shows "finite_measure_space M" -proof - - have "sigma_algebra M" - using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow]) - - have "measure_space M" - by (rule Measure.finite_additivity_sufficient) (fact+) - thus ?thesis - unfolding finite_measure_space_def finite_measure_space_axioms_def - using assms by simp -qed - -lemma finite_measure_spaceI: - assumes "measure_space M" and "finite (space M)" and "sets M = Pow (space M)" - shows "finite_measure_space M" - unfolding finite_measure_space_def finite_measure_space_axioms_def - using assms by simp - -lemma (in finite_measure_space) integral_finite_singleton: - "integral f = (\x \ space M. f x * measure M {x})" -proof - - have "f \ borel_measurable M" - unfolding borel_measurable_le_iff - using sets_eq_Pow by auto - { fix r let ?x = "f -` {r} \ space M" - have "?x \ space M" by auto - with finite_space sets_eq_Pow have "measure M ?x = (\i \ ?x. measure M {i})" - by (auto intro!: measure_real_sum_image) } - note measure_eq_setsum = this - show ?thesis - unfolding integral_finite[OF `f \ borel_measurable M` finite_space] - measure_eq_setsum setsum_right_distrib - apply (subst setsum_Sigma) - apply (simp add: finite_space) - apply (simp add: finite_space) - proof (rule setsum_reindex_cong[symmetric]) - fix a assume "a \ Sigma (f ` space M) (\x. f -` {x} \ space M)" - thus "(\(x, y). x * measure M {y}) a = f (snd a) * measure_space.measure M {snd a}" - by auto - qed (auto intro!: image_eqI inj_onI) -qed - -lemma (in finite_measure_space) RN_deriv_finite_singleton: - fixes v :: "'a set \ real" - assumes ms_v: "measure_space (M\measure := v\)" - and eq_0: "\x. \ x \ space M ; measure M {x} = 0 \ \ v {x} = 0" - and "x \ space M" and "measure M {x} \ 0" - shows "RN_deriv v x = v {x} / (measure M {x})" (is "_ = ?v x") - unfolding RN_deriv_def -proof (rule someI2_ex[where Q = "\f. f x = ?v x"], rule exI[where x = ?v], safe) - show "(\a. v {a} / measure_space.measure M {a}) \ borel_measurable M" - unfolding borel_measurable_le_iff using sets_eq_Pow by auto -next - fix a assume "a \ sets M" - hence "a \ space M" and "finite a" - using sets_into_space finite_space by (auto intro: finite_subset) - have *: "\x a. x \ space M \ (if measure M {x} = 0 then 0 else v {x} * indicator_fn a x) = - v {x} * indicator_fn a x" using eq_0 by auto - - from measure_space.measure_real_sum_image[OF ms_v, of a] - sets_eq_Pow `a \ sets M` sets_into_space `finite a` - have "v a = (\x\a. v {x})" by auto - thus "integral (\x. v {x} / measure_space.measure M {x} * indicator_fn a x) = v a" - apply (simp add: eq_0 integral_finite_singleton) - apply (unfold divide_1) - by (simp add: * indicator_fn_def if_distrib setsum_cases finite_space `a \ space M` Int_absorb1) -next - fix w assume "w \ borel_measurable M" - assume int_eq_v: "\a\sets M. integral (\x. w x * indicator_fn a x) = v a" - have "{x} \ sets M" using sets_eq_Pow `x \ space M` by auto - - have "w x * measure M {x} = - (\y\space M. w y * indicator_fn {x} y * measure M {y})" - apply (subst (3) mult_commute) - unfolding indicator_fn_def if_distrib setsum_cases[OF finite_space] - using `x \ space M` by simp - also have "... = v {x}" - using int_eq_v[rule_format, OF `{x} \ sets M`] - by (simp add: integral_finite_singleton) - finally show "w x = v {x} / measure M {x}" - using `measure M {x} \ 0` by (simp add: eq_divide_eq) -qed fact - -end