diff -r 653e56c6c963 -r f174712d0a84 src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Fri May 30 18:13:40 2014 +0200 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Fri May 30 15:56:30 2014 +0200 @@ -1733,32 +1733,102 @@ subsubsection {* Measures with Restricted Space *} +lemma simple_function_iff_borel_measurable: + fixes f :: "'a \ 'x::{t2_space}" + shows "simple_function M f \ finite (f ` space M) \ f \ borel_measurable M" + by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable) + +lemma simple_function_restrict_space_ereal: + fixes f :: "'a \ ereal" + assumes "\ \ space M \ sets M" + shows "simple_function (restrict_space M \) f \ simple_function M (\x. f x * indicator \ x)" +proof - + { assume "finite (f ` space (restrict_space M \))" + then have "finite (f ` space (restrict_space M \) \ {0})" by simp + then have "finite ((\x. f x * indicator \ x) ` space M)" + by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } + moreover + { assume "finite ((\x. f x * indicator \ x) ` space M)" + then have "finite (f ` space (restrict_space M \))" + by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } + ultimately show ?thesis + unfolding simple_function_iff_borel_measurable + borel_measurable_restrict_space_iff_ereal[OF assms] + by auto +qed + +lemma simple_function_restrict_space: + fixes f :: "'a \ 'b::real_normed_vector" + assumes "\ \ space M \ sets M" + shows "simple_function (restrict_space M \) f \ simple_function M (\x. indicator \ x *\<^sub>R f x)" +proof - + { assume "finite (f ` space (restrict_space M \))" + then have "finite (f ` space (restrict_space M \) \ {0})" by simp + then have "finite ((\x. indicator \ x *\<^sub>R f x) ` space M)" + by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } + moreover + { assume "finite ((\x. indicator \ x *\<^sub>R f x) ` space M)" + then have "finite (f ` space (restrict_space M \))" + by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } + ultimately show ?thesis + unfolding simple_function_iff_borel_measurable + borel_measurable_restrict_space_iff[OF assms] + by auto +qed + + +lemma split_indicator_asm: "P (indicator Q x) = (\ (x \ Q \ \ P 1 \ x \ Q \ \ P 0))" + by (auto split: split_indicator) + +lemma simple_integral_restrict_space: + assumes \: "\ \ space M \ sets M" "simple_function (restrict_space M \) f" + shows "simple_integral (restrict_space M \) f = simple_integral M (\x. f x * indicator \ x)" + using simple_function_restrict_space_ereal[THEN iffD1, OF \, THEN simple_functionD(1)] + by (auto simp add: space_restrict_space emeasure_restrict_space[OF \(1)] le_infI2 simple_integral_def + split: split_indicator split_indicator_asm + intro!: setsum_mono_zero_cong_left ereal_left_mult_cong arg_cong2[where f=emeasure]) + +lemma one_not_less_zero[simp]: "\ 1 < (0::ereal)" + by (simp add: zero_ereal_def one_ereal_def) + lemma nn_integral_restrict_space: - assumes \: "\ \ sets M" and f: "f \ borel_measurable M" "\x. 0 \ f x" "\x. x \ space M - \ \ f x = 0" - shows "nn_integral (restrict_space M \) f = nn_integral M f" -using f proof (induct rule: borel_measurable_induct) - case (cong f g) then show ?case - using nn_integral_cong[of M f g] nn_integral_cong[of "restrict_space M \" f g] - sets.sets_into_space[OF `\ \ sets M`] - by (simp add: subset_eq space_restrict_space) -next - case (set A) - then have "A \ \" - unfolding indicator_eq_0_iff by (auto dest: sets.sets_into_space) - with set `\ \ sets M` sets.sets_into_space[OF `\ \ sets M`] show ?case - by (subst nn_integral_indicator') - (auto simp add: sets_restrict_space_iff space_restrict_space - emeasure_restrict_space Int_absorb2 - dest: sets.sets_into_space) -next - case (mult f c) then show ?case - by (cases "c = 0") (simp_all add: measurable_restrict_space1 \ nn_integral_cmult) -next - case (add f g) then show ?case - by (simp add: measurable_restrict_space1 \ nn_integral_add ereal_add_nonneg_eq_0_iff) -next - case (seq F) then show ?case - by (auto simp add: SUP_eq_iff measurable_restrict_space1 \ nn_integral_monotone_convergence_SUP) + assumes \[simp]: "\ \ space M \ sets M" + shows "nn_integral (restrict_space M \) f = nn_integral M (\x. f x * indicator \ x)" +proof - + let ?R = "restrict_space M \" and ?X = "\M f. {s. simple_function M s \ s \ max 0 \ f \ range s \ {0 ..< \}}" + have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\x. f x * indicator \ x)" + proof (safe intro!: image_eqI) + fix s assume s: "simple_function ?R s" "s \ max 0 \ f" "range s \ {0..<\}" + from s show "integral\<^sup>S (restrict_space M \) s = integral\<^sup>S M (\x. s x * indicator \ x)" + by (intro simple_integral_restrict_space) auto + from s show "simple_function M (\x. s x * indicator \ x)" + by (simp add: simple_function_restrict_space_ereal) + from s show "(\x. s x * indicator \ x) \ max 0 \ (\x. f x * indicator \ x)" + "\x. s x * indicator \ x \ {0..<\}" + by (auto split: split_indicator simp: le_fun_def image_subset_iff) + next + fix s assume s: "simple_function M s" "s \ max 0 \ (\x. f x * indicator \ x)" "range s \ {0..<\}" + then have "simple_function M (\x. s x * indicator (\ \ space M) x)" (is ?s') + by (intro simple_function_mult simple_function_indicator) auto + also have "?s' \ simple_function M (\x. s x * indicator \ x)" + by (rule simple_function_cong) (auto split: split_indicator) + finally show sf: "simple_function (restrict_space M \) s" + by (simp add: simple_function_restrict_space_ereal) + + from s have s_eq: "s = (\x. s x * indicator \ x)" + by (auto simp add: fun_eq_iff le_fun_def image_subset_iff + split: split_indicator split_indicator_asm + intro: antisym) + + show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \) s" + by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \ sf]) + show "\x. s x \ {0..<\}" + using s by (auto simp: image_subset_iff) + from s show "s \ max 0 \ f" + by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) + qed + then show ?thesis + unfolding nn_integral_def_finite SUP_def by simp qed subsubsection {* Measure spaces with an associated density *}