diff -r 261d42f0bfac -r 979cdfdf7a79 src/HOL/Analysis/Set_Integral.thy --- a/src/HOL/Analysis/Set_Integral.thy Mon Oct 17 15:20:06 2016 +0200 +++ b/src/HOL/Analysis/Set_Integral.thy Thu Oct 13 18:36:06 2016 +0200 @@ -1,5 +1,6 @@ (* Title: HOL/Analysis/Set_Integral.thy Author: Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU) + Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr Notation and useful facts for working with integrals over a set. @@ -7,7 +8,7 @@ *) theory Set_Integral - imports Bochner_Integration + imports Radon_Nikodym begin (* @@ -523,5 +524,235 @@ shows "set_borel_measurable borel {a..b} f" by (rule set_borel_measurable_continuous[OF _ assms]) simp + +text{* This notation is from Sébastien Gouëzel: His use is not directly in line with the +notations in this file, they are more in line with sum, and more readable he thinks. *} + +abbreviation "set_nn_integral M A f \ nn_integral M (\x. f x * indicator A x)" + +syntax +"_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" +("(\\<^sup>+((_)\(_)./ _)/\_)" [0,60,110,61] 60) + +"_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" +("(\((_)\(_)./ _)/\_)" [0,60,110,61] 60) + +translations +"\\<^sup>+x \ A. f \M" == "CONST set_nn_integral M A (\x. f)" +"\x \ A. f \M" == "CONST set_lebesgue_integral M A (\x. f)" + +lemma nn_integral_disjoint_pair: + assumes [measurable]: "f \ borel_measurable M" + "B \ sets M" "C \ sets M" + "B \ C = {}" + shows "(\\<^sup>+x \ B \ C. f x \M) = (\\<^sup>+x \ B. f x \M) + (\\<^sup>+x \ C. f x \M)" +proof - + have mes: "\D. D \ sets M \ (\x. f x * indicator D x) \ borel_measurable M" by simp + have pos: "\D. AE x in M. f x * indicator D x \ 0" using assms(2) by auto + have "\x. f x * indicator (B \ C) x = f x * indicator B x + f x * indicator C x" using assms(4) + by (auto split: split_indicator) + then have "(\\<^sup>+x. f x * indicator (B \ C) x \M) = (\\<^sup>+x. f x * indicator B x + f x * indicator C x \M)" + by simp + also have "... = (\\<^sup>+x. f x * indicator B x \M) + (\\<^sup>+x. f x * indicator C x \M)" + by (rule nn_integral_add) (auto simp add: assms mes pos) + finally show ?thesis by simp +qed + +lemma nn_integral_disjoint_pair_countspace: + assumes "B \ C = {}" + shows "(\\<^sup>+x \ B \ C. f x \count_space UNIV) = (\\<^sup>+x \ B. f x \count_space UNIV) + (\\<^sup>+x \ C. f x \count_space UNIV)" +by (rule nn_integral_disjoint_pair) (simp_all add: assms) + +lemma nn_integral_null_delta: + assumes "A \ sets M" "B \ sets M" + "(A - B) \ (B - A) \ null_sets M" + shows "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ B. f x \M)" +proof (rule nn_integral_cong_AE, auto simp add: indicator_def) + have *: "AE a in M. a \ (A - B) \ (B - A)" + using assms(3) AE_not_in by blast + then show "AE a in M. a \ A \ a \ B \ f a = 0" + by auto + show "AE x\A in M. x \ B \ f x = 0" + using * by auto +qed + +lemma nn_integral_disjoint_family: + assumes [measurable]: "f \ borel_measurable M" "\(n::nat). B n \ sets M" + and "disjoint_family B" + shows "(\\<^sup>+x \ (\n. B n). f x \M) = (\n. (\\<^sup>+x \ B n. f x \M))" +proof - + have "(\\<^sup>+ x. (\n. f x * indicator (B n) x) \M) = (\n. (\\<^sup>+ x. f x * indicator (B n) x \M))" + by (rule nn_integral_suminf) simp + moreover have "(\n. f x * indicator (B n) x) = f x * indicator (\n. B n) x" for x + proof (cases) + assume "x \ (\n. B n)" + then obtain n where "x \ B n" by blast + have a: "finite {n}" by simp + have "\i. i \ n \ x \ B i" using `x \ B n` assms(3) disjoint_family_on_def + by (metis IntI UNIV_I empty_iff) + then have "\i. i \ {n} \ indicator (B i) x = (0::ennreal)" using indicator_def by simp + then have b: "\i. i \ {n} \ f x * indicator (B i) x = (0::ennreal)" by simp + + define h where "h = (\i. f x * indicator (B i) x)" + then have "\i. i \ {n} \ h i = 0" using b by simp + then have "(\i. h i) = (\i\{n}. h i)" + by (metis sums_unique[OF sums_finite[OF a]]) + then have "(\i. h i) = h n" by simp + then have "(\n. f x * indicator (B n) x) = f x * indicator (B n) x" using h_def by simp + then have "(\n. f x * indicator (B n) x) = f x" using `x \ B n` indicator_def by simp + then show ?thesis using `x \ (\n. B n)` by auto + next + assume "x \ (\n. B n)" + then have "\n. f x * indicator (B n) x = 0" by simp + have "(\n. f x * indicator (B n) x) = 0" + by (simp add: `\n. f x * indicator (B n) x = 0`) + then show ?thesis using `x \ (\n. B n)` by auto + qed + ultimately show ?thesis by simp +qed + +lemma nn_set_integral_add: + assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" + "A \ sets M" + shows "(\\<^sup>+x \ A. (f x + g x) \M) = (\\<^sup>+x \ A. f x \M) + (\\<^sup>+x \ A. g x \M)" +proof - + have "(\\<^sup>+x \ A. (f x + g x) \M) = (\\<^sup>+x. (f x * indicator A x + g x * indicator A x) \M)" + by (auto simp add: indicator_def intro!: nn_integral_cong) + also have "... = (\\<^sup>+x. f x * indicator A x \M) + (\\<^sup>+x. g x * indicator A x \M)" + apply (rule nn_integral_add) using assms(1) assms(2) by auto + finally show ?thesis by simp +qed + +lemma nn_set_integral_cong: + assumes "AE x in M. f x = g x" + shows "(\\<^sup>+x \ A. f x \M) = (\\<^sup>+x \ A. g x \M)" +apply (rule nn_integral_cong_AE) using assms(1) by auto + +lemma nn_set_integral_set_mono: + "A \ B \ (\\<^sup>+ x \ A. f x \M) \ (\\<^sup>+ x \ B. f x \M)" +by (auto intro!: nn_integral_mono split: split_indicator) + +lemma nn_set_integral_mono: + assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" + "A \ sets M" + and "AE x\A in M. f x \ g x" + shows "(\\<^sup>+x \ A. f x \M) \ (\\<^sup>+x \ A. g x \M)" +by (auto intro!: nn_integral_mono_AE split: split_indicator simp: assms) + +lemma nn_set_integral_space [simp]: + shows "(\\<^sup>+ x \ space M. f x \M) = (\\<^sup>+x. f x \M)" +by (metis (mono_tags, lifting) indicator_simps(1) mult.right_neutral nn_integral_cong) + +lemma nn_integral_count_compose_inj: + assumes "inj_on g A" + shows "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ g`A. f y \count_space UNIV)" +proof - + have "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+x. f (g x) \count_space A)" + by (auto simp add: nn_integral_count_space_indicator[symmetric]) + also have "... = (\\<^sup>+y. f y \count_space (g`A))" + by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw) + also have "... = (\\<^sup>+y \ g`A. f y \count_space UNIV)" + by (auto simp add: nn_integral_count_space_indicator[symmetric]) + finally show ?thesis by simp +qed + +lemma nn_integral_count_compose_bij: + assumes "bij_betw g A B" + shows "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ B. f y \count_space UNIV)" +proof - + have "inj_on g A" using assms bij_betw_def by auto + then have "(\\<^sup>+x \ A. f (g x) \count_space UNIV) = (\\<^sup>+y \ g`A. f y \count_space UNIV)" + by (rule nn_integral_count_compose_inj) + then show ?thesis using assms by (simp add: bij_betw_def) +qed + +lemma set_integral_null_delta: + fixes f::"_ \ _ :: {banach, second_countable_topology}" + assumes [measurable]: "integrable M f" "A \ sets M" "B \ sets M" + and "(A - B) \ (B - A) \ null_sets M" + shows "(\x \ A. f x \M) = (\x \ B. f x \M)" +proof (rule set_integral_cong_set, auto) + have *: "AE a in M. a \ (A - B) \ (B - A)" + using assms(4) AE_not_in by blast + then show "AE x in M. (x \ B) = (x \ A)" + by auto +qed + +lemma set_integral_space: + assumes "integrable M f" + shows "(\x \ space M. f x \M) = (\x. f x \M)" +by (metis (mono_tags, lifting) indicator_simps(1) Bochner_Integration.integral_cong real_vector.scale_one) + +lemma null_if_pos_func_has_zero_nn_int: + fixes f::"'a \ ennreal" + assumes [measurable]: "f \ borel_measurable M" "A \ sets M" + and "AE x\A in M. f x > 0" "(\\<^sup>+x\A. f x \M) = 0" + shows "A \ null_sets M" +proof - + have "AE x in M. f x * indicator A x = 0" + by (subst nn_integral_0_iff_AE[symmetric], auto simp add: assms(4)) + then have "AE x\A in M. False" using assms(3) by auto + then show "A \ null_sets M" using assms(2) by (simp add: AE_iff_null_sets) +qed + +lemma null_if_pos_func_has_zero_int: + assumes [measurable]: "integrable M f" "A \ sets M" + and "AE x\A in M. f x > 0" "(\x\A. f x \M) = (0::real)" + shows "A \ null_sets M" +proof - + have "AE x in M. indicator A x * f x = 0" + apply (subst integral_nonneg_eq_0_iff_AE[symmetric]) + using assms integrable_mult_indicator[OF `A \ sets M` assms(1)] by auto + then have "AE x\A in M. f x = 0" by auto + then have "AE x\A in M. False" using assms(3) by auto + then show "A \ null_sets M" using assms(2) by (simp add: AE_iff_null_sets) +qed + +text{*The next lemma is a variant of \density_unique\. Note that it uses the notation +for nonnegative set integrals introduced earlier.*} + +lemma (in sigma_finite_measure) density_unique2: + assumes [measurable]: "f \ borel_measurable M" "f' \ borel_measurable M" + assumes density_eq: "\A. A \ sets M \ (\\<^sup>+ x \ A. f x \M) = (\\<^sup>+ x \ A. f' x \M)" + shows "AE x in M. f x = f' x" +proof (rule density_unique) + show "density M f = density M f'" + by (intro measure_eqI) (auto simp: emeasure_density intro!: density_eq) +qed (auto simp add: assms) + + +text {* The next lemma implies the same statement for Banach-space valued functions +using Hahn-Banach theorem and linear forms. Since they are not yet easily available, I +only formulate it for real-valued functions.*} + +lemma density_unique_real: + fixes f f'::"_ \ real" + assumes [measurable]: "integrable M f" "integrable M f'" + assumes density_eq: "\A. A \ sets M \ (\x \ A. f x \M) = (\x \ A. f' x \M)" + shows "AE x in M. f x = f' x" +proof - + define A where "A = {x \ space M. f x < f' x}" + then have [measurable]: "A \ sets M" by simp + have "(\x \ A. (f' x - f x) \M) = (\x \ A. f' x \M) - (\x \ A. f x \M)" + using `A \ sets M` assms(1) assms(2) integrable_mult_indicator by blast + then have "(\x \ A. (f' x - f x) \M) = 0" using assms(3) by simp + then have "A \ null_sets M" + using A_def null_if_pos_func_has_zero_int[where ?f = "\x. f' x - f x" and ?A = A] assms by auto + then have "AE x in M. x \ A" by (simp add: AE_not_in) + then have *: "AE x in M. f' x \ f x" unfolding A_def by auto + + + define B where "B = {x \ space M. f' x < f x}" + then have [measurable]: "B \ sets M" by simp + have "(\x \ B. (f x - f' x) \M) = (\x \ B. f x \M) - (\x \ B. f' x \M)" + using `B \ sets M` assms(1) assms(2) integrable_mult_indicator by blast + then have "(\x \ B. (f x - f' x) \M) = 0" using assms(3) by simp + then have "B \ null_sets M" + using B_def null_if_pos_func_has_zero_int[where ?f = "\x. f x - f' x" and ?A = B] assms by auto + then have "AE x in M. x \ B" by (simp add: AE_not_in) + then have "AE x in M. f' x \ f x" unfolding B_def by auto + + then show ?thesis using * by auto +qed + end -