# HG changeset patch # User hoelzl # Date 1400493885 -7200 # Node ID e5366291d6aaa0cf08f1e2ec23f0b9e9ae32889f # Parent d0e5225601d3c6daee52537fa12e8f7e65b4bbe7 introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces diff -r d0e5225601d3 -r e5366291d6aa NEWS --- a/NEWS Mon May 19 11:27:02 2014 +0200 +++ b/NEWS Mon May 19 12:04:45 2014 +0200 @@ -670,6 +670,68 @@ derivative_is_linear ~> has_derivative_linear bounded_linear_imp_linear ~> bounded_linear.linear +* HOL-Probability: + - replaced the Lebesgue integral on real numbers by the more general Bochner + integral for functions into a real-normed vector space. + + integral_zero ~> integral_zero / integrable_zero + integral_minus ~> integral_minus / integrable_minus + integral_add ~> integral_add / integrable_add + integral_diff ~> integral_diff / integrable_diff + integral_setsum ~> integral_setsum / integrable_setsum + integral_multc ~> integral_mult_left / integrable_mult_left + integral_cmult ~> integral_mult_right / integrable_mult_right + integral_triangle_inequality~> integral_norm_bound + integrable_nonneg ~> integrableI_nonneg + integral_positive ~> integral_nonneg_AE + integrable_abs_iff ~> integrable_abs_cancel + positive_integral_lim_INF ~> positive_integral_liminf + lebesgue_real_affine ~> lborel_real_affine + borel_integral_has_integral ~> has_integral_lebesgue_integral + integral_indicator ~> + integral_real_indicator / integrable_real_indicator + positive_integral_fst ~> positive_integral_fst' + positive_integral_fst_measurable ~> positive_integral_fst + positive_integral_snd_measurable ~> positive_integral_snd + + integrable_fst_measurable ~> + integral_fst / integrable_fst / AE_integrable_fst + + integrable_snd_measurable ~> + integral_snd / integrable_snd / AE_integrable_snd + + integral_monotone_convergence ~> + integral_monotone_convergence / integrable_monotone_convergence + + integral_monotone_convergence_at_top ~> + integral_monotone_convergence_at_top / + integrable_monotone_convergence_at_top + + has_integral_iff_positive_integral_lebesgue ~> + has_integral_iff_has_bochner_integral_lebesgue_nonneg + + lebesgue_integral_has_integral ~> + has_integral_integrable_lebesgue_nonneg + + positive_integral_lebesgue_has_integral ~> + integral_has_integral_lebesgue_nonneg / + integrable_has_integral_lebesgue_nonneg + + lebesgue_integral_real_affine ~> + positive_integral_real_affine + + has_integral_iff_positive_integral_lborel ~> + integral_has_integral_nonneg / integrable_has_integral_nonneg + + The following theorems where removed: + + lebesgue_integral_nonneg + lebesgue_integral_uminus + lebesgue_integral_cmult + lebesgue_integral_multc + lebesgue_integral_cmult_nonneg + integral_cmul_indicator + integral_real *** Scala *** diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Lattices_Big.thy --- a/src/HOL/Lattices_Big.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Lattices_Big.thy Mon May 19 12:04:45 2014 +0200 @@ -125,6 +125,9 @@ finally show ?case . qed +lemma infinite: "\ finite A \ F A = the None" + unfolding eq_fold' by (cases "finite (UNIV::'a set)") (auto intro: finite_subset fold_infinite) + end locale semilattice_order_set = binary?: semilattice_order + semilattice_set diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Library/Extended_Real.thy --- a/src/HOL/Library/Extended_Real.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Library/Extended_Real.thy Mon May 19 12:04:45 2014 +0200 @@ -453,7 +453,7 @@ lemma ereal_add_strict_mono: fixes a b c d :: ereal - assumes "a = b" + assumes "a \ b" and "0 \ a" and "a \ \" and "c < d" @@ -2022,6 +2022,22 @@ by auto qed (auto simp add: image_Union image_Int) + +lemma eventually_finite: + fixes x :: ereal + assumes "\x\ \ \" "(f ---> x) F" + shows "eventually (\x. \f x\ \ \) F" +proof - + have "(f ---> ereal (real x)) F" + using assms by (cases x) auto + then have "eventually (\x. f x \ ereal ` UNIV) F" + by (rule topological_tendstoD) (auto intro: open_ereal) + also have "(\x. f x \ ereal ` UNIV) = (\x. \f x\ \ \)" + by auto + finally show ?thesis . +qed + + lemma open_ereal_def: "open A \ open (ereal -` A) \ (\ \ A \ (\x. {ereal x <..} \ A)) \ (-\ \ A \ (\x. {.. A))" (is "open A \ ?rhs") @@ -2265,6 +2281,27 @@ shows "a * b = a * c \ (\a\ = \ \ 0 < b * c) \ a = 0 \ b = c" by (cases rule: ereal3_cases[of a b c]) (simp_all add: zero_less_mult_iff) +lemma tendsto_add_ereal: + fixes x y :: ereal + assumes x: "\x\ \ \" and y: "\y\ \ \" + assumes f: "(f ---> x) F" and g: "(g ---> y) F" + shows "((\x. f x + g x) ---> x + y) F" +proof - + from x obtain r where x': "x = ereal r" by (cases x) auto + with f have "((\i. real (f i)) ---> r) F" by simp + moreover + from y obtain p where y': "y = ereal p" by (cases y) auto + with g have "((\i. real (g i)) ---> p) F" by simp + ultimately have "((\i. real (f i) + real (g i)) ---> r + p) F" + by (rule tendsto_add) + moreover + from eventually_finite[OF x f] eventually_finite[OF y g] + have "eventually (\x. f x + g x = ereal (real (f x) + real (g x))) F" + by eventually_elim auto + ultimately show ?thesis + by (simp add: x' y' cong: filterlim_cong) +qed + lemma ereal_inj_affinity: fixes m t :: ereal assumes "\m\ \ \" diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Library/Indicator_Function.thy --- a/src/HOL/Library/Indicator_Function.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Library/Indicator_Function.thy Mon May 19 12:04:45 2014 +0200 @@ -5,7 +5,7 @@ header {* Indicator Function *} theory Indicator_Function -imports Main +imports Complex_Main begin definition "indicator S x = (if x \ S then 1 else 0)" @@ -65,4 +65,9 @@ using setsum_mult_indicator[OF assms, of "%x. 1::nat"] unfolding card_eq_setsum by simp +lemma setsum_indicator_scaleR[simp]: + "finite A \ + (\x \ A. indicator (B x) (g x) *\<^sub>R f x) = (\x \ {x\A. g x \ B x}. f x::'a::real_vector)" + using assms by (auto intro!: setsum_mono_zero_cong_right split: split_if_asm simp: indicator_def) + end \ No newline at end of file diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Mon May 19 12:04:45 2014 +0200 @@ -5,7 +5,7 @@ header {*Binary product measures*} theory Binary_Product_Measure -imports Lebesgue_Integration +imports Nonnegative_Lebesgue_Integration begin lemma Pair_vimage_times[simp]: "Pair x -` (A \ B) = (if x \ A then B else {})" @@ -510,7 +510,7 @@ positive_integral_monotone_convergence_SUP incseq_def le_fun_def cong: measurable_cong) -lemma (in sigma_finite_measure) positive_integral_fst: +lemma (in sigma_finite_measure) positive_integral_fst': assumes f: "f \ borel_measurable (M1 \\<^sub>M M)" "\x. 0 \ f x" shows "(\\<^sup>+ x. \\<^sup>+ y. f (x, y) \M \M1) = integral\<^sup>P (M1 \\<^sub>M M) f" (is "?I f = _") using f proof induct @@ -525,30 +525,25 @@ borel_measurable_positive_integral_fst' positive_integral_mono incseq_def le_fun_def cong: positive_integral_cong) -lemma (in sigma_finite_measure) positive_integral_fst_measurable: +lemma (in sigma_finite_measure) positive_integral_fst: assumes f: "f \ borel_measurable (M1 \\<^sub>M M)" - shows "(\x. \\<^sup>+ y. f (x, y) \M) \ borel_measurable M1" - (is "?C f \ borel_measurable M1") - and "(\\<^sup>+ x. (\\<^sup>+ y. f (x, y) \M) \M1) = integral\<^sup>P (M1 \\<^sub>M M) f" + shows "(\\<^sup>+ x. (\\<^sup>+ y. f (x, y) \M) \M1) = integral\<^sup>P (M1 \\<^sub>M M) f" using f borel_measurable_positive_integral_fst'[of "\x. max 0 (f x)"] - positive_integral_fst[of "\x. max 0 (f x)"] + positive_integral_fst'[of "\x. max 0 (f x)"] unfolding positive_integral_max_0 by auto lemma (in sigma_finite_measure) borel_measurable_positive_integral[measurable (raw)]: "split f \ borel_measurable (N \\<^sub>M M) \ (\x. \\<^sup>+ y. f x y \M) \ borel_measurable N" - using positive_integral_fst_measurable(1)[of "split f" N] by simp + using borel_measurable_positive_integral_fst'[of "\x. max 0 (split f x)" N] + by (simp add: positive_integral_max_0) -lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]: - "split f \ borel_measurable (N \\<^sub>M M) \ (\x. \ y. f x y \M) \ borel_measurable N" - by (simp add: lebesgue_integral_def) - -lemma (in pair_sigma_finite) positive_integral_snd_measurable: +lemma (in pair_sigma_finite) positive_integral_snd: assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = integral\<^sup>P (M1 \\<^sub>M M2) f" proof - note measurable_pair_swap[OF f] - from M1.positive_integral_fst_measurable[OF this] + from M1.positive_integral_fst[OF this] have "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = (\\<^sup>+ (x, y). f (y, x) \(M2 \\<^sub>M M1))" by simp also have "(\\<^sup>+ (x, y). f (y, x) \(M2 \\<^sub>M M1)) = integral\<^sup>P (M1 \\<^sub>M M2) f" @@ -560,113 +555,7 @@ lemma (in pair_sigma_finite) Fubini: assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" shows "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = (\\<^sup>+ x. (\\<^sup>+ y. f (x, y) \M2) \M1)" - unfolding positive_integral_snd_measurable[OF assms] - unfolding M2.positive_integral_fst_measurable[OF assms] .. - -lemma (in pair_sigma_finite) integrable_product_swap: - assumes "integrable (M1 \\<^sub>M M2) f" - shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) - show ?thesis unfolding * - by (rule integrable_distr[OF measurable_pair_swap']) - (simp add: distr_pair_swap[symmetric] assms) -qed - -lemma (in pair_sigma_finite) integrable_product_swap_iff: - "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^sub>M M2) f" -proof - - interpret Q: pair_sigma_finite M2 M1 by default - from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] - show ?thesis by auto -qed - -lemma (in pair_sigma_finite) integral_product_swap: - assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" - shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" -proof - - have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) - show ?thesis unfolding * - by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) -qed - -lemma (in pair_sigma_finite) integrable_fst_measurable: - assumes f: "integrable (M1 \\<^sub>M M2) f" - shows "AE x in M1. integrable M2 (\ y. f (x, y))" (is "?AE") - and "(\x. (\y. f (x, y) \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) f" (is "?INT") -proof - - have f_borel: "f \ borel_measurable (M1 \\<^sub>M M2)" - using f by auto - let ?pf = "\x. ereal (f x)" and ?nf = "\x. ereal (- f x)" - have - borel: "?nf \ borel_measurable (M1 \\<^sub>M M2)""?pf \ borel_measurable (M1 \\<^sub>M M2)" and - int: "integral\<^sup>P (M1 \\<^sub>M M2) ?nf \ \" "integral\<^sup>P (M1 \\<^sub>M M2) ?pf \ \" - using assms by auto - have "(\\<^sup>+x. (\\<^sup>+y. ereal (f (x, y)) \M2) \M1) \ \" - "(\\<^sup>+x. (\\<^sup>+y. ereal (- f (x, y)) \M2) \M1) \ \" - using borel[THEN M2.positive_integral_fst_measurable(1)] int - unfolding borel[THEN M2.positive_integral_fst_measurable(2)] by simp_all - with borel[THEN M2.positive_integral_fst_measurable(1)] - have AE_pos: "AE x in M1. (\\<^sup>+y. ereal (f (x, y)) \M2) \ \" - "AE x in M1. (\\<^sup>+y. ereal (- f (x, y)) \M2) \ \" - by (auto intro!: positive_integral_PInf_AE ) - then have AE: "AE x in M1. \\\<^sup>+y. ereal (f (x, y)) \M2\ \ \" - "AE x in M1. \\\<^sup>+y. ereal (- f (x, y)) \M2\ \ \" - by (auto simp: positive_integral_positive) - from AE_pos show ?AE using assms - by (simp add: measurable_Pair2[OF f_borel] integrable_def) - { fix f have "(\\<^sup>+ x. - \\<^sup>+ y. ereal (f x y) \M2 \M1) = (\\<^sup>+x. 0 \M1)" - using positive_integral_positive - by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder) - then have "(\\<^sup>+ x. - \\<^sup>+ y. ereal (f x y) \M2 \M1) = 0" by simp } - note this[simp] - { fix f assume borel: "(\x. ereal (f x)) \ borel_measurable (M1 \\<^sub>M M2)" - and int: "integral\<^sup>P (M1 \\<^sub>M M2) (\x. ereal (f x)) \ \" - and AE: "AE x in M1. (\\<^sup>+y. ereal (f (x, y)) \M2) \ \" - have "integrable M1 (\x. real (\\<^sup>+y. ereal (f (x, y)) \M2))" (is "integrable M1 ?f") - proof (intro integrable_def[THEN iffD2] conjI) - show "?f \ borel_measurable M1" - using borel by (auto intro!: M2.positive_integral_fst_measurable) - have "(\\<^sup>+x. ereal (?f x) \M1) = (\\<^sup>+x. (\\<^sup>+y. ereal (f (x, y)) \M2) \M1)" - using AE positive_integral_positive[of M2] - by (auto intro!: positive_integral_cong_AE simp: ereal_real) - then show "(\\<^sup>+x. ereal (?f x) \M1) \ \" - using M2.positive_integral_fst_measurable[OF borel] int by simp - have "(\\<^sup>+x. ereal (- ?f x) \M1) = (\\<^sup>+x. 0 \M1)" - by (intro positive_integral_cong_pos) - (simp add: positive_integral_positive real_of_ereal_pos) - then show "(\\<^sup>+x. ereal (- ?f x) \M1) \ \" by simp - qed } - with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)] - show ?INT - unfolding lebesgue_integral_def[of "M1 \\<^sub>M M2"] lebesgue_integral_def[of M2] - borel[THEN M2.positive_integral_fst_measurable(2), symmetric] - using AE[THEN integral_real] - by simp -qed - -lemma (in pair_sigma_finite) integrable_snd_measurable: - assumes f: "integrable (M1 \\<^sub>M M2) f" - shows "AE y in M2. integrable M1 (\x. f (x, y))" (is "?AE") - and "(\y. (\x. f (x, y) \M1) \M2) = integral\<^sup>L (M1 \\<^sub>M M2) f" (is "?INT") -proof - - interpret Q: pair_sigma_finite M2 M1 by default - have Q_int: "integrable (M2 \\<^sub>M M1) (\(x, y). f (y, x))" - using f unfolding integrable_product_swap_iff . - show ?INT - using Q.integrable_fst_measurable(2)[OF Q_int] - using integral_product_swap[of f] f by auto - show ?AE - using Q.integrable_fst_measurable(1)[OF Q_int] - by simp -qed - -lemma (in pair_sigma_finite) Fubini_integral: - assumes f: "integrable (M1 \\<^sub>M M2) f" - shows "(\y. (\x. f (x, y) \M1) \M2) = (\x. (\y. f (x, y) \M2) \M1)" - unfolding integrable_snd_measurable[OF assms] - unfolding integrable_fst_measurable[OF assms] .. + unfolding positive_integral_snd[OF assms] M2.positive_integral_fst[OF assms] .. section {* Products on counting spaces, densities and distributions *} @@ -741,7 +630,7 @@ (auto simp add: positive_integral_cmult[symmetric] ac_simps) with A f g show "emeasure ?L A = emeasure ?R A" by (simp add: D2.emeasure_pair_measure emeasure_density positive_integral_density - M2.positive_integral_fst_measurable(2)[symmetric] + M2.positive_integral_fst[symmetric] cong: positive_integral_cong) qed simp diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Bochner_Integration.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Bochner_Integration.thy Mon May 19 12:04:45 2014 +0200 @@ -0,0 +1,2615 @@ +(* Title: HOL/Probability/Bochner_Integration.thy + Author: Johannes Hölzl, TU München +*) + +header {* Bochner Integration for Vector-Valued Functions *} + +theory Bochner_Integration + imports Finite_Product_Measure +begin + +text {* + +In the following development of the Bochner integral we use second countable topologies instead +of separable spaces. A second countable topology is also separable. + +*} + +lemma borel_measurable_implies_sequence_metric: + fixes f :: "'a \ 'b :: {metric_space, second_countable_topology}" + assumes [measurable]: "f \ borel_measurable M" + shows "\F. (\i. simple_function M (F i)) \ (\x\space M. (\i. F i x) ----> f x) \ + (\i. \x\space M. dist (F i x) z \ 2 * dist (f x) z)" +proof - + obtain D :: "'b set" where "countable D" and D: "\X. open X \ X \ {} \ \d\D. d \ X" + by (erule countable_dense_setE) + + def e \ "from_nat_into D" + { fix n x + obtain d where "d \ D" and d: "d \ ball x (1 / Suc n)" + using D[of "ball x (1 / Suc n)"] by auto + from `d \ D` D[of UNIV] `countable D` obtain i where "d = e i" + unfolding e_def by (auto dest: from_nat_into_surj) + with d have "\i. dist x (e i) < 1 / Suc n" + by auto } + note e = this + + def A \ "\m n. {x\space M. dist (f x) (e n) < 1 / (Suc m) \ 1 / (Suc m) \ dist (f x) z}" + def B \ "\m. disjointed (A m)" + + def m \ "\N x. Max {m::nat. m \ N \ x \ (\n\N. B m n)}" + def F \ "\N::nat. \x. if (\m\N. x \ (\n\N. B m n)) \ (\n\N. x \ B (m N x) n) + then e (LEAST n. x \ B (m N x) n) else z" + + have B_imp_A[intro, simp]: "\x m n. x \ B m n \ x \ A m n" + using disjointed_subset[of "A m" for m] unfolding B_def by auto + + { fix m + have "\n. A m n \ sets M" + by (auto simp: A_def) + then have "\n. B m n \ sets M" + using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) } + note this[measurable] + + { fix N i x assume "\m\N. x \ (\n\N. B m n)" + then have "m N x \ {m::nat. m \ N \ x \ (\n\N. B m n)}" + unfolding m_def by (intro Max_in) auto + then have "m N x \ N" "\n\N. x \ B (m N x) n" + by auto } + note m = this + + { fix j N i x assume "j \ N" "i \ N" "x \ B j i" + then have "j \ m N x" + unfolding m_def by (intro Max_ge) auto } + note m_upper = this + + show ?thesis + unfolding simple_function_def + proof (safe intro!: exI[of _ F]) + have [measurable]: "\i. F i \ borel_measurable M" + unfolding F_def m_def by measurable + show "\x i. F i -` {x} \ space M \ sets M" + by measurable + + { fix i + { fix n x assume "x \ B (m i x) n" + then have "(LEAST n. x \ B (m i x) n) \ n" + by (intro Least_le) + also assume "n \ i" + finally have "(LEAST n. x \ B (m i x) n) \ i" . } + then have "F i ` space M \ {z} \ e ` {.. i}" + by (auto simp: F_def) + then show "finite (F i ` space M)" + by (rule finite_subset) auto } + + { fix N i n x assume "i \ N" "n \ N" "x \ B i n" + then have 1: "\m\N. x \ (\ n\N. B m n)" by auto + from m[OF this] obtain n where n: "m N x \ N" "n \ N" "x \ B (m N x) n" by auto + moreover + def L \ "LEAST n. x \ B (m N x) n" + have "dist (f x) (e L) < 1 / Suc (m N x)" + proof - + have "x \ B (m N x) L" + using n(3) unfolding L_def by (rule LeastI) + then have "x \ A (m N x) L" + by auto + then show ?thesis + unfolding A_def by simp + qed + ultimately have "dist (f x) (F N x) < 1 / Suc (m N x)" + by (auto simp add: F_def L_def) } + note * = this + + fix x assume "x \ space M" + show "(\i. F i x) ----> f x" + proof cases + assume "f x = z" + then have "\i n. x \ A i n" + unfolding A_def by auto + then have "\i. F i x = z" + by (auto simp: F_def) + then show ?thesis + using `f x = z` by auto + next + assume "f x \ z" + + show ?thesis + proof (rule tendstoI) + fix e :: real assume "0 < e" + with `f x \ z` obtain n where "1 / Suc n < e" "1 / Suc n < dist (f x) z" + by (metis dist_nz order_less_trans neq_iff nat_approx_posE) + with `x\space M` `f x \ z` have "x \ (\i. B n i)" + unfolding A_def B_def UN_disjointed_eq using e by auto + then obtain i where i: "x \ B n i" by auto + + show "eventually (\i. dist (F i x) (f x) < e) sequentially" + using eventually_ge_at_top[of "max n i"] + proof eventually_elim + fix j assume j: "max n i \ j" + with i have "dist (f x) (F j x) < 1 / Suc (m j x)" + by (intro *[OF _ _ i]) auto + also have "\ \ 1 / Suc n" + using j m_upper[OF _ _ i] + by (auto simp: field_simps) + also note `1 / Suc n < e` + finally show "dist (F j x) (f x) < e" + by (simp add: less_imp_le dist_commute) + qed + qed + qed + fix i + { fix n m assume "x \ A n m" + then have "dist (e m) (f x) + dist (f x) z \ 2 * dist (f x) z" + unfolding A_def by (auto simp: dist_commute) + also have "dist (e m) z \ dist (e m) (f x) + dist (f x) z" + by (rule dist_triangle) + finally (xtrans) have "dist (e m) z \ 2 * dist (f x) z" . } + then show "dist (F i x) z \ 2 * dist (f x) z" + unfolding F_def + apply auto + apply (rule LeastI2) + apply auto + done + qed +qed + +lemma real_indicator: "real (indicator A x :: ereal) = indicator A x" + unfolding indicator_def by auto + +lemma split_indicator_asm: + "P (indicator S x) \ \ ((x \ S \ \ P 1) \ (x \ S \ \ P 0))" + unfolding indicator_def by auto + +lemma + fixes f :: "'a \ 'b::semiring_1" assumes "finite A" + shows setsum_mult_indicator[simp]: "(\x \ A. f x * indicator (B x) (g x)) = (\x\{x\A. g x \ B x}. f x)" + and setsum_indicator_mult[simp]: "(\x \ A. indicator (B x) (g x) * f x) = (\x\{x\A. g x \ B x}. f x)" + unfolding indicator_def + using assms by (auto intro!: setsum_mono_zero_cong_right split: split_if_asm) + +lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]: + fixes P :: "('a \ real) \ bool" + assumes u: "u \ borel_measurable M" "\x. 0 \ u x" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" + assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" + assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ (\x. x \ space M \ (\i. U i x) ----> u x) \ P u" + shows "P u" +proof - + have "(\x. ereal (u x)) \ borel_measurable M" using u by auto + from borel_measurable_implies_simple_function_sequence'[OF this] + obtain U where U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and + sup: "\x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\i x. 0 \ U i x" + by blast + + def U' \ "\i x. indicator (space M) x * real (U i x)" + then have U'_sf[measurable]: "\i. simple_function M (U' i)" + using U by (auto intro!: simple_function_compose1[where g=real]) + + show "P u" + proof (rule seq) + fix i show "U' i \ borel_measurable M" "\x. 0 \ U' i x" + using U nn by (auto + intro: borel_measurable_simple_function + intro!: borel_measurable_real_of_ereal real_of_ereal_pos borel_measurable_times + simp: U'_def zero_le_mult_iff) + show "incseq U'" + using U(2,3) nn + by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def + intro!: real_of_ereal_positive_mono) + next + fix x assume x: "x \ space M" + have "(\i. U i x) ----> (SUP i. U i x)" + using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def) + moreover have "(\i. U i x) = (\i. ereal (U' i x))" + using x nn U(3) by (auto simp: fun_eq_iff U'_def ereal_real image_iff eq_commute) + moreover have "(SUP i. U i x) = ereal (u x)" + using sup u(2) by (simp add: max_def) + ultimately show "(\i. U' i x) ----> u x" + by simp + next + fix i + have "U' i ` space M \ real ` (U i ` space M)" "finite (U i ` space M)" + unfolding U'_def using U(1) by (auto dest: simple_functionD) + then have fin: "finite (U' i ` space M)" + by (metis finite_subset finite_imageI) + moreover have "\z. {y. U' i z = y \ y \ U' i ` space M \ z \ space M} = (if z \ space M then {U' i z} else {})" + by auto + ultimately have U': "(\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z) = U' i" + by (simp add: U'_def fun_eq_iff) + have "\x. x \ U' i ` space M \ 0 \ x" + using nn by (auto simp: U'_def real_of_ereal_pos) + with fin have "P (\z. \y\U' i`space M. y * indicator {x\space M. U' i x = y} z)" + proof induct + case empty from set[of "{}"] show ?case + by (simp add: indicator_def[abs_def]) + next + case (insert x F) + then show ?case + by (auto intro!: add mult set setsum_nonneg split: split_indicator split_indicator_asm + simp del: setsum_mult_indicator simp: setsum_nonneg_eq_0_iff ) + qed + with U' show "P (U' i)" by simp + qed +qed + +lemma scaleR_cong_right: + fixes x :: "'a :: real_vector" + shows "(x \ 0 \ r = p) \ r *\<^sub>R x = p *\<^sub>R x" + by (cases "x = 0") auto + +inductive simple_bochner_integrable :: "'a measure \ ('a \ 'b::real_vector) \ bool" for M f where + "simple_function M f \ emeasure M {y\space M. f y \ 0} \ \ \ + simple_bochner_integrable M f" + +lemma simple_bochner_integrable_compose2: + assumes p_0: "p 0 0 = 0" + shows "simple_bochner_integrable M f \ simple_bochner_integrable M g \ + simple_bochner_integrable M (\x. p (f x) (g x))" +proof (safe intro!: simple_bochner_integrable.intros elim!: simple_bochner_integrable.cases del: notI) + assume sf: "simple_function M f" "simple_function M g" + then show "simple_function M (\x. p (f x) (g x))" + by (rule simple_function_compose2) + + from sf have [measurable]: + "f \ measurable M (count_space UNIV)" + "g \ measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function) + + assume fin: "emeasure M {y \ space M. f y \ 0} \ \" "emeasure M {y \ space M. g y \ 0} \ \" + + have "emeasure M {x\space M. p (f x) (g x) \ 0} \ + emeasure M ({x\space M. f x \ 0} \ {x\space M. g x \ 0})" + by (intro emeasure_mono) (auto simp: p_0) + also have "\ \ emeasure M {x\space M. f x \ 0} + emeasure M {x\space M. g x \ 0}" + by (intro emeasure_subadditive) auto + finally show "emeasure M {y \ space M. p (f y) (g y) \ 0} \ \" + using fin by auto +qed + +lemma simple_function_finite_support: + assumes f: "simple_function M f" and fin: "(\\<^sup>+x. f x \M) < \" and nn: "\x. 0 \ f x" + shows "emeasure M {x\space M. f x \ 0} \ \" +proof cases + from f have meas[measurable]: "f \ borel_measurable M" + by (rule borel_measurable_simple_function) + + assume non_empty: "\x\space M. f x \ 0" + + def m \ "Min (f`space M - {0})" + have "m \ f`space M - {0}" + unfolding m_def using f non_empty by (intro Min_in) (auto simp: simple_function_def) + then have m: "0 < m" + using nn by (auto simp: less_le) + + from m have "m * emeasure M {x\space M. 0 \ f x} = + (\\<^sup>+x. m * indicator {x\space M. 0 \ f x} x \M)" + using f by (intro positive_integral_cmult_indicator[symmetric]) auto + also have "\ \ (\\<^sup>+x. f x \M)" + using AE_space + proof (intro positive_integral_mono_AE, eventually_elim) + fix x assume "x \ space M" + with nn show "m * indicator {x \ space M. 0 \ f x} x \ f x" + using f by (auto split: split_indicator simp: simple_function_def m_def) + qed + also note `\ < \` + finally show ?thesis + using m by auto +next + assume "\ (\x\space M. f x \ 0)" + with nn have *: "{x\space M. f x \ 0} = {}" + by auto + show ?thesis unfolding * by simp +qed + +lemma simple_bochner_integrableI_bounded: + assumes f: "simple_function M f" and fin: "(\\<^sup>+x. norm (f x) \M) < \" + shows "simple_bochner_integrable M f" +proof + have "emeasure M {y \ space M. ereal (norm (f y)) \ 0} \ \" + proof (rule simple_function_finite_support) + show "simple_function M (\x. ereal (norm (f x)))" + using f by (rule simple_function_compose1) + show "(\\<^sup>+ y. ereal (norm (f y)) \M) < \" by fact + qed simp + then show "emeasure M {y \ space M. f y \ 0} \ \" by simp +qed fact + +definition simple_bochner_integral :: "'a measure \ ('a \ 'b::real_vector) \ 'b" where + "simple_bochner_integral M f = (\y\f`space M. measure M {x\space M. f x = y} *\<^sub>R y)" + +lemma simple_bochner_integral_partition: + assumes f: "simple_bochner_integrable M f" and g: "simple_function M g" + assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" + assumes v: "\x. x \ space M \ f x = v (g x)" + shows "simple_bochner_integral M f = (\y\g ` space M. measure M {x\space M. g x = y} *\<^sub>R v y)" + (is "_ = ?r") +proof - + from f g have [simp]: "finite (f`space M)" "finite (g`space M)" + by (auto simp: simple_function_def elim: simple_bochner_integrable.cases) + + from f have [measurable]: "f \ measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) + + from g have [measurable]: "g \ measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases) + + { fix y assume "y \ space M" + then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" + by (auto cong: sub simp: v[symmetric]) } + note eq = this + + have "simple_bochner_integral M f = + (\y\f`space M. (\z\g`space M. + if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} else 0) *\<^sub>R y)" + unfolding simple_bochner_integral_def + proof (safe intro!: setsum_cong scaleR_cong_right) + fix y assume y: "y \ space M" "f y \ 0" + have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = + {z. \x\space M. f y = f x \ z = g x}" + by auto + have eq:"{x \ space M. f x = f y} = + (\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i})" + by (auto simp: eq_commute cong: sub rev_conj_cong) + have "finite (g`space M)" by simp + then have "finite {z. \x\space M. f y = f x \ z = g x}" + by (rule rev_finite_subset) auto + moreover + { fix x assume "x \ space M" "f x = f y" + then have "x \ space M" "f x \ 0" + using y by auto + then have "emeasure M {y \ space M. g y = g x} \ emeasure M {y \ space M. f y \ 0}" + by (auto intro!: emeasure_mono cong: sub) + then have "emeasure M {xa \ space M. g xa = g x} < \" + using f by (auto simp: simple_bochner_integrable.simps) } + ultimately + show "measure M {x \ space M. f x = f y} = + (\z\g ` space M. if \x\space M. f y = f x \ z = g x then measure M {x \ space M. g x = z} else 0)" + apply (simp add: setsum_cases eq) + apply (subst measure_finite_Union[symmetric]) + apply (auto simp: disjoint_family_on_def) + done + qed + also have "\ = (\y\f`space M. (\z\g`space M. + if \x\space M. y = f x \ z = g x then measure M {x\space M. g x = z} *\<^sub>R y else 0))" + by (auto intro!: setsum_cong simp: scaleR_setsum_left) + also have "\ = ?r" + by (subst setsum_commute) + (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq) + finally show "simple_bochner_integral M f = ?r" . +qed + +lemma simple_bochner_integral_add: + assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" + shows "simple_bochner_integral M (\x. f x + g x) = + simple_bochner_integral M f + simple_bochner_integral M g" +proof - + from f g have "simple_bochner_integral M (\x. f x + g x) = + (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R (fst y + snd y))" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + moreover from f g have "simple_bochner_integral M f = + (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R fst y)" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + moreover from f g have "simple_bochner_integral M g = + (\y\(\x. (f x, g x)) ` space M. measure M {x \ space M. (f x, g x) = y} *\<^sub>R snd y)" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + ultimately show ?thesis + by (simp add: setsum_addf[symmetric] scaleR_add_right) +qed + +lemma (in linear) simple_bochner_integral_linear: + assumes g: "simple_bochner_integrable M g" + shows "simple_bochner_integral M (\x. f (g x)) = f (simple_bochner_integral M g)" +proof - + from g have "simple_bochner_integral M (\x. f (g x)) = + (\y\g ` space M. measure M {x \ space M. g x = y} *\<^sub>R f y)" + by (intro simple_bochner_integral_partition) + (auto simp: simple_bochner_integrable_compose2[where p="\x y. f x"] zero + elim: simple_bochner_integrable.cases) + also have "\ = f (simple_bochner_integral M g)" + by (simp add: simple_bochner_integral_def setsum scaleR) + finally show ?thesis . +qed + +lemma simple_bochner_integral_minus: + assumes f: "simple_bochner_integrable M f" + shows "simple_bochner_integral M (\x. - f x) = - simple_bochner_integral M f" +proof - + interpret linear uminus by unfold_locales auto + from f show ?thesis + by (rule simple_bochner_integral_linear) +qed + +lemma simple_bochner_integral_diff: + assumes f: "simple_bochner_integrable M f" and g: "simple_bochner_integrable M g" + shows "simple_bochner_integral M (\x. f x - g x) = + simple_bochner_integral M f - simple_bochner_integral M g" + unfolding diff_conv_add_uminus using f g + by (subst simple_bochner_integral_add) + (auto simp: simple_bochner_integral_minus simple_bochner_integrable_compose2[where p="\x y. - y"]) + +lemma simple_bochner_integral_norm_bound: + assumes f: "simple_bochner_integrable M f" + shows "norm (simple_bochner_integral M f) \ simple_bochner_integral M (\x. norm (f x))" +proof - + have "norm (simple_bochner_integral M f) \ + (\y\f ` space M. norm (measure M {x \ space M. f x = y} *\<^sub>R y))" + unfolding simple_bochner_integral_def by (rule norm_setsum) + also have "\ = (\y\f ` space M. measure M {x \ space M. f x = y} *\<^sub>R norm y)" + by (simp add: measure_nonneg) + also have "\ = simple_bochner_integral M (\x. norm (f x))" + using f + by (intro simple_bochner_integral_partition[symmetric]) + (auto intro: f simple_bochner_integrable_compose2 elim: simple_bochner_integrable.cases) + finally show ?thesis . +qed + +lemma simple_bochner_integral_eq_positive_integral: + assumes f: "simple_bochner_integrable M f" "\x. 0 \ f x" + shows "simple_bochner_integral M f = (\\<^sup>+x. f x \M)" +proof - + { fix x y z have "(x \ 0 \ y = z) \ ereal x * y = ereal x * z" + by (cases "x = 0") (auto simp: zero_ereal_def[symmetric]) } + note ereal_cong_mult = this + + have [measurable]: "f \ borel_measurable M" + using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + { fix y assume y: "y \ space M" "f y \ 0" + have "ereal (measure M {x \ space M. f x = f y}) = emeasure M {x \ space M. f x = f y}" + proof (rule emeasure_eq_ereal_measure[symmetric]) + have "emeasure M {x \ space M. f x = f y} \ emeasure M {x \ space M. f x \ 0}" + using y by (intro emeasure_mono) auto + with f show "emeasure M {x \ space M. f x = f y} \ \" + by (auto simp: simple_bochner_integrable.simps) + qed + moreover have "{x \ space M. f x = f y} = (\x. ereal (f x)) -` {ereal (f y)} \ space M" + by auto + ultimately have "ereal (measure M {x \ space M. f x = f y}) = + emeasure M ((\x. ereal (f x)) -` {ereal (f y)} \ space M)" by simp } + with f have "simple_bochner_integral M f = (\\<^sup>Sx. f x \M)" + unfolding simple_integral_def + by (subst simple_bochner_integral_partition[OF f(1), where g="\x. ereal (f x)" and v=real]) + (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases + intro!: setsum_cong ereal_cong_mult + simp: setsum_ereal[symmetric] times_ereal.simps(1)[symmetric] mult_ac + simp del: setsum_ereal times_ereal.simps(1)) + also have "\ = (\\<^sup>+x. f x \M)" + using f + by (intro positive_integral_eq_simple_integral[symmetric]) + (auto simp: simple_function_compose1 simple_bochner_integrable.simps) + finally show ?thesis . +qed + +lemma simple_bochner_integral_bounded: + fixes f :: "'a \ 'b::{real_normed_vector, second_countable_topology}" + assumes f[measurable]: "f \ borel_measurable M" + assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t" + shows "ereal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \ + (\\<^sup>+ x. norm (f x - s x) \M) + (\\<^sup>+ x. norm (f x - t x) \M)" + (is "ereal (norm (?s - ?t)) \ ?S + ?T") +proof - + have [measurable]: "s \ borel_measurable M" "t \ borel_measurable M" + using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + have "ereal (norm (?s - ?t)) = norm (simple_bochner_integral M (\x. s x - t x))" + using s t by (subst simple_bochner_integral_diff) auto + also have "\ \ simple_bochner_integral M (\x. norm (s x - t x))" + using simple_bochner_integrable_compose2[of "op -" M "s" "t"] s t + by (auto intro!: simple_bochner_integral_norm_bound) + also have "\ = (\\<^sup>+x. norm (s x - t x) \M)" + using simple_bochner_integrable_compose2[of "\x y. norm (x - y)" M "s" "t"] s t + by (auto intro!: simple_bochner_integral_eq_positive_integral) + also have "\ \ (\\<^sup>+x. ereal (norm (f x - s x)) + ereal (norm (f x - t x)) \M)" + by (auto intro!: positive_integral_mono) + (metis (erased, hide_lams) add_diff_cancel_left add_diff_eq diff_add_eq order_trans + norm_minus_commute norm_triangle_ineq4 order_refl) + also have "\ = ?S + ?T" + by (rule positive_integral_add) auto + finally show ?thesis . +qed + +inductive has_bochner_integral :: "'a measure \ ('a \ 'b) \ 'b::{real_normed_vector, second_countable_topology} \ bool" + for M f x where + "f \ borel_measurable M \ + (\i. simple_bochner_integrable M (s i)) \ + (\i. \\<^sup>+x. norm (f x - s i x) \M) ----> 0 \ + (\i. simple_bochner_integral M (s i)) ----> x \ + has_bochner_integral M f x" + +lemma has_bochner_integral_cong: + assumes "M = N" "\x. x \ space N \ f x = g x" "x = y" + shows "has_bochner_integral M f x \ has_bochner_integral N g y" + unfolding has_bochner_integral.simps assms(1,3) + using assms(2) by (simp cong: measurable_cong_strong positive_integral_cong_strong) + +lemma has_bochner_integral_cong_AE: + "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ + has_bochner_integral M f x \ has_bochner_integral M g x" + unfolding has_bochner_integral.simps + by (intro arg_cong[where f=Ex] ext conj_cong rev_conj_cong refl arg_cong[where f="\x. x ----> 0"] + positive_integral_cong_AE) + auto + +lemma borel_measurable_has_bochner_integral[measurable_dest]: + "has_bochner_integral M f x \ f \ borel_measurable M" + by (auto elim: has_bochner_integral.cases) + +lemma has_bochner_integral_simple_bochner_integrable: + "simple_bochner_integrable M f \ has_bochner_integral M f (simple_bochner_integral M f)" + by (rule has_bochner_integral.intros[where s="\_. f"]) + (auto intro: borel_measurable_simple_function + elim: simple_bochner_integrable.cases + simp: zero_ereal_def[symmetric]) + +lemma has_bochner_integral_real_indicator: + assumes [measurable]: "A \ sets M" and A: "emeasure M A < \" + shows "has_bochner_integral M (indicator A) (measure M A)" +proof - + have sbi: "simple_bochner_integrable M (indicator A::'a \ real)" + proof + have "{y \ space M. (indicator A y::real) \ 0} = A" + using sets.sets_into_space[OF `A\sets M`] by (auto split: split_indicator) + then show "emeasure M {y \ space M. (indicator A y::real) \ 0} \ \" + using A by auto + qed (rule simple_function_indicator assms)+ + moreover have "simple_bochner_integral M (indicator A) = measure M A" + using simple_bochner_integral_eq_positive_integral[OF sbi] A + by (simp add: ereal_indicator emeasure_eq_ereal_measure) + ultimately show ?thesis + by (metis has_bochner_integral_simple_bochner_integrable) +qed + +lemma has_bochner_integral_add: + "has_bochner_integral M f x \ has_bochner_integral M g y \ + has_bochner_integral M (\x. f x + g x) (x + y)" +proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) + fix sf sg + assume f_sf: "(\i. \\<^sup>+ x. norm (f x - sf i x) \M) ----> 0" + assume g_sg: "(\i. \\<^sup>+ x. norm (g x - sg i x) \M) ----> 0" + + assume sf: "\i. simple_bochner_integrable M (sf i)" + and sg: "\i. simple_bochner_integrable M (sg i)" + then have [measurable]: "\i. sf i \ borel_measurable M" "\i. sg i \ borel_measurable M" + by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + assume [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" + + show "\i. simple_bochner_integrable M (\x. sf i x + sg i x)" + using sf sg by (simp add: simple_bochner_integrable_compose2) + + show "(\i. \\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \M) ----> 0" + (is "?f ----> 0") + proof (rule tendsto_sandwich) + show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) ----> 0" + by (auto simp: positive_integral_positive) + show "eventually (\i. ?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) \M) + \\<^sup>+ x. (norm (g x - sg i x)) \M) sequentially" + (is "eventually (\i. ?f i \ ?g i) sequentially") + proof (intro always_eventually allI) + fix i have "?f i \ (\\<^sup>+ x. (norm (f x - sf i x)) + ereal (norm (g x - sg i x)) \M)" + by (auto intro!: positive_integral_mono norm_diff_triangle_ineq) + also have "\ = ?g i" + by (intro positive_integral_add) auto + finally show "?f i \ ?g i" . + qed + show "?g ----> 0" + using tendsto_add_ereal[OF _ _ f_sf g_sg] by simp + qed +qed (auto simp: simple_bochner_integral_add tendsto_add) + +lemma has_bochner_integral_bounded_linear: + assumes "bounded_linear T" + shows "has_bochner_integral M f x \ has_bochner_integral M (\x. T (f x)) (T x)" +proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) + interpret T: bounded_linear T by fact + have [measurable]: "T \ borel_measurable borel" + by (intro borel_measurable_continuous_on1 T.continuous_on continuous_on_id) + assume [measurable]: "f \ borel_measurable M" + then show "(\x. T (f x)) \ borel_measurable M" + by auto + + fix s assume f_s: "(\i. \\<^sup>+ x. norm (f x - s i x) \M) ----> 0" + assume s: "\i. simple_bochner_integrable M (s i)" + then show "\i. simple_bochner_integrable M (\x. T (s i x))" + by (auto intro: simple_bochner_integrable_compose2 T.zero) + + have [measurable]: "\i. s i \ borel_measurable M" + using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + obtain K where K: "K > 0" "\x i. norm (T (f x) - T (s i x)) \ norm (f x - s i x) * K" + using T.pos_bounded by (auto simp: T.diff[symmetric]) + + show "(\i. \\<^sup>+ x. norm (T (f x) - T (s i x)) \M) ----> 0" + (is "?f ----> 0") + proof (rule tendsto_sandwich) + show "eventually (\n. 0 \ ?f n) sequentially" "(\_. 0) ----> 0" + by (auto simp: positive_integral_positive) + + show "eventually (\i. ?f i \ K * (\\<^sup>+ x. norm (f x - s i x) \M)) sequentially" + (is "eventually (\i. ?f i \ ?g i) sequentially") + proof (intro always_eventually allI) + fix i have "?f i \ (\\<^sup>+ x. ereal K * norm (f x - s i x) \M)" + using K by (intro positive_integral_mono) (auto simp: mult_ac) + also have "\ = ?g i" + using K by (intro positive_integral_cmult) auto + finally show "?f i \ ?g i" . + qed + show "?g ----> 0" + using ereal_lim_mult[OF f_s, of "ereal K"] by simp + qed + + assume "(\i. simple_bochner_integral M (s i)) ----> x" + with s show "(\i. simple_bochner_integral M (\x. T (s i x))) ----> T x" + by (auto intro!: T.tendsto simp: T.simple_bochner_integral_linear) +qed + +lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\x. 0) 0" + by (auto intro!: has_bochner_integral.intros[where s="\_ _. 0"] + simp: zero_ereal_def[symmetric] simple_bochner_integrable.simps + simple_bochner_integral_def image_constant_conv) + +lemma has_bochner_integral_scaleR_left[intro]: + "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x *\<^sub>R c) (x *\<^sub>R c)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_left]) + +lemma has_bochner_integral_scaleR_right[intro]: + "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c *\<^sub>R f x) (c *\<^sub>R x)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_scaleR_right]) + +lemma has_bochner_integral_mult_left[intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x * c) (x * c)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_left]) + +lemma has_bochner_integral_mult_right[intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c * f x) (c * x)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_mult_right]) + +lemmas has_bochner_integral_divide = + has_bochner_integral_bounded_linear[OF bounded_linear_divide] + +lemma has_bochner_integral_divide_zero[intro]: + fixes c :: "_::{real_normed_field, field_inverse_zero, second_countable_topology}" + shows "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x / c) (x / c)" + using has_bochner_integral_divide by (cases "c = 0") auto + +lemma has_bochner_integral_inner_left[intro]: + "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. f x \ c) (x \ c)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_left]) + +lemma has_bochner_integral_inner_right[intro]: + "(c \ 0 \ has_bochner_integral M f x) \ has_bochner_integral M (\x. c \ f x) (c \ x)" + by (cases "c = 0") (auto simp add: has_bochner_integral_bounded_linear[OF bounded_linear_inner_right]) + +lemmas has_bochner_integral_minus = + has_bochner_integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] +lemmas has_bochner_integral_Re = + has_bochner_integral_bounded_linear[OF bounded_linear_Re] +lemmas has_bochner_integral_Im = + has_bochner_integral_bounded_linear[OF bounded_linear_Im] +lemmas has_bochner_integral_cnj = + has_bochner_integral_bounded_linear[OF bounded_linear_cnj] +lemmas has_bochner_integral_of_real = + has_bochner_integral_bounded_linear[OF bounded_linear_of_real] +lemmas has_bochner_integral_fst = + has_bochner_integral_bounded_linear[OF bounded_linear_fst] +lemmas has_bochner_integral_snd = + has_bochner_integral_bounded_linear[OF bounded_linear_snd] + +lemma has_bochner_integral_indicator: + "A \ sets M \ emeasure M A < \ \ + has_bochner_integral M (\x. indicator A x *\<^sub>R c) (measure M A *\<^sub>R c)" + by (intro has_bochner_integral_scaleR_left has_bochner_integral_real_indicator) + +lemma has_bochner_integral_diff: + "has_bochner_integral M f x \ has_bochner_integral M g y \ + has_bochner_integral M (\x. f x - g x) (x - y)" + unfolding diff_conv_add_uminus + by (intro has_bochner_integral_add has_bochner_integral_minus) + +lemma has_bochner_integral_setsum: + "(\i. i \ I \ has_bochner_integral M (f i) (x i)) \ + has_bochner_integral M (\x. \i\I. f i x) (\i\I. x i)" + by (induct I rule: infinite_finite_induct) + (auto intro: has_bochner_integral_zero has_bochner_integral_add) + +lemma has_bochner_integral_implies_finite_norm: + "has_bochner_integral M f x \ (\\<^sup>+x. norm (f x) \M) < \" +proof (elim has_bochner_integral.cases) + fix s v + assume [measurable]: "f \ borel_measurable M" and s: "\i. simple_bochner_integrable M (s i)" and + lim_0: "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" + from order_tendstoD[OF lim_0, of "\"] + obtain i where f_s_fin: "(\\<^sup>+ x. ereal (norm (f x - s i x)) \M) < \" + by (metis (mono_tags, lifting) eventually_False_sequentially eventually_elim1 + less_ereal.simps(4) zero_ereal_def) + + have [measurable]: "\i. s i \ borel_measurable M" + using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + def m \ "if space M = {} then 0 else Max ((\x. norm (s i x))`space M)" + have "finite (s i ` space M)" + using s by (auto simp: simple_function_def simple_bochner_integrable.simps) + then have "finite (norm ` s i ` space M)" + by (rule finite_imageI) + then have "\x. x \ space M \ norm (s i x) \ m" "0 \ m" + by (auto simp: m_def image_comp comp_def Max_ge_iff) + then have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. ereal m * indicator {x\space M. s i x \ 0} x \M)" + by (auto split: split_indicator intro!: Max_ge positive_integral_mono simp:) + also have "\ < \" + using s by (subst positive_integral_cmult_indicator) (auto simp: `0 \ m` simple_bochner_integrable.simps) + finally have s_fin: "(\\<^sup>+x. norm (s i x) \M) < \" . + + have "(\\<^sup>+ x. norm (f x) \M) \ (\\<^sup>+ x. ereal (norm (f x - s i x)) + ereal (norm (s i x)) \M)" + by (auto intro!: positive_integral_mono) (metis add_commute norm_triangle_sub) + also have "\ = (\\<^sup>+x. norm (f x - s i x) \M) + (\\<^sup>+x. norm (s i x) \M)" + by (rule positive_integral_add) auto + also have "\ < \" + using s_fin f_s_fin by auto + finally show "(\\<^sup>+ x. ereal (norm (f x)) \M) < \" . +qed + +lemma has_bochner_integral_norm_bound: + assumes i: "has_bochner_integral M f x" + shows "norm x \ (\\<^sup>+x. norm (f x) \M)" +using assms proof + fix s assume + x: "(\i. simple_bochner_integral M (s i)) ----> x" (is "?s ----> x") and + s[simp]: "\i. simple_bochner_integrable M (s i)" and + lim: "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" and + f[measurable]: "f \ borel_measurable M" + + have [measurable]: "\i. s i \ borel_measurable M" + using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function) + + show "norm x \ (\\<^sup>+x. norm (f x) \M)" + proof (rule LIMSEQ_le) + show "(\i. ereal (norm (?s i))) ----> norm x" + using x by (intro tendsto_intros lim_ereal[THEN iffD2]) + show "\N. \n\N. norm (?s n) \ (\\<^sup>+x. norm (f x - s n x) \M) + (\\<^sup>+x. norm (f x) \M)" + (is "\N. \n\N. _ \ ?t n") + proof (intro exI allI impI) + fix n + have "ereal (norm (?s n)) \ simple_bochner_integral M (\x. norm (s n x))" + by (auto intro!: simple_bochner_integral_norm_bound) + also have "\ = (\\<^sup>+x. norm (s n x) \M)" + by (intro simple_bochner_integral_eq_positive_integral) + (auto intro: s simple_bochner_integrable_compose2) + also have "\ \ (\\<^sup>+x. ereal (norm (f x - s n x)) + norm (f x) \M)" + by (auto intro!: positive_integral_mono) + (metis add_commute norm_minus_commute norm_triangle_sub) + also have "\ = ?t n" + by (rule positive_integral_add) auto + finally show "norm (?s n) \ ?t n" . + qed + have "?t ----> 0 + (\\<^sup>+ x. ereal (norm (f x)) \M)" + using has_bochner_integral_implies_finite_norm[OF i] + by (intro tendsto_add_ereal tendsto_const lim) auto + then show "?t ----> \\<^sup>+ x. ereal (norm (f x)) \M" + by simp + qed +qed + +lemma has_bochner_integral_eq: + "has_bochner_integral M f x \ has_bochner_integral M f y \ x = y" +proof (elim has_bochner_integral.cases) + assume f[measurable]: "f \ borel_measurable M" + + fix s t + assume "(\i. \\<^sup>+ x. norm (f x - s i x) \M) ----> 0" (is "?S ----> 0") + assume "(\i. \\<^sup>+ x. norm (f x - t i x) \M) ----> 0" (is "?T ----> 0") + assume s: "\i. simple_bochner_integrable M (s i)" + assume t: "\i. simple_bochner_integrable M (t i)" + + have [measurable]: "\i. s i \ borel_measurable M" "\i. t i \ borel_measurable M" + using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases) + + let ?s = "\i. simple_bochner_integral M (s i)" + let ?t = "\i. simple_bochner_integral M (t i)" + assume "?s ----> x" "?t ----> y" + then have "(\i. norm (?s i - ?t i)) ----> norm (x - y)" + by (intro tendsto_intros) + moreover + have "(\i. ereal (norm (?s i - ?t i))) ----> ereal 0" + proof (rule tendsto_sandwich) + show "eventually (\i. 0 \ ereal (norm (?s i - ?t i))) sequentially" "(\_. 0) ----> ereal 0" + by (auto simp: positive_integral_positive zero_ereal_def[symmetric]) + + show "eventually (\i. norm (?s i - ?t i) \ ?S i + ?T i) sequentially" + by (intro always_eventually allI simple_bochner_integral_bounded s t f) + show "(\i. ?S i + ?T i) ----> ereal 0" + using tendsto_add_ereal[OF _ _ `?S ----> 0` `?T ----> 0`] + by (simp add: zero_ereal_def[symmetric]) + qed + then have "(\i. norm (?s i - ?t i)) ----> 0" + by simp + ultimately have "norm (x - y) = 0" + by (rule LIMSEQ_unique) + then show "x = y" by simp +qed + +lemma has_bochner_integralI_AE: + assumes f: "has_bochner_integral M f x" + and g: "g \ borel_measurable M" + and ae: "AE x in M. f x = g x" + shows "has_bochner_integral M g x" + using f +proof (safe intro!: has_bochner_integral.intros elim!: has_bochner_integral.cases) + fix s assume "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" + also have "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) = (\i. \\<^sup>+ x. ereal (norm (g x - s i x)) \M)" + using ae + by (intro ext positive_integral_cong_AE, eventually_elim) simp + finally show "(\i. \\<^sup>+ x. ereal (norm (g x - s i x)) \M) ----> 0" . +qed (auto intro: g) + +lemma has_bochner_integral_eq_AE: + assumes f: "has_bochner_integral M f x" + and g: "has_bochner_integral M g y" + and ae: "AE x in M. f x = g x" + shows "x = y" +proof - + from assms have "has_bochner_integral M g x" + by (auto intro: has_bochner_integralI_AE) + from this g show "x = y" + by (rule has_bochner_integral_eq) +qed + +inductive integrable for M f where + "has_bochner_integral M f x \ integrable M f" + +definition lebesgue_integral ("integral\<^sup>L") where + "integral\<^sup>L M f = (THE x. has_bochner_integral M f x)" + +syntax + "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\ _. _ \_" [60,61] 110) + +translations + "\ x. f \M" == "CONST lebesgue_integral M (\x. f)" + +lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \ integral\<^sup>L M f = x" + by (metis the_equality has_bochner_integral_eq lebesgue_integral_def) + +lemma has_bochner_integral_integrable: + "integrable M f \ has_bochner_integral M f (integral\<^sup>L M f)" + by (auto simp: has_bochner_integral_integral_eq integrable.simps) + +lemma has_bochner_integral_iff: + "has_bochner_integral M f x \ integrable M f \ integral\<^sup>L M f = x" + by (metis has_bochner_integral_integrable has_bochner_integral_integral_eq integrable.intros) + +lemma simple_bochner_integrable_eq_integral: + "simple_bochner_integrable M f \ simple_bochner_integral M f = integral\<^sup>L M f" + using has_bochner_integral_simple_bochner_integrable[of M f] + by (simp add: has_bochner_integral_integral_eq) + +lemma not_integrable_integral_eq: "\ integrable M f \ integral\<^sup>L M f = (THE x. False)" + unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The]) + +lemma integral_eq_cases: + "integrable M f \ integrable N g \ + (integrable M f \ integrable N g \ integral\<^sup>L M f = integral\<^sup>L N g) \ + integral\<^sup>L M f = integral\<^sup>L N g" + by (metis not_integrable_integral_eq) + +lemma borel_measurable_integrable[measurable_dest]: "integrable M f \ f \ borel_measurable M" + by (auto elim: integrable.cases has_bochner_integral.cases) + +lemma integrable_cong: + "M = N \ (\x. x \ space N \ f x = g x) \ integrable M f \ integrable N g" + using assms by (simp cong: has_bochner_integral_cong add: integrable.simps) + +lemma integrable_cong_AE: + "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ + integrable M f \ integrable M g" + unfolding integrable.simps + by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext) + +lemma integral_cong: + "M = N \ (\x. x \ space N \ f x = g x) \ integral\<^sup>L M f = integral\<^sup>L N g" + using assms by (simp cong: has_bochner_integral_cong add: lebesgue_integral_def) + +lemma integral_cong_AE: + "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. f x = g x \ + integral\<^sup>L M f = integral\<^sup>L M g" + unfolding lebesgue_integral_def + by (intro has_bochner_integral_cong_AE arg_cong[where f=The] ext) + +lemma integrable_add[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x + g x)" + by (auto simp: integrable.simps intro: has_bochner_integral_add) + +lemma integrable_zero[simp, intro]: "integrable M (\x. 0)" + by (metis has_bochner_integral_zero integrable.simps) + +lemma integrable_setsum[simp, intro]: "(\i. i \ I \ integrable M (f i)) \ integrable M (\x. \i\I. f i x)" + by (metis has_bochner_integral_setsum integrable.simps) + +lemma integrable_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ + integrable M (\x. indicator A x *\<^sub>R c)" + by (metis has_bochner_integral_indicator integrable.simps) + +lemma integrable_real_indicator[simp, intro]: "A \ sets M \ emeasure M A < \ \ + integrable M (indicator A :: 'a \ real)" + by (metis has_bochner_integral_real_indicator integrable.simps) + +lemma integrable_diff[simp, intro]: "integrable M f \ integrable M g \ integrable M (\x. f x - g x)" + by (auto simp: integrable.simps intro: has_bochner_integral_diff) + +lemma integrable_bounded_linear: "bounded_linear T \ integrable M f \ integrable M (\x. T (f x))" + by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear) + +lemma integrable_scaleR_left[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. f x *\<^sub>R c)" + unfolding integrable.simps by fastforce + +lemma integrable_scaleR_right[simp, intro]: "(c \ 0 \ integrable M f) \ integrable M (\x. c *\<^sub>R f x)" + unfolding integrable.simps by fastforce + +lemma integrable_mult_left[simp, intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x * c)" + unfolding integrable.simps by fastforce + +lemma integrable_mult_right[simp, intro]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \ 0 \ integrable M f) \ integrable M (\x. c * f x)" + unfolding integrable.simps by fastforce + +lemma integrable_divide_zero[simp, intro]: + fixes c :: "_::{real_normed_field, field_inverse_zero, second_countable_topology}" + shows "(c \ 0 \ integrable M f) \ integrable M (\x. f x / c)" + unfolding integrable.simps by fastforce + +lemma integrable_inner_left[simp, intro]: + "(c \ 0 \ integrable M f) \ integrable M (\x. f x \ c)" + unfolding integrable.simps by fastforce + +lemma integrable_inner_right[simp, intro]: + "(c \ 0 \ integrable M f) \ integrable M (\x. c \ f x)" + unfolding integrable.simps by fastforce + +lemmas integrable_minus[simp, intro] = + integrable_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] +lemmas integrable_divide[simp, intro] = + integrable_bounded_linear[OF bounded_linear_divide] +lemmas integrable_Re[simp, intro] = + integrable_bounded_linear[OF bounded_linear_Re] +lemmas integrable_Im[simp, intro] = + integrable_bounded_linear[OF bounded_linear_Im] +lemmas integrable_cnj[simp, intro] = + integrable_bounded_linear[OF bounded_linear_cnj] +lemmas integrable_of_real[simp, intro] = + integrable_bounded_linear[OF bounded_linear_of_real] +lemmas integrable_fst[simp, intro] = + integrable_bounded_linear[OF bounded_linear_fst] +lemmas integrable_snd[simp, intro] = + integrable_bounded_linear[OF bounded_linear_snd] + +lemma integral_zero[simp]: "integral\<^sup>L M (\x. 0) = 0" + by (intro has_bochner_integral_integral_eq has_bochner_integral_zero) + +lemma integral_add[simp]: "integrable M f \ integrable M g \ + integral\<^sup>L M (\x. f x + g x) = integral\<^sup>L M f + integral\<^sup>L M g" + by (intro has_bochner_integral_integral_eq has_bochner_integral_add has_bochner_integral_integrable) + +lemma integral_diff[simp]: "integrable M f \ integrable M g \ + integral\<^sup>L M (\x. f x - g x) = integral\<^sup>L M f - integral\<^sup>L M g" + by (intro has_bochner_integral_integral_eq has_bochner_integral_diff has_bochner_integral_integrable) + +lemma integral_setsum[simp]: "(\i. i \ I \ integrable M (f i)) \ + integral\<^sup>L M (\x. \i\I. f i x) = (\i\I. integral\<^sup>L M (f i))" + by (intro has_bochner_integral_integral_eq has_bochner_integral_setsum has_bochner_integral_integrable) + +lemma integral_bounded_linear: "bounded_linear T \ integrable M f \ + integral\<^sup>L M (\x. T (f x)) = T (integral\<^sup>L M f)" + by (metis has_bochner_integral_bounded_linear has_bochner_integral_integrable has_bochner_integral_integral_eq) + +lemma integral_indicator[simp]: "A \ sets M \ emeasure M A < \ \ + integral\<^sup>L M (\x. indicator A x *\<^sub>R c) = measure M A *\<^sub>R c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_indicator has_bochner_integral_integrable) + +lemma integral_real_indicator[simp]: "A \ sets M \ emeasure M A < \ \ + integral\<^sup>L M (indicator A :: 'a \ real) = measure M A" + by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator has_bochner_integral_integrable) + +lemma integral_scaleR_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x *\<^sub>R c \M) = integral\<^sup>L M f *\<^sub>R c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_left) + +lemma integral_scaleR_right[simp]: "(c \ 0 \ integrable M f) \ (\ x. c *\<^sub>R f x \M) = c *\<^sub>R integral\<^sup>L M f" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_scaleR_right) + +lemma integral_mult_left[simp]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \ 0 \ integrable M f) \ (\ x. f x * c \M) = integral\<^sup>L M f * c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_left) + +lemma integral_mult_right[simp]: + fixes c :: "_::{real_normed_algebra,second_countable_topology}" + shows "(c \ 0 \ integrable M f) \ (\ x. c * f x \M) = c * integral\<^sup>L M f" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_mult_right) + +lemma integral_inner_left[simp]: "(c \ 0 \ integrable M f) \ (\ x. f x \ c \M) = integral\<^sup>L M f \ c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left) + +lemma integral_inner_right[simp]: "(c \ 0 \ integrable M f) \ (\ x. c \ f x \M) = c \ integral\<^sup>L M f" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_right) + +lemma integral_divide_zero[simp]: + fixes c :: "_::{real_normed_field, field_inverse_zero, second_countable_topology}" + shows "(c \ 0 \ integrable M f) \ integral\<^sup>L M (\x. f x / c) = integral\<^sup>L M f / c" + by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_divide_zero) + +lemmas integral_minus[simp] = + integral_bounded_linear[OF bounded_linear_minus[OF bounded_linear_ident]] +lemmas integral_divide[simp] = + integral_bounded_linear[OF bounded_linear_divide] +lemmas integral_Re[simp] = + integral_bounded_linear[OF bounded_linear_Re] +lemmas integral_Im[simp] = + integral_bounded_linear[OF bounded_linear_Im] +lemmas integral_cnj[simp] = + integral_bounded_linear[OF bounded_linear_cnj] +lemmas integral_of_real[simp] = + integral_bounded_linear[OF bounded_linear_of_real] +lemmas integral_fst[simp] = + integral_bounded_linear[OF bounded_linear_fst] +lemmas integral_snd[simp] = + integral_bounded_linear[OF bounded_linear_snd] + +lemma integral_norm_bound_ereal: + "integrable M f \ norm (integral\<^sup>L M f) \ (\\<^sup>+x. norm (f x) \M)" + by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound) + +lemma integrableI_sequence: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes f[measurable]: "f \ borel_measurable M" + assumes s: "\i. simple_bochner_integrable M (s i)" + assumes lim: "(\i. \\<^sup>+x. norm (f x - s i x) \M) ----> 0" (is "?S ----> 0") + shows "integrable M f" +proof - + let ?s = "\n. simple_bochner_integral M (s n)" + + have "\x. ?s ----> x" + unfolding convergent_eq_cauchy + proof (rule metric_CauchyI) + fix e :: real assume "0 < e" + then have "0 < ereal (e / 2)" by auto + from order_tendstoD(2)[OF lim this] + obtain M where M: "\n. M \ n \ ?S n < e / 2" + by (auto simp: eventually_sequentially) + show "\M. \m\M. \n\M. dist (?s m) (?s n) < e" + proof (intro exI allI impI) + fix m n assume m: "M \ m" and n: "M \ n" + have "?S n \ \" + using M[OF n] by auto + have "norm (?s n - ?s m) \ ?S n + ?S m" + by (intro simple_bochner_integral_bounded s f) + also have "\ < ereal (e / 2) + e / 2" + using ereal_add_strict_mono[OF less_imp_le[OF M[OF n]] _ `?S n \ \` M[OF m]] + by (auto simp: positive_integral_positive) + also have "\ = e" by simp + finally show "dist (?s n) (?s m) < e" + by (simp add: dist_norm) + qed + qed + then obtain x where "?s ----> x" .. + show ?thesis + by (rule, rule) fact+ +qed + +lemma positive_integral_dominated_convergence_norm: + fixes u' :: "_ \ _::{real_normed_vector, second_countable_topology}" + assumes [measurable]: + "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" + and bound: "\j. AE x in M. norm (u j x) \ w x" + and w: "(\\<^sup>+x. w x \M) < \" + and u': "AE x in M. (\i. u i x) ----> u' x" + shows "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) ----> 0" +proof - + have "AE x in M. \j. norm (u j x) \ w x" + unfolding AE_all_countable by rule fact + with u' have bnd: "AE x in M. \j. norm (u' x - u j x) \ 2 * w x" + proof (eventually_elim, intro allI) + fix i x assume "(\i. u i x) ----> u' x" "\j. norm (u j x) \ w x" "\j. norm (u j x) \ w x" + then have "norm (u' x) \ w x" "norm (u i x) \ w x" + by (auto intro: LIMSEQ_le_const2 tendsto_norm) + then have "norm (u' x) + norm (u i x) \ 2 * w x" + by simp + also have "norm (u' x - u i x) \ norm (u' x) + norm (u i x)" + by (rule norm_triangle_ineq4) + finally (xtrans) show "norm (u' x - u i x) \ 2 * w x" . + qed + + have "(\i. (\\<^sup>+x. norm (u' x - u i x) \M)) ----> (\\<^sup>+x. 0 \M)" + proof (rule positive_integral_dominated_convergence) + show "(\\<^sup>+x. 2 * w x \M) < \" + by (rule positive_integral_mult_bounded_inf[OF _ w, of 2]) auto + show "AE x in M. (\i. ereal (norm (u' x - u i x))) ----> 0" + using u' + proof eventually_elim + fix x assume "(\i. u i x) ----> u' x" + from tendsto_diff[OF tendsto_const[of "u' x"] this] + show "(\i. ereal (norm (u' x - u i x))) ----> 0" + by (simp add: zero_ereal_def tendsto_norm_zero_iff) + qed + qed (insert bnd, auto) + then show ?thesis by simp +qed + +lemma integrableI_bounded: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes f[measurable]: "f \ borel_measurable M" and fin: "(\\<^sup>+x. norm (f x) \M) < \" + shows "integrable M f" +proof - + from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where + s: "\i. simple_function M (s i)" and + pointwise: "\x. x \ space M \ (\i. s i x) ----> f x" and + bound: "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" + by (simp add: norm_conv_dist) metis + + show ?thesis + proof (rule integrableI_sequence) + { fix i + have "(\\<^sup>+x. norm (s i x) \M) \ (\\<^sup>+x. 2 * ereal (norm (f x)) \M)" + by (intro positive_integral_mono) (simp add: bound) + also have "\ = 2 * (\\<^sup>+x. ereal (norm (f x)) \M)" + by (rule positive_integral_cmult) auto + finally have "(\\<^sup>+x. norm (s i x) \M) < \" + using fin by auto } + note fin_s = this + + show "\i. simple_bochner_integrable M (s i)" + by (rule simple_bochner_integrableI_bounded) fact+ + + show "(\i. \\<^sup>+ x. ereal (norm (f x - s i x)) \M) ----> 0" + proof (rule positive_integral_dominated_convergence_norm) + show "\j. AE x in M. norm (s j x) \ 2 * norm (f x)" + using bound by auto + show "\i. s i \ borel_measurable M" "(\x. 2 * norm (f x)) \ borel_measurable M" + using s by (auto intro: borel_measurable_simple_function) + show "(\\<^sup>+ x. ereal (2 * norm (f x)) \M) < \" + using fin unfolding times_ereal.simps(1)[symmetric] by (subst positive_integral_cmult) auto + show "AE x in M. (\i. s i x) ----> f x" + using pointwise by auto + qed fact + qed fact +qed + +lemma integrableI_nonneg: + fixes f :: "'a \ real" + assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "(\\<^sup>+x. f x \M) < \" + shows "integrable M f" +proof - + have "(\\<^sup>+x. norm (f x) \M) = (\\<^sup>+x. f x \M)" + using assms by (intro positive_integral_cong_AE) auto + then show ?thesis + using assms by (intro integrableI_bounded) auto +qed + +lemma integrable_iff_bounded: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + shows "integrable M f \ f \ borel_measurable M \ (\\<^sup>+x. norm (f x) \M) < \" + using integrableI_bounded[of f M] has_bochner_integral_implies_finite_norm[of M f] + unfolding integrable.simps has_bochner_integral.simps[abs_def] by auto + +lemma integrable_bound: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + and g :: "'a \ 'c::{banach, second_countable_topology}" + shows "integrable M f \ g \ borel_measurable M \ (AE x in M. norm (g x) \ norm (f x)) \ + integrable M g" + unfolding integrable_iff_bounded +proof safe + assume "f \ borel_measurable M" "g \ borel_measurable M" + assume "AE x in M. norm (g x) \ norm (f x)" + then have "(\\<^sup>+ x. ereal (norm (g x)) \M) \ (\\<^sup>+ x. ereal (norm (f x)) \M)" + by (intro positive_integral_mono_AE) auto + also assume "(\\<^sup>+ x. ereal (norm (f x)) \M) < \" + finally show "(\\<^sup>+ x. ereal (norm (g x)) \M) < \" . +qed + +lemma integrable_abs[simp, intro]: + fixes f :: "'a \ real" + assumes [measurable]: "integrable M f" shows "integrable M (\x. \f x\)" + using assms by (rule integrable_bound) auto + +lemma integrable_norm[simp, intro]: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes [measurable]: "integrable M f" shows "integrable M (\x. norm (f x))" + using assms by (rule integrable_bound) auto + +lemma integrable_norm_cancel: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes [measurable]: "integrable M (\x. norm (f x))" "f \ borel_measurable M" shows "integrable M f" + using assms by (rule integrable_bound) auto + +lemma integrable_abs_cancel: + fixes f :: "'a \ real" + assumes [measurable]: "integrable M (\x. \f x\)" "f \ borel_measurable M" shows "integrable M f" + using assms by (rule integrable_bound) auto + +lemma integrable_max[simp, intro]: + fixes f :: "'a \ real" + assumes fg[measurable]: "integrable M f" "integrable M g" + shows "integrable M (\x. max (f x) (g x))" + using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] + by (rule integrable_bound) auto + +lemma integrable_min[simp, intro]: + fixes f :: "'a \ real" + assumes fg[measurable]: "integrable M f" "integrable M g" + shows "integrable M (\x. min (f x) (g x))" + using integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]] + by (rule integrable_bound) auto + +lemma integral_minus_iff[simp]: + "integrable M (\x. - f x ::'a::{banach, second_countable_topology}) \ integrable M f" + unfolding integrable_iff_bounded + by (auto intro: borel_measurable_uminus[of "\x. - f x" M, simplified]) + +lemma integrable_indicator_iff: + "integrable M (indicator A::_ \ real) \ A \ space M \ sets M \ emeasure M (A \ space M) < \" + by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ereal_indicator positive_integral_indicator' + cong: conj_cong) + +lemma integral_dominated_convergence: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" and w :: "'a \ real" + assumes "f \ borel_measurable M" "\i. s i \ borel_measurable M" "integrable M w" + assumes lim: "AE x in M. (\i. s i x) ----> f x" + assumes bound: "\i. AE x in M. norm (s i x) \ w x" + shows "integrable M f" + and "\i. integrable M (s i)" + and "(\i. integral\<^sup>L M (s i)) ----> integral\<^sup>L M f" +proof - + have "AE x in M. 0 \ w x" + using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans) + then have "(\\<^sup>+x. w x \M) = (\\<^sup>+x. norm (w x) \M)" + by (intro positive_integral_cong_AE) auto + with `integrable M w` have w: "w \ borel_measurable M" "(\\<^sup>+x. w x \M) < \" + unfolding integrable_iff_bounded by auto + + show int_s: "\i. integrable M (s i)" + unfolding integrable_iff_bounded + proof + fix i + have "(\\<^sup>+ x. ereal (norm (s i x)) \M) \ (\\<^sup>+x. w x \M)" + using bound by (intro positive_integral_mono_AE) auto + with w show "(\\<^sup>+ x. ereal (norm (s i x)) \M) < \" by auto + qed fact + + have all_bound: "AE x in M. \i. norm (s i x) \ w x" + using bound unfolding AE_all_countable by auto + + show int_f: "integrable M f" + unfolding integrable_iff_bounded + proof + have "(\\<^sup>+ x. ereal (norm (f x)) \M) \ (\\<^sup>+x. w x \M)" + using all_bound lim + proof (intro positive_integral_mono_AE, eventually_elim) + fix x assume "\i. norm (s i x) \ w x" "(\i. s i x) ----> f x" + then show "ereal (norm (f x)) \ ereal (w x)" + by (intro LIMSEQ_le_const2[where X="\i. ereal (norm (s i x))"] tendsto_intros lim_ereal[THEN iffD2]) auto + qed + with w show "(\\<^sup>+ x. ereal (norm (f x)) \M) < \" by auto + qed fact + + have "(\n. ereal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) ----> ereal 0" (is "?d ----> ereal 0") + proof (rule tendsto_sandwich) + show "eventually (\n. ereal 0 \ ?d n) sequentially" "(\_. ereal 0) ----> ereal 0" by auto + show "eventually (\n. ?d n \ (\\<^sup>+x. norm (s n x - f x) \M)) sequentially" + proof (intro always_eventually allI) + fix n + have "?d n = norm (integral\<^sup>L M (\x. s n x - f x))" + using int_f int_s by simp + also have "\ \ (\\<^sup>+x. norm (s n x - f x) \M)" + by (intro int_f int_s integrable_diff integral_norm_bound_ereal) + finally show "?d n \ (\\<^sup>+x. norm (s n x - f x) \M)" . + qed + show "(\n. \\<^sup>+x. norm (s n x - f x) \M) ----> ereal 0" + unfolding zero_ereal_def[symmetric] + apply (subst norm_minus_commute) + proof (rule positive_integral_dominated_convergence_norm[where w=w]) + show "\n. s n \ borel_measurable M" + using int_s unfolding integrable_iff_bounded by auto + qed fact+ + qed + then have "(\n. integral\<^sup>L M (s n) - integral\<^sup>L M f) ----> 0" + unfolding lim_ereal tendsto_norm_zero_iff . + from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]] + show "(\i. integral\<^sup>L M (s i)) ----> integral\<^sup>L M f" by simp +qed + +lemma integrable_mult_left_iff: + fixes f :: "'a \ real" + shows "integrable M (\x. c * f x) \ c = 0 \ integrable M f" + using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\x. c * f x"] + by (cases "c = 0") auto + +lemma positive_integral_eq_integral: + assumes f: "integrable M f" + assumes nonneg: "AE x in M. 0 \ f x" + shows "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" +proof - + { fix f :: "'a \ real" assume f: "f \ borel_measurable M" "\x. 0 \ f x" "integrable M f" + then have "(\\<^sup>+ x. f x \M) = integral\<^sup>L M f" + proof (induct rule: borel_measurable_induct_real) + case (set A) then show ?case + by (simp add: integrable_indicator_iff ereal_indicator emeasure_eq_ereal_measure) + next + case (mult f c) then show ?case + unfolding times_ereal.simps(1)[symmetric] + by (subst positive_integral_cmult) + (auto simp add: integrable_mult_left_iff zero_ereal_def[symmetric]) + next + case (add g f) + then have "integrable M f" "integrable M g" + by (auto intro!: integrable_bound[OF add(8)]) + with add show ?case + unfolding plus_ereal.simps(1)[symmetric] + by (subst positive_integral_add) auto + next + case (seq s) + { fix i x assume "x \ space M" with seq(4) have "s i x \ f x" + by (intro LIMSEQ_le_const[OF seq(5)] exI[of _ i]) (auto simp: incseq_def le_fun_def) } + note s_le_f = this + + show ?case + proof (rule LIMSEQ_unique) + show "(\i. ereal (integral\<^sup>L M (s i))) ----> ereal (integral\<^sup>L M f)" + unfolding lim_ereal + proof (rule integral_dominated_convergence[where w=f]) + show "integrable M f" by fact + from s_le_f seq show "\i. AE x in M. norm (s i x) \ f x" + by auto + qed (insert seq, auto) + have int_s: "\i. integrable M (s i)" + using seq f s_le_f by (intro integrable_bound[OF f(3)]) auto + have "(\i. \\<^sup>+ x. s i x \M) ----> \\<^sup>+ x. f x \M" + using seq s_le_f f + by (intro positive_integral_dominated_convergence[where w=f]) + (auto simp: integrable_iff_bounded) + also have "(\i. \\<^sup>+x. s i x \M) = (\i. \x. s i x \M)" + using seq int_s by simp + finally show "(\i. \x. s i x \M) ----> \\<^sup>+x. f x \M" + by simp + qed + qed } + from this[of "\x. max 0 (f x)"] assms have "(\\<^sup>+ x. max 0 (f x) \M) = integral\<^sup>L M (\x. max 0 (f x))" + by simp + also have "\ = integral\<^sup>L M f" + using assms by (auto intro!: integral_cong_AE) + also have "(\\<^sup>+ x. max 0 (f x) \M) = (\\<^sup>+ x. f x \M)" + using assms by (auto intro!: positive_integral_cong_AE simp: max_def) + finally show ?thesis . +qed + +lemma integral_norm_bound: + fixes f :: "_ \ 'a :: {banach, second_countable_topology}" + shows "integrable M f \ norm (integral\<^sup>L M f) \ (\x. norm (f x) \M)" + using positive_integral_eq_integral[of M "\x. norm (f x)"] + using integral_norm_bound_ereal[of M f] by simp + +lemma integral_eq_positive_integral: + "integrable M f \ (\x. 0 \ f x) \ integral\<^sup>L M f = real (\\<^sup>+ x. ereal (f x) \M)" + by (subst positive_integral_eq_integral) auto + +lemma integrableI_simple_bochner_integrable: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + shows "simple_bochner_integrable M f \ integrable M f" + by (intro integrableI_sequence[where s="\_. f"] borel_measurable_simple_function) + (auto simp: zero_ereal_def[symmetric] simple_bochner_integrable.simps) + +lemma integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes "integrable M f" + assumes base: "\A c. A \ sets M \ emeasure M A < \ \ P (\x. indicator A x *\<^sub>R c)" + assumes add: "\f g. integrable M f \ P f \ integrable M g \ P g \ P (\x. f x + g x)" + assumes lim: "\f s. (\i. integrable M (s i)) \ (\i. P (s i)) \ + (\x. x \ space M \ (\i. s i x) ----> f x) \ + (\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)) \ integrable M f \ P f" + shows "P f" +proof - + from `integrable M f` have f: "f \ borel_measurable M" "(\\<^sup>+x. norm (f x) \M) < \" + unfolding integrable_iff_bounded by auto + from borel_measurable_implies_sequence_metric[OF f(1)] + obtain s where s: "\i. simple_function M (s i)" "\x. x \ space M \ (\i. s i x) ----> f x" + "\i x. x \ space M \ norm (s i x) \ 2 * norm (f x)" + unfolding norm_conv_dist by metis + + { fix f A + have [simp]: "P (\x. 0)" + using base[of "{}" undefined] by simp + have "(\i::'b. i \ A \ integrable M (f i::'a \ 'b)) \ + (\i. i \ A \ P (f i)) \ P (\x. \i\A. f i x)" + by (induct A rule: infinite_finite_induct) (auto intro!: add) } + note setsum = this + + def s' \ "\i z. indicator (space M) z *\<^sub>R s i z" + then have s'_eq_s: "\i x. x \ space M \ s' i x = s i x" + by simp + + have sf[measurable]: "\i. simple_function M (s' i)" + unfolding s'_def using s(1) + by (intro simple_function_compose2[where h="op *\<^sub>R"] simple_function_indicator) auto + + { fix i + have "\z. {y. s' i z = y \ y \ s' i ` space M \ y \ 0 \ z \ space M} = + (if z \ space M \ s' i z \ 0 then {s' i z} else {})" + by (auto simp add: s'_def split: split_indicator) + then have "\z. s' i = (\z. \y\s' i`space M - {0}. indicator {x\space M. s' i x = y} z *\<^sub>R y)" + using sf by (auto simp: fun_eq_iff simple_function_def s'_def) } + note s'_eq = this + + show "P f" + proof (rule lim) + fix i + + have "(\\<^sup>+x. norm (s' i x) \M) \ (\\<^sup>+x. 2 * ereal (norm (f x)) \M)" + using s by (intro positive_integral_mono) (auto simp: s'_eq_s) + also have "\ < \" + using f by (subst positive_integral_cmult) auto + finally have sbi: "simple_bochner_integrable M (s' i)" + using sf by (intro simple_bochner_integrableI_bounded) auto + then show "integrable M (s' i)" + by (rule integrableI_simple_bochner_integrable) + + { fix x assume"x \ space M" "s' i x \ 0" + then have "emeasure M {y \ space M. s' i y = s' i x} \ emeasure M {y \ space M. s' i y \ 0}" + by (intro emeasure_mono) auto + also have "\ < \" + using sbi by (auto elim: simple_bochner_integrable.cases) + finally have "emeasure M {y \ space M. s' i y = s' i x} \ \" by simp } + then show "P (s' i)" + by (subst s'_eq) (auto intro!: setsum base) + + fix x assume "x \ space M" with s show "(\i. s' i x) ----> f x" + by (simp add: s'_eq_s) + show "norm (s' i x) \ 2 * norm (f x)" + using `x \ space M` s by (simp add: s'_eq_s) + qed fact +qed + +lemma integral_nonneg_AE: + fixes f :: "'a \ real" + assumes [measurable]: "integrable M f" "AE x in M. 0 \ f x" + shows "0 \ integral\<^sup>L M f" +proof - + have "0 \ ereal (integral\<^sup>L M (\x. max 0 (f x)))" + by (subst integral_eq_positive_integral) + (auto intro: real_of_ereal_pos positive_integral_positive integrable_max assms) + also have "integral\<^sup>L M (\x. max 0 (f x)) = integral\<^sup>L M f" + using assms(2) by (intro integral_cong_AE assms integrable_max) auto + finally show ?thesis + by simp +qed + +lemma integral_eq_zero_AE: + "f \ borel_measurable M \ (AE x in M. f x = 0) \ integral\<^sup>L M f = 0" + using integral_cong_AE[of f M "\_. 0"] by simp + +lemma integral_nonneg_eq_0_iff_AE: + fixes f :: "_ \ real" + assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \ f x" + shows "integral\<^sup>L M f = 0 \ (AE x in M. f x = 0)" +proof + assume "integral\<^sup>L M f = 0" + then have "integral\<^sup>P M f = 0" + using positive_integral_eq_integral[OF f nonneg] by simp + then have "AE x in M. ereal (f x) \ 0" + by (simp add: positive_integral_0_iff_AE) + with nonneg show "AE x in M. f x = 0" + by auto +qed (auto simp add: integral_eq_zero_AE) + +lemma integral_mono_AE: + fixes f :: "'a \ real" + assumes "integrable M f" "integrable M g" "AE x in M. f x \ g x" + shows "integral\<^sup>L M f \ integral\<^sup>L M g" +proof - + have "0 \ integral\<^sup>L M (\x. g x - f x)" + using assms by (intro integral_nonneg_AE integrable_diff assms) auto + also have "\ = integral\<^sup>L M g - integral\<^sup>L M f" + by (intro integral_diff assms) + finally show ?thesis by simp +qed + +lemma integral_mono: + fixes f :: "'a \ real" + shows "integrable M f \ integrable M g \ (\x. x \ space M \ f x \ g x) \ + integral\<^sup>L M f \ integral\<^sup>L M g" + by (intro integral_mono_AE) auto + +section {* Measure spaces with an associated density *} + +lemma integrable_density: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" + assumes [measurable]: "f \ borel_measurable M" "g \ borel_measurable M" + and nn: "AE x in M. 0 \ g x" + shows "integrable (density M g) f \ integrable M (\x. g x *\<^sub>R f x)" + unfolding integrable_iff_bounded using nn + apply (simp add: positive_integral_density ) + apply (intro arg_cong2[where f="op ="] refl positive_integral_cong_AE) + apply auto + done + +lemma integral_density: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" + assumes f: "f \ borel_measurable M" + and g[measurable]: "g \ borel_measurable M" "AE x in M. 0 \ g x" + shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\x. g x *\<^sub>R f x)" +proof (rule integral_eq_cases) + assume "integrable (density M g) f" + then show ?thesis + proof induct + case (base A c) + then have [measurable]: "A \ sets M" by auto + + have int: "integrable M (\x. g x * indicator A x)" + using g base integrable_density[of "indicator A :: 'a \ real" M g] by simp + then have "integral\<^sup>L M (\x. g x * indicator A x) = (\\<^sup>+ x. ereal (g x * indicator A x) \M)" + using g by (subst positive_integral_eq_integral) auto + also have "\ = (\\<^sup>+ x. ereal (g x) * indicator A x \M)" + by (intro positive_integral_cong) (auto split: split_indicator) + also have "\ = emeasure (density M g) A" + by (rule emeasure_density[symmetric]) auto + also have "\ = ereal (measure (density M g) A)" + using base by (auto intro: emeasure_eq_ereal_measure) + also have "\ = integral\<^sup>L (density M g) (indicator A)" + using base by simp + finally show ?case + using base by (simp add: int) + next + case (add f h) + then have [measurable]: "f \ borel_measurable M" "h \ borel_measurable M" + by (auto dest!: borel_measurable_integrable) + from add g show ?case + by (simp add: scaleR_add_right integrable_density) + next + case (lim f s) + have [measurable]: "f \ borel_measurable M" "\i. s i \ borel_measurable M" + using lim(1,5)[THEN borel_measurable_integrable] by auto + + show ?case + proof (rule LIMSEQ_unique) + show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) ----> integral\<^sup>L M (\x. g x *\<^sub>R f x)" + proof (rule integral_dominated_convergence(3)) + show "integrable M (\x. 2 * norm (g x *\<^sub>R f x))" + by (intro integrable_mult_right integrable_norm integrable_density[THEN iffD1] lim g) auto + show "AE x in M. (\i. g x *\<^sub>R s i x) ----> g x *\<^sub>R f x" + using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M]) + show "\i. AE x in M. norm (g x *\<^sub>R s i x) \ 2 * norm (g x *\<^sub>R f x)" + using lim(4) g by (auto intro!: AE_I2[of M] mult_left_mono simp: field_simps) + qed auto + show "(\i. integral\<^sup>L M (\x. g x *\<^sub>R s i x)) ----> integral\<^sup>L (density M g) f" + unfolding lim(2)[symmetric] + by (rule integral_dominated_convergence(3)[where w="\x. 2 * norm (f x)"]) + (insert lim(3-5), auto intro: integrable_norm) + qed + qed +qed (simp add: f g integrable_density) + +lemma + fixes g :: "'a \ real" + assumes "f \ borel_measurable M" "AE x in M. 0 \ f x" "g \ borel_measurable M" + shows integral_real_density: "integral\<^sup>L (density M f) g = (\ x. f x * g x \M)" + and integrable_real_density: "integrable (density M f) g \ integrable M (\x. f x * g x)" + using assms integral_density[of g M f] integrable_density[of g M f] by auto + +lemma has_bochner_integral_density: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" + shows "f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. 0 \ g x) \ + has_bochner_integral M (\x. g x *\<^sub>R f x) x \ has_bochner_integral (density M g) f x" + by (simp add: has_bochner_integral_iff integrable_density integral_density) + +subsection {* Distributions *} + +lemma integrable_distr_eq: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes [measurable]: "g \ measurable M N" "f \ borel_measurable N" + shows "integrable (distr M N g) f \ integrable M (\x. f (g x))" + unfolding integrable_iff_bounded by (simp_all add: positive_integral_distr) + +lemma integrable_distr: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + shows "T \ measurable M M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" + by (subst integrable_distr_eq[symmetric, where g=T]) + (auto dest: borel_measurable_integrable) + +lemma integral_distr: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes g[measurable]: "g \ measurable M N" and f: "f \ borel_measurable N" + shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\x. f (g x))" +proof (rule integral_eq_cases) + assume "integrable (distr M N g) f" + then show ?thesis + proof induct + case (base A c) + then have [measurable]: "A \ sets N" by auto + from base have int: "integrable (distr M N g) (\a. indicator A a *\<^sub>R c)" + by (intro integrable_indicator) + + have "integral\<^sup>L (distr M N g) (\a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c" + using base by (subst integral_indicator) auto + also have "\ = measure M (g -` A \ space M) *\<^sub>R c" + by (subst measure_distr) auto + also have "\ = integral\<^sup>L M (\a. indicator (g -` A \ space M) a *\<^sub>R c)" + using base by (subst integral_indicator) (auto simp: emeasure_distr) + also have "\ = integral\<^sup>L M (\a. indicator A (g a) *\<^sub>R c)" + using int base by (intro integral_cong_AE) (auto simp: emeasure_distr split: split_indicator) + finally show ?case . + next + case (add f h) + then have [measurable]: "f \ borel_measurable N" "h \ borel_measurable N" + by (auto dest!: borel_measurable_integrable) + from add g show ?case + by (simp add: scaleR_add_right integrable_distr_eq) + next + case (lim f s) + have [measurable]: "f \ borel_measurable N" "\i. s i \ borel_measurable N" + using lim(1,5)[THEN borel_measurable_integrable] by auto + + show ?case + proof (rule LIMSEQ_unique) + show "(\i. integral\<^sup>L M (\x. s i (g x))) ----> integral\<^sup>L M (\x. f (g x))" + proof (rule integral_dominated_convergence(3)) + show "integrable M (\x. 2 * norm (f (g x)))" + using lim by (auto intro!: integrable_norm simp: integrable_distr_eq) + show "AE x in M. (\i. s i (g x)) ----> f (g x)" + using lim(3) g[THEN measurable_space] by auto + show "\i. AE x in M. norm (s i (g x)) \ 2 * norm (f (g x))" + using lim(4) g[THEN measurable_space] by auto + qed auto + show "(\i. integral\<^sup>L M (\x. s i (g x))) ----> integral\<^sup>L (distr M N g) f" + unfolding lim(2)[symmetric] + by (rule integral_dominated_convergence(3)[where w="\x. 2 * norm (f x)"]) + (insert lim(3-5), auto intro: integrable_norm) + qed + qed +qed (simp add: f g integrable_distr_eq) + +lemma has_bochner_integral_distr: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + shows "f \ borel_measurable N \ g \ measurable M N \ + has_bochner_integral M (\x. f (g x)) x \ has_bochner_integral (distr M N g) f x" + by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) + +section {* Lebesgue integration on @{const count_space} *} + +lemma integrable_count_space: + fixes f :: "'a \ 'b::{banach,second_countable_topology}" + shows "finite X \ integrable (count_space X) f" + by (auto simp: positive_integral_count_space integrable_iff_bounded) + +lemma measure_count_space[simp]: + "B \ A \ finite B \ measure (count_space A) B = card B" + unfolding measure_def by (subst emeasure_count_space ) auto + +lemma lebesgue_integral_count_space_finite_support: + assumes f: "finite {a\A. f a \ 0}" + shows "(\x. f x \count_space A) = (\a | a \ A \ f a \ 0. f a)" +proof - + have eq: "\x. x \ A \ (\a | x = a \ a \ A \ f a \ 0. f a) = (\x\{x}. f x)" + by (intro setsum_mono_zero_cong_left) auto + + have "(\x. f x \count_space A) = (\x. (\a | a \ A \ f a \ 0. indicator {a} x *\<^sub>R f a) \count_space A)" + by (intro integral_cong refl) (simp add: f eq) + also have "\ = (\a | a \ A \ f a \ 0. measure (count_space A) {a} *\<^sub>R f a)" + by (subst integral_setsum) (auto intro!: setsum_cong) + finally show ?thesis + by auto +qed + +lemma lebesgue_integral_count_space_finite: "finite A \ (\x. f x \count_space A) = (\a\A. f a)" + by (subst lebesgue_integral_count_space_finite_support) + (auto intro!: setsum_mono_zero_cong_left) + +section {* Point measure *} + +lemma lebesgue_integral_point_measure_finite: + fixes g :: "'a \ 'b::{banach, second_countable_topology}" + shows "finite A \ (\a. a \ A \ 0 \ f a) \ + integral\<^sup>L (point_measure A f) g = (\a\A. f a *\<^sub>R g a)" + by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) + +lemma integrable_point_measure_finite: + fixes g :: "'a \ 'b::{banach, second_countable_topology}" and f :: "'a \ real" + shows "finite A \ integrable (point_measure A f) g" + unfolding point_measure_def + apply (subst density_ereal_max_0) + apply (subst integrable_density) + apply (auto simp: AE_count_space integrable_count_space) + done + +subsection {* Legacy lemmas for the real-valued Lebesgue integral\<^sup>L *} + +lemma real_lebesgue_integral_def: + assumes f: "integrable M f" + shows "integral\<^sup>L M f = real (\\<^sup>+x. f x \M) - real (\\<^sup>+x. - f x \M)" +proof - + have "integral\<^sup>L M f = integral\<^sup>L M (\x. max 0 (f x) - max 0 (- f x))" + by (auto intro!: arg_cong[where f="integral\<^sup>L M"]) + also have "\ = integral\<^sup>L M (\x. max 0 (f x)) - integral\<^sup>L M (\x. max 0 (- f x))" + by (intro integral_diff integrable_max integrable_minus integrable_zero f) + also have "integral\<^sup>L M (\x. max 0 (f x)) = real (\\<^sup>+x. max 0 (f x) \M)" + by (subst integral_eq_positive_integral[symmetric]) (auto intro!: integrable_max f) + also have "integral\<^sup>L M (\x. max 0 (- f x)) = real (\\<^sup>+x. max 0 (- f x) \M)" + by (subst integral_eq_positive_integral[symmetric]) (auto intro!: integrable_max f) + also have "(\x. ereal (max 0 (f x))) = (\x. max 0 (ereal (f x)))" + by (auto simp: max_def) + also have "(\x. ereal (max 0 (- f x))) = (\x. max 0 (- ereal (f x)))" + by (auto simp: max_def) + finally show ?thesis + unfolding positive_integral_max_0 . +qed + +lemma real_integrable_def: + "integrable M f \ f \ borel_measurable M \ + (\\<^sup>+ x. ereal (f x) \M) \ \ \ (\\<^sup>+ x. ereal (- f x) \M) \ \" + unfolding integrable_iff_bounded +proof (safe del: notI) + assume *: "(\\<^sup>+ x. ereal (norm (f x)) \M) < \" + have "(\\<^sup>+ x. ereal (f x) \M) \ (\\<^sup>+ x. ereal (norm (f x)) \M)" + by (intro positive_integral_mono) auto + also note * + finally show "(\\<^sup>+ x. ereal (f x) \M) \ \" + by simp + have "(\\<^sup>+ x. ereal (- f x) \M) \ (\\<^sup>+ x. ereal (norm (f x)) \M)" + by (intro positive_integral_mono) auto + also note * + finally show "(\\<^sup>+ x. ereal (- f x) \M) \ \" + by simp +next + assume [measurable]: "f \ borel_measurable M" + assume fin: "(\\<^sup>+ x. ereal (f x) \M) \ \" "(\\<^sup>+ x. ereal (- f x) \M) \ \" + have "(\\<^sup>+ x. norm (f x) \M) = (\\<^sup>+ x. max 0 (ereal (f x)) + max 0 (ereal (- f x)) \M)" + by (intro positive_integral_cong) (auto simp: max_def) + also have"\ = (\\<^sup>+ x. max 0 (ereal (f x)) \M) + (\\<^sup>+ x. max 0 (ereal (- f x)) \M)" + by (intro positive_integral_add) auto + also have "\ < \" + using fin by (auto simp: positive_integral_max_0) + finally show "(\\<^sup>+ x. norm (f x) \M) < \" . +qed + +lemma integrableD[dest]: + assumes "integrable M f" + shows "f \ borel_measurable M" "(\\<^sup>+ x. ereal (f x) \M) \ \" "(\\<^sup>+ x. ereal (- f x) \M) \ \" + using assms unfolding real_integrable_def by auto + +lemma integrableE: + assumes "integrable M f" + obtains r q where + "(\\<^sup>+x. ereal (f x)\M) = ereal r" + "(\\<^sup>+x. ereal (-f x)\M) = ereal q" + "f \ borel_measurable M" "integral\<^sup>L M f = r - q" + using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms] + using positive_integral_positive[of M "\x. ereal (f x)"] + using positive_integral_positive[of M "\x. ereal (-f x)"] + by (cases rule: ereal2_cases[of "(\\<^sup>+x. ereal (-f x)\M)" "(\\<^sup>+x. ereal (f x)\M)"]) auto + +lemma integral_monotone_convergence_nonneg: + fixes f :: "nat \ 'a \ real" + assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" + and pos: "\i. AE x in M. 0 \ f i x" + and lim: "AE x in M. (\i. f i x) ----> u x" + and ilim: "(\i. integral\<^sup>L M (f i)) ----> x" + and u: "u \ borel_measurable M" + shows "integrable M u" + and "integral\<^sup>L M u = x" +proof - + have "(\\<^sup>+ x. ereal (u x) \M) = (SUP n. (\\<^sup>+ x. ereal (f n x) \M))" + proof (subst positive_integral_monotone_convergence_SUP_AE[symmetric]) + fix i + from mono pos show "AE x in M. ereal (f i x) \ ereal (f (Suc i) x) \ 0 \ ereal (f i x)" + by eventually_elim (auto simp: mono_def) + show "(\x. ereal (f i x)) \ borel_measurable M" + using i by auto + next + show "(\\<^sup>+ x. ereal (u x) \M) = \\<^sup>+ x. (SUP i. ereal (f i x)) \M" + apply (rule positive_integral_cong_AE) + using lim mono + by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2]) + qed + also have "\ = ereal x" + using mono i unfolding positive_integral_eq_integral[OF i pos] + by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim) + finally have "(\\<^sup>+ x. ereal (u x) \M) = ereal x" . + moreover have "(\\<^sup>+ x. ereal (- u x) \M) = 0" + proof (subst positive_integral_0_iff_AE) + show "(\x. ereal (- u x)) \ borel_measurable M" + using u by auto + from mono pos[of 0] lim show "AE x in M. ereal (- u x) \ 0" + proof eventually_elim + fix x assume "mono (\n. f n x)" "0 \ f 0 x" "(\i. f i x) ----> u x" + then show "ereal (- u x) \ 0" + using incseq_le[of "\n. f n x" "u x" 0] by (simp add: mono_def incseq_def) + qed + qed + ultimately show "integrable M u" "integral\<^sup>L M u = x" + by (auto simp: real_integrable_def real_lebesgue_integral_def u) +qed + +lemma + fixes f :: "nat \ 'a \ real" + assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" + and lim: "AE x in M. (\i. f i x) ----> u x" + and ilim: "(\i. integral\<^sup>L M (f i)) ----> x" + and u: "u \ borel_measurable M" + shows integrable_monotone_convergence: "integrable M u" + and integral_monotone_convergence: "integral\<^sup>L M u = x" + and has_bochner_integral_monotone_convergence: "has_bochner_integral M u x" +proof - + have 1: "\i. integrable M (\x. f i x - f 0 x)" + using f by auto + have 2: "AE x in M. mono (\n. f n x - f 0 x)" + using mono by (auto simp: mono_def le_fun_def) + have 3: "\n. AE x in M. 0 \ f n x - f 0 x" + using mono by (auto simp: field_simps mono_def le_fun_def) + have 4: "AE x in M. (\i. f i x - f 0 x) ----> u x - f 0 x" + using lim by (auto intro!: tendsto_diff) + have 5: "(\i. (\x. f i x - f 0 x \M)) ----> x - integral\<^sup>L M (f 0)" + using f ilim by (auto intro!: tendsto_diff) + have 6: "(\x. u x - f 0 x) \ borel_measurable M" + using f[of 0] u by auto + note diff = integral_monotone_convergence_nonneg[OF 1 2 3 4 5 6] + have "integrable M (\x. (u x - f 0 x) + f 0 x)" + using diff(1) f by (rule integrable_add) + with diff(2) f show "integrable M u" "integral\<^sup>L M u = x" + by auto + then show "has_bochner_integral M u x" + by (metis has_bochner_integral_integrable) +qed + +lemma integral_norm_eq_0_iff: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes f[measurable]: "integrable M f" + shows "(\x. norm (f x) \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" +proof - + have "(\\<^sup>+x. norm (f x) \M) = (\x. norm (f x) \M)" + using f by (intro positive_integral_eq_integral integrable_norm) auto + then have "(\x. norm (f x) \M) = 0 \ (\\<^sup>+x. norm (f x) \M) = 0" + by simp + also have "\ \ emeasure M {x\space M. ereal (norm (f x)) \ 0} = 0" + by (intro positive_integral_0_iff) auto + finally show ?thesis + by simp +qed + +lemma integral_0_iff: + fixes f :: "'a \ real" + shows "integrable M f \ (\x. abs (f x) \M) = 0 \ emeasure M {x\space M. f x \ 0} = 0" + using integral_norm_eq_0_iff[of M f] by simp + +lemma (in finite_measure) lebesgue_integral_const[simp]: + "(\x. a \M) = measure M (space M) *\<^sub>R a" + using integral_indicator[of "space M" M a] + by (simp del: integral_indicator integral_scaleR_left cong: integral_cong) + +lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\x. a)" + using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong) + +lemma (in finite_measure) integrable_const_bound: + fixes f :: "'a \ 'b::{banach,second_countable_topology}" + shows "AE x in M. norm (f x) \ B \ f \ borel_measurable M \ integrable M f" + apply (rule integrable_bound[OF integrable_const[of B], of f]) + apply assumption + apply (cases "0 \ B") + apply auto + done + +lemma (in finite_measure) integral_less_AE: + fixes X Y :: "'a \ real" + assumes int: "integrable M X" "integrable M Y" + assumes A: "(emeasure M) A \ 0" "A \ sets M" "AE x in M. x \ A \ X x \ Y x" + assumes gt: "AE x in M. X x \ Y x" + shows "integral\<^sup>L M X < integral\<^sup>L M Y" +proof - + have "integral\<^sup>L M X \ integral\<^sup>L M Y" + using gt int by (intro integral_mono_AE) auto + moreover + have "integral\<^sup>L M X \ integral\<^sup>L M Y" + proof + assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y" + have "integral\<^sup>L M (\x. \Y x - X x\) = integral\<^sup>L M (\x. Y x - X x)" + using gt int by (intro integral_cong_AE) auto + also have "\ = 0" + using eq int by simp + finally have "(emeasure M) {x \ space M. Y x - X x \ 0} = 0" + using int by (simp add: integral_0_iff) + moreover + have "(\\<^sup>+x. indicator A x \M) \ (\\<^sup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" + using A by (intro positive_integral_mono_AE) auto + then have "(emeasure M) A \ (emeasure M) {x \ space M. Y x - X x \ 0}" + using int A by (simp add: integrable_def) + ultimately have "emeasure M A = 0" + using emeasure_nonneg[of M A] by simp + with `(emeasure M) A \ 0` show False by auto + qed + ultimately show ?thesis by auto +qed + +lemma (in finite_measure) integral_less_AE_space: + fixes X Y :: "'a \ real" + assumes int: "integrable M X" "integrable M Y" + assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \ 0" + shows "integral\<^sup>L M X < integral\<^sup>L M Y" + using gt by (intro integral_less_AE[OF int, where A="space M"]) auto + +(* GENERALIZE to banach, sct *) +lemma integral_sums: + fixes f :: "nat \ 'a \ real" + assumes integrable[measurable]: "\i. integrable M (f i)" + and summable: "\x. x \ space M \ summable (\i. \f i x\)" + and sums: "summable (\i. (\x. \f i x\ \M))" + shows "integrable M (\x. (\i. f i x))" (is "integrable M ?S") + and "(\i. integral\<^sup>L M (f i)) sums (\x. (\i. f i x) \M)" (is ?integral) +proof - + have "\x\space M. \w. (\i. \f i x\) sums w" + using summable unfolding summable_def by auto + from bchoice[OF this] + obtain w where w: "\x. x \ space M \ (\i. \f i x\) sums w x" by auto + then have w_borel: "w \ borel_measurable M" unfolding sums_def + by (rule borel_measurable_LIMSEQ) auto + + let ?w = "\y. if y \ space M then w y else 0" + + obtain x where abs_sum: "(\i. (\x. \f i x\ \M)) sums x" + using sums unfolding summable_def .. + + have 1: "\n. integrable M (\x. \ij. AE x in M. norm (\i ?w x" + using AE_space + proof eventually_elim + fix j x assume [simp]: "x \ space M" + have "\\i \ (\if i x\)" by (rule setsum_abs) + also have "\ \ w x" using w[of x] setsum_le_suminf[of "\i. \f i x\"] unfolding sums_iff by auto + finally show "norm (\i ?w x" by simp + qed + + have 3: "integrable M ?w" + proof (rule integrable_monotone_convergence(1)) + let ?F = "\n y. (\if i y\)" + let ?w' = "\n y. if y \ space M then ?F n y else 0" + have "\n. integrable M (?F n)" + using integrable by (auto intro!: integrable_abs) + thus "\n. integrable M (?w' n)" by (simp cong: integrable_cong) + show "AE x in M. mono (\n. ?w' n x)" + by (auto simp: mono_def le_fun_def intro!: setsum_mono2) + show "AE x in M. (\n. ?w' n x) ----> ?w x" + using w by (simp_all add: tendsto_const sums_def) + have *: "\n. integral\<^sup>L M (?w' n) = (\i< n. (\x. \f i x\ \M))" + using integrable by (simp add: integrable_abs cong: integral_cong) + from abs_sum + show "(\i. integral\<^sup>L M (?w' i)) ----> x" unfolding * sums_def . + qed (simp add: w_borel measurable_If_set) + + from summable[THEN summable_rabs_cancel] + have 4: "AE x in M. (\n. \i (\i. f i x)" + by (auto intro: summable_LIMSEQ) + + note int = integral_dominated_convergence(1,3)[OF _ _ 3 4 2] + + from int show "integrable M ?S" by simp + + show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF integrable] + using int(2) by simp +qed + +lemma integrable_mult_indicator: + fixes f :: "'a \ real" + shows "A \ sets M \ integrable M f \ integrable M (\x. f x * indicator A x)" + by (rule integrable_bound[where f="\x. \f x\"]) + (auto intro: integrable_abs split: split_indicator) + +lemma tendsto_integral_at_top: + fixes f :: "real \ real" + assumes M: "sets M = sets borel" and f[measurable]: "integrable M f" + shows "((\y. \ x. f x * indicator {.. y} x \M) ---> \ x. f x \M) at_top" +proof - + have M_measure[simp]: "borel_measurable M = borel_measurable borel" + using M by (simp add: sets_eq_imp_space_eq measurable_def) + { fix f :: "real \ real" assume f: "integrable M f" "\x. 0 \ f x" + then have [measurable]: "f \ borel_measurable borel" + by (simp add: real_integrable_def) + have "((\y. \ x. f x * indicator {.. y} x \M) ---> \ x. f x \M) at_top" + proof (rule tendsto_at_topI_sequentially) + have int: "\n. integrable M (\x. f x * indicator {.. n} x)" + by (rule integrable_mult_indicator) (auto simp: M f) + show "(\n. \ x. f x * indicator {..real n} x \M) ----> integral\<^sup>L M f" + proof (rule integral_dominated_convergence) + { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" + by (rule eventually_sequentiallyI[of "natceiling x"]) + (auto split: split_indicator simp: natceiling_le_eq) } + from filterlim_cong[OF refl refl this] + show "AE x in M. (\n. f x * indicator {..real n} x) ----> f x" + by (simp add: tendsto_const) + show "\j. AE x in M. norm (f x * indicator {.. j} x) \ f x" + using f(2) by (intro AE_I2) (auto split: split_indicator) + qed (simp | fact)+ + show "mono (\y. \ x. f x * indicator {..y} x \M)" + by (intro monoI integral_mono int) (auto split: split_indicator intro: f) + qed } + note nonneg = this + let ?P = "\y. \ x. max 0 (f x) * indicator {..y} x \M" + let ?N = "\y. \ x. max 0 (- f x) * indicator {..y} x \M" + let ?p = "integral\<^sup>L M (\x. max 0 (f x))" + let ?n = "integral\<^sup>L M (\x. max 0 (- f x))" + have "(?P ---> ?p) at_top" "(?N ---> ?n) at_top" + by (auto intro!: nonneg integrable_max f) + note tendsto_diff[OF this] + also have "(\y. ?P y - ?N y) = (\y. \ x. f x * indicator {..y} x \M)" + by (subst integral_diff[symmetric]) + (auto intro!: integrable_mult_indicator integrable_max f integral_cong + simp: M split: split_max) + also have "?p - ?n = integral\<^sup>L M f" + by (subst integral_diff[symmetric]) + (auto intro!: integrable_max f integral_cong simp: M split: split_max) + finally show ?thesis . +qed + +lemma + fixes f :: "real \ real" + assumes M: "sets M = sets borel" + assumes nonneg: "AE x in M. 0 \ f x" + assumes borel: "f \ borel_measurable borel" + assumes int: "\y. integrable M (\x. f x * indicator {.. y} x)" + assumes conv: "((\y. \ x. f x * indicator {.. y} x \M) ---> x) at_top" + shows has_bochner_integral_monotone_convergence_at_top: "has_bochner_integral M f x" + and integrable_monotone_convergence_at_top: "integrable M f" + and integral_monotone_convergence_at_top:"integral\<^sup>L M f = x" +proof - + from nonneg have "AE x in M. mono (\n::nat. f x * indicator {..real n} x)" + by (auto split: split_indicator intro!: monoI) + { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" + by (rule eventually_sequentiallyI[of "natceiling x"]) + (auto split: split_indicator simp: natceiling_le_eq) } + from filterlim_cong[OF refl refl this] + have "AE x in M. (\i. f x * indicator {..real i} x) ----> f x" + by (simp add: tendsto_const) + have "(\i. \ x. f x * indicator {..real i} x \M) ----> x" + using conv filterlim_real_sequentially by (rule filterlim_compose) + have M_measure[simp]: "borel_measurable M = borel_measurable borel" + using M by (simp add: sets_eq_imp_space_eq measurable_def) + have "f \ borel_measurable M" + using borel by simp + show "has_bochner_integral M f x" + by (rule has_bochner_integral_monotone_convergence) fact+ + then show "integrable M f" "integral\<^sup>L M f = x" + by (auto simp: _has_bochner_integral_iff) +qed + + +section "Lebesgue integration on countable spaces" + +lemma integral_on_countable: + fixes f :: "real \ real" + assumes f: "f \ borel_measurable M" + and bij: "bij_betw enum S (f ` space M)" + and enum_zero: "enum ` (-S) \ {0}" + and fin: "\x. x \ 0 \ (emeasure M) (f -` {x} \ space M) \ \" + and abs_summable: "summable (\r. \enum r * real ((emeasure M) (f -` {enum r} \ space M))\)" + shows "integrable M f" + and "(\r. enum r * real ((emeasure M) (f -` {enum r} \ space M))) sums integral\<^sup>L M f" (is ?sums) +proof - + let ?A = "\r. f -` {enum r} \ space M" + let ?F = "\r x. enum r * indicator (?A r) x" + have enum_eq: "\r. enum r * real ((emeasure M) (?A r)) = integral\<^sup>L M (?F r)" + using f fin by (simp add: measure_def cong: disj_cong) + + { fix x assume "x \ space M" + hence "f x \ enum ` S" using bij unfolding bij_betw_def by auto + then obtain i where "i\S" "enum i = f x" by auto + have F: "\j. ?F j x = (if j = i then f x else 0)" + proof cases + fix j assume "j = i" + thus "?thesis j" using `x \ space M` `enum i = f x` by (simp add: indicator_def) + next + fix j assume "j \ i" + show "?thesis j" using bij `i \ S` `j \ i` `enum i = f x` enum_zero + by (cases "j \ S") (auto simp add: indicator_def bij_betw_def inj_on_def) + qed + hence F_abs: "\j. \if j = i then f x else 0\ = (if j = i then \f x\ else 0)" by auto + have "(\i. ?F i x) sums f x" + "(\i. \?F i x\) sums \f x\" + by (auto intro!: sums_single simp: F F_abs) } + note F_sums_f = this(1) and F_abs_sums_f = this(2) + + have int_f: "integral\<^sup>L M f = (\x. (\r. ?F r x) \M)" "integrable M f = integrable M (\x. \r. ?F r x)" + using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff) + + { fix r + have "(\x. \?F r x\ \M) = (\x. \enum r\ * indicator (?A r) x \M)" + by (auto simp: indicator_def intro!: integral_cong) + also have "\ = \enum r\ * real ((emeasure M) (?A r))" + using f fin by (simp add: measure_def cong: disj_cong) + finally have "(\x. \?F r x\ \M) = \enum r * real ((emeasure M) (?A r))\" + using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) } + note int_abs_F = this + + have 1: "\i. integrable M (\x. ?F i x)" + using f fin by auto + + have 2: "\x. x \ space M \ summable (\i. \?F i x\)" + using F_abs_sums_f unfolding sums_iff by auto + + from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] + show ?sums unfolding enum_eq int_f by simp + + from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] + show "integrable M f" unfolding int_f by simp +qed + +subsection {* Product measure *} + +lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]: + fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" + assumes [measurable]: "split f \ borel_measurable (N \\<^sub>M M)" + shows "Measurable.pred N (\x. integrable M (f x))" +proof - + have [simp]: "\x. x \ space N \ integrable M (f x) \ (\\<^sup>+y. norm (f x y) \M) < \" + unfolding integrable_iff_bounded by simp + show ?thesis + by (simp cong: measurable_cong) +qed + +lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]: + "(\x. x \ space N \ A x \ space M) \ + {x \ space (N \\<^sub>M M). snd x \ A (fst x)} \ sets (N \\<^sub>M M) \ + (\x. measure M (A x)) \ borel_measurable N" + unfolding measure_def by (intro measurable_emeasure borel_measurable_real_of_ereal) + +lemma Collect_subset [simp]: "{x\A. P x} \ A" by auto + +lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]: + fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" + assumes f[measurable]: "split f \ borel_measurable (N \\<^sub>M M)" + shows "(\x. \y. f x y \M) \ borel_measurable N" +proof - + from borel_measurable_implies_sequence_metric[OF f, of 0] guess s .. + then have s: "\i. simple_function (N \\<^sub>M M) (s i)" + "\x y. x \ space N \ y \ space M \ (\i. s i (x, y)) ----> f x y" + "\i x y. x \ space N \ y \ space M \ norm (s i (x, y)) \ 2 * norm (f x y)" + by (auto simp: space_pair_measure norm_conv_dist) + + have [measurable]: "\i. s i \ borel_measurable (N \\<^sub>M M)" + by (rule borel_measurable_simple_function) fact + + have "\i. s i \ measurable (N \\<^sub>M M) (count_space UNIV)" + by (rule measurable_simple_function) fact + + def f' \ "\i x. if integrable M (f x) + then simple_bochner_integral M (\y. s i (x, y)) + else (THE x. False)" + + { fix i x assume "x \ space N" + then have "simple_bochner_integral M (\y. s i (x, y)) = + (\z\s i ` (space N \ space M). measure M {y \ space M. s i (x, y) = z} *\<^sub>R z)" + using s(1)[THEN simple_functionD(1)] + unfolding simple_bochner_integral_def + by (intro setsum_mono_zero_cong_left) + (auto simp: eq_commute space_pair_measure image_iff cong: conj_cong) } + note eq = this + + show ?thesis + proof (rule borel_measurable_LIMSEQ_metric) + fix i show "f' i \ borel_measurable N" + unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong) + next + fix x assume x: "x \ space N" + { assume int_f: "integrable M (f x)" + have int_2f: "integrable M (\y. 2 * norm (f x y))" + by (intro integrable_norm integrable_mult_right int_f) + have "(\i. integral\<^sup>L M (\y. s i (x, y))) ----> integral\<^sup>L M (f x)" + proof (rule integral_dominated_convergence) + from int_f show "f x \ borel_measurable M" by auto + show "\i. (\y. s i (x, y)) \ borel_measurable M" + using x by simp + show "AE xa in M. (\i. s i (x, xa)) ----> f x xa" + using x s(2) by auto + show "\i. AE xa in M. norm (s i (x, xa)) \ 2 * norm (f x xa)" + using x s(3) by auto + qed fact + moreover + { fix i + have "simple_bochner_integrable M (\y. s i (x, y))" + proof (rule simple_bochner_integrableI_bounded) + have "(\y. s i (x, y)) ` space M \ s i ` (space N \ space M)" + using x by auto + then show "simple_function M (\y. s i (x, y))" + using simple_functionD(1)[OF s(1), of i] x + by (intro simple_function_borel_measurable) + (auto simp: space_pair_measure dest: finite_subset) + have "(\\<^sup>+ y. ereal (norm (s i (x, y))) \M) \ (\\<^sup>+ y. 2 * norm (f x y) \M)" + using x s by (intro positive_integral_mono) auto + also have "(\\<^sup>+ y. 2 * norm (f x y) \M) < \" + using int_2f by (simp add: integrable_iff_bounded) + finally show "(\\<^sup>+ xa. ereal (norm (s i (x, xa))) \M) < \" . + qed + then have "integral\<^sup>L M (\y. s i (x, y)) = simple_bochner_integral M (\y. s i (x, y))" + by (rule simple_bochner_integrable_eq_integral[symmetric]) } + ultimately have "(\i. simple_bochner_integral M (\y. s i (x, y))) ----> integral\<^sup>L M (f x)" + by simp } + then + show "(\i. f' i x) ----> integral\<^sup>L M (f x)" + unfolding f'_def + by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq tendsto_const) + qed +qed + +lemma (in pair_sigma_finite) integrable_product_swap: + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes "integrable (M1 \\<^sub>M M2) f" + shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x))" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) + show ?thesis unfolding * + by (rule integrable_distr[OF measurable_pair_swap']) + (simp add: distr_pair_swap[symmetric] assms) +qed + +lemma (in pair_sigma_finite) integrable_product_swap_iff: + fixes f :: "_ \ _::{banach, second_countable_topology}" + shows "integrable (M2 \\<^sub>M M1) (\(x,y). f (y,x)) \ integrable (M1 \\<^sub>M M2) f" +proof - + interpret Q: pair_sigma_finite M2 M1 by default + from Q.integrable_product_swap[of "\(x,y). f (y,x)"] integrable_product_swap[of f] + show ?thesis by auto +qed + +lemma (in pair_sigma_finite) integral_product_swap: + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes f: "f \ borel_measurable (M1 \\<^sub>M M2)" + shows "(\(x,y). f (y,x) \(M2 \\<^sub>M M1)) = integral\<^sup>L (M1 \\<^sub>M M2) f" +proof - + have *: "(\(x,y). f (y,x)) = (\x. f (case x of (x,y)\(y,x)))" by (auto simp: fun_eq_iff) + show ?thesis unfolding * + by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) +qed + +lemma (in pair_sigma_finite) emeasure_pair_measure_finite: + assumes A: "A \ sets (M1 \\<^sub>M M2)" and finite: "emeasure (M1 \\<^sub>M M2) A < \" + shows "AE x in M1. emeasure M2 {y\space M2. (x, y) \ A} < \" +proof - + from M2.emeasure_pair_measure_alt[OF A] finite + have "(\\<^sup>+ x. emeasure M2 (Pair x -` A) \M1) \ \" + by simp + then have "AE x in M1. emeasure M2 (Pair x -` A) \ \" + by (rule positive_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A) + moreover have "\x. x \ space M1 \ Pair x -` A = {y\space M2. (x, y) \ A}" + using sets.sets_into_space[OF A] by (auto simp: space_pair_measure) + ultimately show ?thesis by auto +qed + +lemma (in pair_sigma_finite) AE_integrable_fst': + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" + shows "AE x in M1. integrable M2 (\y. f (x, y))" +proof - + have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" + by (rule M2.positive_integral_fst) simp + also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) \ \" + using f unfolding integrable_iff_bounded by simp + finally have "AE x in M1. (\\<^sup>+y. norm (f (x, y)) \M2) \ \" + by (intro positive_integral_PInf_AE M2.borel_measurable_positive_integral ) + (auto simp: measurable_split_conv) + with AE_space show ?thesis + by eventually_elim + (auto simp: integrable_iff_bounded measurable_compose[OF _ borel_measurable_integrable[OF f]]) +qed + +lemma (in pair_sigma_finite) integrable_fst': + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes f[measurable]: "integrable (M1 \\<^sub>M M2) f" + shows "integrable M1 (\x. \y. f (x, y) \M2)" + unfolding integrable_iff_bounded +proof + show "(\x. \ y. f (x, y) \M2) \ borel_measurable M1" + by (rule M2.borel_measurable_lebesgue_integral) simp + have "(\\<^sup>+ x. ereal (norm (\ y. f (x, y) \M2)) \M1) \ (\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1)" + using AE_integrable_fst'[OF f] by (auto intro!: positive_integral_mono_AE integral_norm_bound_ereal) + also have "(\\<^sup>+x. (\\<^sup>+y. norm (f (x, y)) \M2) \M1) = (\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2))" + by (rule M2.positive_integral_fst) simp + also have "(\\<^sup>+x. norm (f x) \(M1 \\<^sub>M M2)) < \" + using f unfolding integrable_iff_bounded by simp + finally show "(\\<^sup>+ x. ereal (norm (\ y. f (x, y) \M2)) \M1) < \" . +qed + +lemma (in pair_sigma_finite) integral_fst': + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes f: "integrable (M1 \\<^sub>M M2) f" + shows "(\x. (\y. f (x, y) \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) f" +using f proof induct + case (base A c) + have A[measurable]: "A \ sets (M1 \\<^sub>M M2)" by fact + + have eq: "\x y. x \ space M1 \ indicator A (x, y) = indicator {y\space M2. (x, y) \ A} y" + using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure) + + have int_A: "integrable (M1 \\<^sub>M M2) (indicator A :: _ \ real)" + using base by (rule integrable_real_indicator) + + have "(\ x. \ y. indicator A (x, y) *\<^sub>R c \M2 \M1) = (\x. measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c \M1)" + proof (intro integral_cong_AE, simp, simp) + from AE_integrable_fst'[OF int_A] AE_space + show "AE x in M1. (\y. indicator A (x, y) *\<^sub>R c \M2) = measure M2 {y\space M2. (x, y) \ A} *\<^sub>R c" + by eventually_elim (simp add: eq integrable_indicator_iff) + qed + also have "\ = measure (M1 \\<^sub>M M2) A *\<^sub>R c" + proof (subst integral_scaleR_left) + have "(\\<^sup>+x. ereal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = + (\\<^sup>+x. emeasure M2 {y \ space M2. (x, y) \ A} \M1)" + using emeasure_pair_measure_finite[OF base] + by (intro positive_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ereal_measure) + also have "\ = emeasure (M1 \\<^sub>M M2) A" + using sets.sets_into_space[OF A] + by (subst M2.emeasure_pair_measure_alt) + (auto intro!: positive_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure) + finally have *: "(\\<^sup>+x. ereal (measure M2 {y \ space M2. (x, y) \ A}) \M1) = emeasure (M1 \\<^sub>M M2) A" . + + from base * show "integrable M1 (\x. measure M2 {y \ space M2. (x, y) \ A})" + by (simp add: measure_nonneg integrable_iff_bounded) + then have "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) = + (\\<^sup>+x. ereal (measure M2 {y \ space M2. (x, y) \ A}) \M1)" + by (rule positive_integral_eq_integral[symmetric]) (simp add: measure_nonneg) + also note * + finally show "(\x. measure M2 {y \ space M2. (x, y) \ A} \M1) *\<^sub>R c = measure (M1 \\<^sub>M M2) A *\<^sub>R c" + using base by (simp add: emeasure_eq_ereal_measure) + qed + also have "\ = (\ a. indicator A a *\<^sub>R c \(M1 \\<^sub>M M2))" + using base by simp + finally show ?case . +next + case (add f g) + then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "g \ borel_measurable (M1 \\<^sub>M M2)" + by auto + have "(\ x. \ y. f (x, y) + g (x, y) \M2 \M1) = + (\ x. (\ y. f (x, y) \M2) + (\ y. g (x, y) \M2) \M1)" + apply (rule integral_cong_AE) + apply simp_all + using AE_integrable_fst'[OF add(1)] AE_integrable_fst'[OF add(3)] + apply eventually_elim + apply simp + done + also have "\ = (\ x. f x \(M1 \\<^sub>M M2)) + (\ x. g x \(M1 \\<^sub>M M2))" + using integrable_fst'[OF add(1)] integrable_fst'[OF add(3)] add(2,4) by simp + finally show ?case + using add by simp +next + case (lim f s) + then have [measurable]: "f \ borel_measurable (M1 \\<^sub>M M2)" "\i. s i \ borel_measurable (M1 \\<^sub>M M2)" + by auto + + show ?case + proof (rule LIMSEQ_unique) + show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) ----> integral\<^sup>L (M1 \\<^sub>M M2) f" + proof (rule integral_dominated_convergence) + show "integrable (M1 \\<^sub>M M2) (\x. 2 * norm (f x))" + using lim(5) by (auto intro: integrable_norm) + qed (insert lim, auto) + have "(\i. \ x. \ y. s i (x, y) \M2 \M1) ----> \ x. \ y. f (x, y) \M2 \M1" + proof (rule integral_dominated_convergence) + have "AE x in M1. \i. integrable M2 (\y. s i (x, y))" + unfolding AE_all_countable using AE_integrable_fst'[OF lim(1)] .. + with AE_space AE_integrable_fst'[OF lim(5)] + show "AE x in M1. (\i. \ y. s i (x, y) \M2) ----> \ y. f (x, y) \M2" + proof eventually_elim + fix x assume x: "x \ space M1" and + s: "\i. integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" + show "(\i. \ y. s i (x, y) \M2) ----> \ y. f (x, y) \M2" + proof (rule integral_dominated_convergence) + show "integrable M2 (\y. 2 * norm (f (x, y)))" + using f by (auto intro: integrable_norm) + show "AE xa in M2. (\i. s i (x, xa)) ----> f (x, xa)" + using x lim(3) by (auto simp: space_pair_measure) + show "\i. AE xa in M2. norm (s i (x, xa)) \ 2 * norm (f (x, xa))" + using x lim(4) by (auto simp: space_pair_measure) + qed (insert x, measurable) + qed + show "integrable M1 (\x. (\ y. 2 * norm (f (x, y)) \M2))" + by (intro integrable_mult_right integrable_norm integrable_fst' lim) + fix i show "AE x in M1. norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" + using AE_space AE_integrable_fst'[OF lim(1), of i] AE_integrable_fst'[OF lim(5)] + proof eventually_elim + fix x assume x: "x \ space M1" + and s: "integrable M2 (\y. s i (x, y))" and f: "integrable M2 (\y. f (x, y))" + from s have "norm (\ y. s i (x, y) \M2) \ (\\<^sup>+y. norm (s i (x, y)) \M2)" + by (rule integral_norm_bound_ereal) + also have "\ \ (\\<^sup>+y. 2 * norm (f (x, y)) \M2)" + using x lim by (auto intro!: positive_integral_mono simp: space_pair_measure) + also have "\ = (\y. 2 * norm (f (x, y)) \M2)" + using f by (intro positive_integral_eq_integral) auto + finally show "norm (\ y. s i (x, y) \M2) \ (\ y. 2 * norm (f (x, y)) \M2)" + by simp + qed + qed simp_all + then show "(\i. integral\<^sup>L (M1 \\<^sub>M M2) (s i)) ----> \ x. \ y. f (x, y) \M2 \M1" + using lim by simp + qed +qed + +lemma (in pair_sigma_finite) + fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" + assumes f: "integrable (M1 \\<^sub>M M2) (split f)" + shows AE_integrable_fst: "AE x in M1. integrable M2 (\y. f x y)" (is "?AE") + and integrable_fst: "integrable M1 (\x. \y. f x y \M2)" (is "?INT") + and integral_fst: "(\x. (\y. f x y \M2) \M1) = integral\<^sup>L (M1 \\<^sub>M M2) (\(x, y). f x y)" (is "?EQ") + using AE_integrable_fst'[OF f] integrable_fst'[OF f] integral_fst'[OF f] by auto + +lemma (in pair_sigma_finite) + fixes f :: "_ \ _ \ _::{banach, second_countable_topology}" + assumes f[measurable]: "integrable (M1 \\<^sub>M M2) (split f)" + shows AE_integrable_snd: "AE y in M2. integrable M1 (\x. f x y)" (is "?AE") + and integrable_snd: "integrable M2 (\y. \x. f x y \M1)" (is "?INT") + and integral_snd: "(\y. (\x. f x y \M1) \M2) = integral\<^sup>L (M1 \\<^sub>M M2) (split f)" (is "?EQ") +proof - + interpret Q: pair_sigma_finite M2 M1 by default + have Q_int: "integrable (M2 \\<^sub>M M1) (\(x, y). f y x)" + using f unfolding integrable_product_swap_iff[symmetric] by simp + show ?AE using Q.AE_integrable_fst'[OF Q_int] by simp + show ?INT using Q.integrable_fst'[OF Q_int] by simp + show ?EQ using Q.integral_fst'[OF Q_int] + using integral_product_swap[of "split f"] by simp +qed + +lemma (in pair_sigma_finite) Fubini_integral: + fixes f :: "_ \ _ \ _ :: {banach, second_countable_topology}" + assumes f: "integrable (M1 \\<^sub>M M2) (split f)" + shows "(\y. (\x. f x y \M1) \M2) = (\x. (\y. f x y \M2) \M1)" + unfolding integral_snd[OF assms] integral_fst[OF assms] .. + +lemma (in product_sigma_finite) product_integral_singleton: + fixes f :: "_ \ _::{banach, second_countable_topology}" + shows "f \ borel_measurable (M i) \ (\x. f (x i) \Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" + apply (subst distr_singleton[symmetric]) + apply (subst integral_distr) + apply simp_all + done + +lemma (in product_sigma_finite) product_integral_fold: + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" + and f: "integrable (Pi\<^sub>M (I \ J) M) f" + shows "integral\<^sup>L (Pi\<^sub>M (I \ J) M) f = (\x. (\y. f (merge I J (x, y)) \Pi\<^sub>M J M) \Pi\<^sub>M I M)" +proof - + interpret I: finite_product_sigma_finite M I by default fact + interpret J: finite_product_sigma_finite M J by default fact + have "finite (I \ J)" using fin by auto + interpret IJ: finite_product_sigma_finite M "I \ J" by default fact + interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default + let ?M = "merge I J" + let ?f = "\x. f (?M x)" + from f have f_borel: "f \ borel_measurable (Pi\<^sub>M (I \ J) M)" + by auto + have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" + using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) + have f_int: "integrable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) ?f" + by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) + show ?thesis + apply (subst distr_merge[symmetric, OF IJ fin]) + apply (subst integral_distr[OF measurable_merge f_borel]) + apply (subst P.integral_fst'[symmetric, OF f_int]) + apply simp + done +qed + +lemma (in product_sigma_finite) product_integral_insert: + fixes f :: "_ \ _::{banach, second_countable_topology}" + assumes I: "finite I" "i \ I" + and f: "integrable (Pi\<^sub>M (insert i I) M) f" + shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^sub>M I M)" +proof - + have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \ {i}) M) f" + by simp + also have "\ = (\x. (\y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) \Pi\<^sub>M I M)" + using f I by (intro product_integral_fold) auto + also have "\ = (\x. (\y. f (x(i := y)) \M i) \Pi\<^sub>M I M)" + proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric]) + fix x assume x: "x \ space (Pi\<^sub>M I M)" + have f_borel: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" + using f by auto + show "(\y. f (x(i := y))) \ borel_measurable (M i)" + using measurable_comp[OF measurable_component_update f_borel, OF x `i \ I`] + unfolding comp_def . + from x I show "(\ y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^sub>M {i} M)" + by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) + qed + finally show ?thesis . +qed + +lemma (in product_sigma_finite) product_integrable_setprod: + fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" + assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" + shows "integrable (Pi\<^sub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") +proof (unfold integrable_iff_bounded, intro conjI) + interpret finite_product_sigma_finite M I by default fact + show "?f \ borel_measurable (Pi\<^sub>M I M)" + using assms by simp + have "(\\<^sup>+ x. ereal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) = + (\\<^sup>+ x. (\i\I. ereal (norm (f i (x i)))) \Pi\<^sub>M I M)" + by (simp add: setprod_norm setprod_ereal) + also have "\ = (\i\I. \\<^sup>+ x. ereal (norm (f i x)) \M i)" + using assms by (intro product_positive_integral_setprod) auto + also have "\ < \" + using integrable by (simp add: setprod_PInf positive_integral_positive integrable_iff_bounded) + finally show "(\\<^sup>+ x. ereal (norm (\i\I. f i (x i))) \Pi\<^sub>M I M) < \" . +qed + +lemma (in product_sigma_finite) product_integral_setprod: + fixes f :: "'i \ 'a \ _::{real_normed_field,banach,second_countable_topology}" + assumes "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" + shows "(\x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>L (M i) (f i))" +using assms proof induct + case empty + interpret finite_measure "Pi\<^sub>M {} M" + by rule (simp add: space_PiM) + show ?case by (simp add: space_PiM measure_def) +next + case (insert i I) + then have iI: "finite (insert i I)" by auto + then have prod: "\J. J \ insert i I \ + integrable (Pi\<^sub>M J M) (\x. (\i\J. f i (x i)))" + by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset) + interpret I: finite_product_sigma_finite M I by default fact + have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" + using `i \ I` by (auto intro!: setprod_cong) + show ?case + unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] + by (simp add: * insert prod subset_insertI) +qed + +lemma integrable_subalgebra: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes borel: "f \ borel_measurable N" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" + shows "integrable N f \ integrable M f" (is ?P) +proof - + have "f \ borel_measurable M" + using assms by (auto simp: measurable_def) + with assms show ?thesis + using assms by (auto simp: integrable_iff_bounded positive_integral_subalgebra) +qed + +lemma integral_subalgebra: + fixes f :: "'a \ 'b::{banach, second_countable_topology}" + assumes borel: "f \ borel_measurable N" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" + shows "integral\<^sup>L N f = integral\<^sup>L M f" +proof cases + assume "integrable N f" + then show ?thesis + proof induct + case base with assms show ?case by (auto simp: subset_eq measure_def) + next + case (add f g) + then have "(\ a. f a + g a \N) = integral\<^sup>L M f + integral\<^sup>L M g" + by simp + also have "\ = (\ a. f a + g a \M)" + using add integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of g] by simp + finally show ?case . + next + case (lim f s) + then have M: "\i. integrable M (s i)" "integrable M f" + using integrable_subalgebra[OF _ N, of f] integrable_subalgebra[OF _ N, of "s i" for i] by simp_all + show ?case + proof (intro LIMSEQ_unique) + show "(\i. integral\<^sup>L N (s i)) ----> integral\<^sup>L N f" + apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) + using lim + apply auto + done + show "(\i. integral\<^sup>L N (s i)) ----> integral\<^sup>L M f" + unfolding lim + apply (rule integral_dominated_convergence[where w="\x. 2 * norm (f x)"]) + using lim M N(2) + apply auto + done + qed + qed +qed (simp add: not_integrable_integral_eq integrable_subalgebra[OF assms]) + +hide_const simple_bochner_integral +hide_const simple_bochner_integrable + +end diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Borel_Space.thy Mon May 19 12:04:45 2014 +0200 @@ -655,6 +655,9 @@ subsection "Borel measurable operators" +lemma borel_measurable_norm[measurable]: "norm \ borel_measurable borel" + by (intro borel_measurable_continuous_on1 continuous_intros) + lemma borel_measurable_uminus[measurable (raw)]: fixes g :: "'a \ 'b::{second_countable_topology, real_normed_vector}" assumes g: "g \ borel_measurable M" @@ -820,20 +823,6 @@ lemma borel_measurable_exp[measurable]: "exp \ borel_measurable borel" by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp) -lemma measurable_count_space_eq2_countable: - fixes f :: "'a => 'c::countable" - shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" -proof - - { fix X assume "X \ A" "f \ space M \ A" - then have "f -` X \ space M = (\a\X. f -` {a} \ space M)" - by auto - moreover assume "\a. a\A \ f -` {a} \ space M \ sets M" - ultimately have "f -` X \ space M \ sets M" - using `X \ A` by (simp add: subset_eq del: UN_simps) } - then show ?thesis - unfolding measurable_def by auto -qed - lemma measurable_real_floor[measurable]: "(floor :: real \ int) \ measurable borel (count_space UNIV)" proof - @@ -1166,6 +1155,37 @@ ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff) qed +lemma borel_measurable_LIMSEQ_metric: + fixes f :: "nat \ 'a \ 'b :: metric_space" + assumes [measurable]: "\i. f i \ borel_measurable M" + assumes lim: "\x. x \ space M \ (\i. f i x) ----> g x" + shows "g \ borel_measurable M" + unfolding borel_eq_closed +proof (safe intro!: measurable_measure_of) + fix A :: "'b set" assume "closed A" + + have [measurable]: "(\x. infdist (g x) A) \ borel_measurable M" + proof (rule borel_measurable_LIMSEQ) + show "\x. x \ space M \ (\i. infdist (f i x) A) ----> infdist (g x) A" + by (intro tendsto_infdist lim) + show "\i. (\x. infdist (f i x) A) \ borel_measurable M" + by (intro borel_measurable_continuous_on[where f="\x. infdist x A"] + continuous_at_imp_continuous_on ballI continuous_infdist isCont_ident) auto + qed + + show "g -` A \ space M \ sets M" + proof cases + assume "A \ {}" + then have "\x. infdist x A = 0 \ x \ A" + using `closed A` by (simp add: in_closed_iff_infdist_zero) + then have "g -` A \ space M = {x\space M. infdist (g x) A = 0}" + by auto + also have "\ \ sets M" + by measurable + finally show ?thesis . + qed simp +qed auto + lemma sets_Collect_Cauchy[measurable]: fixes f :: "nat \ 'a => real" assumes f[measurable]: "\i. f i \ borel_measurable M" diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Complete_Measure.thy Mon May 19 12:04:45 2014 +0200 @@ -3,7 +3,7 @@ *) theory Complete_Measure -imports Lebesgue_Integration +imports Bochner_Integration begin definition diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Distributions.thy --- a/src/HOL/Probability/Distributions.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Distributions.thy Mon May 19 12:04:45 2014 +0200 @@ -162,8 +162,8 @@ by (auto intro!: exponential_distributedI simp: one_ereal_def emeasure_eq_measure) lemma borel_integral_x_exp: - "(\x. x * exp (- x) * indicator {0::real ..} x \lborel) = 1" -proof (rule integral_monotone_convergence) + "has_bochner_integral lborel (\x. x * exp (- x) * indicator {0::real ..} x) 1" +proof (rule has_bochner_integral_monotone_convergence) let ?f = "\i x. x * exp (- x) * indicator {0::real .. i} x" have "eventually (\b::real. 0 \ b) at_top" by (rule eventually_ge_at_top) @@ -172,7 +172,7 @@ fix b :: real assume [simp]: "0 \ b" have "(\x. (exp (-x)) * indicator {0 .. b} x \lborel) - (integral\<^sup>L lborel (?f b)) = (\x. (exp (-x) - x * exp (-x)) * indicator {0 .. b} x \lborel)" - by (subst integral_diff(2)[symmetric]) + by (subst integral_diff[symmetric]) (auto intro!: borel_integrable_atLeastAtMost integral_cong split: split_indicator) also have "\ = b * exp (-b) - 0 * exp (- 0)" proof (rule integral_FTC_atLeastAtMost) @@ -217,9 +217,11 @@ using `0 < l` by (auto split: split_indicator simp: zero_le_mult_iff exponential_density_def) from borel_integral_x_exp `0 < l` - show "(\ x. exponential_density l x * x \lborel) = 1 / l" - by (subst (asm) lebesgue_integral_real_affine[of "l" _ 0]) - (simp_all add: borel_measurable_exp nonzero_eq_divide_eq ac_simps) + have "has_bochner_integral lborel (\x. exponential_density l x * x) (1 / l)" + by (subst (asm) lborel_has_bochner_integral_real_affine_iff[of l _ _ 0]) + (simp_all add: field_simps) + then show "(\ x. exponential_density l x * x \lborel) = 1 / l" + by (metis has_bochner_integral_integral_eq) qed simp subsection {* Uniform distribution *} @@ -356,7 +358,9 @@ uniform_distributed_bounds[of X a b] uniform_distributed_measure[of X a b] distributed_measurable[of M lborel X] - by (auto intro!: uniform_distrI_borel_atLeastAtMost simp: one_ereal_def emeasure_eq_measure) + by (auto intro!: uniform_distrI_borel_atLeastAtMost + simp: one_ereal_def emeasure_eq_measure + simp del: measure_lborel) lemma (in prob_space) uniform_distributed_expectation: fixes a b :: real diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Mon May 19 12:04:45 2014 +0200 @@ -782,7 +782,7 @@ show ?thesis apply (subst distr_merge[OF IJ, symmetric]) apply (subst positive_integral_distr[OF measurable_merge f]) - apply (subst J.positive_integral_fst_measurable(2)[symmetric, OF P_borel]) + apply (subst J.positive_integral_fst[symmetric, OF P_borel]) apply simp done qed @@ -859,17 +859,6 @@ done qed (simp add: space_PiM) -lemma (in product_sigma_finite) product_integral_singleton: - assumes f: "f \ borel_measurable (M i)" - shows "(\x. f (x i) \Pi\<^sub>M {i} M) = integral\<^sup>L (M i) f" -proof - - interpret I: finite_product_sigma_finite M "{i}" by default simp - have *: "(\x. ereal (f x)) \ borel_measurable (M i)" - "(\x. ereal (- f x)) \ borel_measurable (M i)" - using assms by auto - show ?thesis - unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. -qed lemma (in product_sigma_finite) distr_component: "distr (M i) (Pi\<^sub>M {i} M) (\x. \i\{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P") proof (intro measure_eqI[symmetric]) @@ -888,32 +877,6 @@ finally show "emeasure (Pi\<^sub>M {i} M) A = emeasure ?D A" . qed simp -lemma (in product_sigma_finite) product_integral_fold: - assumes IJ[simp]: "I \ J = {}" and fin: "finite I" "finite J" - and f: "integrable (Pi\<^sub>M (I \ J) M) f" - shows "integral\<^sup>L (Pi\<^sub>M (I \ J) M) f = (\x. (\y. f (merge I J (x, y)) \Pi\<^sub>M J M) \Pi\<^sub>M I M)" -proof - - interpret I: finite_product_sigma_finite M I by default fact - interpret J: finite_product_sigma_finite M J by default fact - have "finite (I \ J)" using fin by auto - interpret IJ: finite_product_sigma_finite M "I \ J" by default fact - interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by default - let ?M = "merge I J" - let ?f = "\x. f (?M x)" - from f have f_borel: "f \ borel_measurable (Pi\<^sub>M (I \ J) M)" - by auto - have P_borel: "(\x. f (merge I J x)) \ borel_measurable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M)" - using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def) - have f_int: "integrable (Pi\<^sub>M I M \\<^sub>M Pi\<^sub>M J M) ?f" - by (rule integrable_distr[OF measurable_merge]) (simp add: distr_merge[OF IJ fin] f) - show ?thesis - apply (subst distr_merge[symmetric, OF IJ fin]) - apply (subst integral_distr[OF measurable_merge f_borel]) - apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int]) - apply simp - done -qed - lemma (in product_sigma_finite) assumes IJ: "I \ J = {}" "finite I" "finite J" and A: "A \ sets (Pi\<^sub>M (I \ J) M)" shows emeasure_fold_integral: @@ -939,82 +902,6 @@ by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong) qed -lemma (in product_sigma_finite) product_integral_insert: - assumes I: "finite I" "i \ I" - and f: "integrable (Pi\<^sub>M (insert i I) M) f" - shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\x. (\y. f (x(i:=y)) \M i) \Pi\<^sub>M I M)" -proof - - have "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = integral\<^sup>L (Pi\<^sub>M (I \ {i}) M) f" - by simp - also have "\ = (\x. (\y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) \Pi\<^sub>M I M)" - using f I by (intro product_integral_fold) auto - also have "\ = (\x. (\y. f (x(i := y)) \M i) \Pi\<^sub>M I M)" - proof (rule integral_cong, subst product_integral_singleton[symmetric]) - fix x assume x: "x \ space (Pi\<^sub>M I M)" - have f_borel: "f \ borel_measurable (Pi\<^sub>M (insert i I) M)" - using f by auto - show "(\y. f (x(i := y))) \ borel_measurable (M i)" - using measurable_comp[OF measurable_component_update f_borel, OF x `i \ I`] - unfolding comp_def . - from x I show "(\ y. f (merge I {i} (x,y)) \Pi\<^sub>M {i} M) = (\ xa. f (x(i := xa i)) \Pi\<^sub>M {i} M)" - by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def PiE_def) - qed - finally show ?thesis . -qed - -lemma (in product_sigma_finite) product_integrable_setprod: - fixes f :: "'i \ 'a \ real" - assumes [simp]: "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" - shows "integrable (Pi\<^sub>M I M) (\x. (\i\I. f i (x i)))" (is "integrable _ ?f") -proof - - interpret finite_product_sigma_finite M I by default fact - have f: "\i. i \ I \ f i \ borel_measurable (M i)" - using integrable unfolding integrable_def by auto - have borel: "?f \ borel_measurable (Pi\<^sub>M I M)" - using measurable_comp[OF measurable_component_singleton[of _ I M] f] by (auto simp: comp_def) - moreover have "integrable (Pi\<^sub>M I M) (\x. \\i\I. f i (x i)\)" - proof (unfold integrable_def, intro conjI) - show "(\x. abs (?f x)) \ borel_measurable (Pi\<^sub>M I M)" - using borel by auto - have "(\\<^sup>+x. ereal (abs (?f x)) \Pi\<^sub>M I M) = (\\<^sup>+x. (\i\I. ereal (abs (f i (x i)))) \Pi\<^sub>M I M)" - by (simp add: setprod_ereal abs_setprod) - also have "\ = (\i\I. (\\<^sup>+x. ereal (abs (f i x)) \M i))" - using f by (subst product_positive_integral_setprod) auto - also have "\ < \" - using integrable[THEN integrable_abs] - by (simp add: setprod_PInf integrable_def positive_integral_positive) - finally show "(\\<^sup>+x. ereal (abs (?f x)) \(Pi\<^sub>M I M)) \ \" by auto - have "(\\<^sup>+x. ereal (- abs (?f x)) \(Pi\<^sub>M I M)) = (\\<^sup>+x. 0 \(Pi\<^sub>M I M))" - by (intro positive_integral_cong_pos) auto - then show "(\\<^sup>+x. ereal (- abs (?f x)) \(Pi\<^sub>M I M)) \ \" by simp - qed - ultimately show ?thesis - by (rule integrable_abs_iff[THEN iffD1]) -qed - -lemma (in product_sigma_finite) product_integral_setprod: - fixes f :: "'i \ 'a \ real" - assumes "finite I" and integrable: "\i. i \ I \ integrable (M i) (f i)" - shows "(\x. (\i\I. f i (x i)) \Pi\<^sub>M I M) = (\i\I. integral\<^sup>L (M i) (f i))" -using assms proof induct - case empty - interpret finite_measure "Pi\<^sub>M {} M" - by rule (simp add: space_PiM) - show ?case by (simp add: space_PiM measure_def) -next - case (insert i I) - then have iI: "finite (insert i I)" by auto - then have prod: "\J. J \ insert i I \ - integrable (Pi\<^sub>M J M) (\x. (\i\J. f i (x i)))" - by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset) - interpret I: finite_product_sigma_finite M I by default fact - have *: "\x y. (\j\I. f j (if j = i then y else x j)) = (\j\I. f j (x j))" - using `i \ I` by (auto intro!: setprod_cong) - show ?case - unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]] - by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI) -qed - lemma sets_Collect_single: "i \ I \ A \ sets (M i) \ { x \ space (Pi\<^sub>M I M). x i \ A } \ sets (Pi\<^sub>M I M)" by simp diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Information.thy Mon May 19 12:04:45 2014 +0200 @@ -79,37 +79,26 @@ definition "KL_divergence b M N = integral\<^sup>L N (entropy_density b M N)" -lemma (in information_space) measurable_entropy_density: - assumes ac: "absolutely_continuous M N" "sets N = events" - shows "entropy_density b M N \ borel_measurable M" -proof - - from borel_measurable_RN_deriv[OF ac] b_gt_1 show ?thesis - unfolding entropy_density_def by auto -qed - -lemma borel_measurable_RN_deriv_density[measurable (raw)]: - "f \ borel_measurable M \ RN_deriv M (density M f) \ borel_measurable M" - using borel_measurable_RN_deriv_density[of "\x. max 0 (f x )" M] - by (simp add: density_max_0[symmetric]) +lemma measurable_entropy_density[measurable]: "entropy_density b M N \ borel_measurable M" + unfolding entropy_density_def by auto lemma (in sigma_finite_measure) KL_density: fixes f :: "'a \ real" assumes "1 < b" - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes f[measurable]: "f \ borel_measurable M" and nn: "AE x in M. 0 \ f x" shows "KL_divergence b M (density M f) = (\x. f x * log b (f x) \M)" unfolding KL_divergence_def -proof (subst integral_density) - show "entropy_density b M (density M (\x. ereal (f x))) \ borel_measurable M" +proof (subst integral_real_density) + show [measurable]: "entropy_density b M (density M (\x. ereal (f x))) \ borel_measurable M" using f by (auto simp: comp_def entropy_density_def) have "density M (RN_deriv M (density M f)) = density M f" - using f by (intro density_RN_deriv_density) auto + using f nn by (intro density_RN_deriv_density) auto then have eq: "AE x in M. RN_deriv M (density M f) x = f x" - using f - by (intro density_unique) - (auto intro!: borel_measurable_log borel_measurable_RN_deriv_density simp: RN_deriv_density_nonneg) + using f nn by (intro density_unique) (auto simp: RN_deriv_nonneg) show "(\x. f x * entropy_density b M (density M (\x. ereal (f x))) x \M) = (\x. f x * log b (f x) \M)" apply (intro integral_cong_AE) + apply measurable using eq apply eventually_elim apply (auto simp: entropy_density_def) @@ -161,8 +150,10 @@ then have D_pos: "(\\<^sup>+ x. ereal (D x) \M) = 1" using N.emeasure_space_1 by simp - have "integrable M D" "integral\<^sup>L M D = 1" - using D D_pos D_neg unfolding integrable_def lebesgue_integral_def by simp_all + have "integrable M D" + using D D_pos D_neg unfolding real_integrable_def real_lebesgue_integral_def by simp_all + then have "integral\<^sup>L M D = 1" + using D D_pos D_neg by (simp add: real_lebesgue_integral_def) have "0 \ 1 - measure M ?D_set" using prob_le_1 by (auto simp: field_simps) @@ -172,10 +163,9 @@ also have "\ < (\ x. D x * (ln b * log b (D x)) \M)" proof (rule integral_less_AE) show "integrable M (\x. D x - indicator ?D_set x)" - using `integrable M D` - by (intro integral_diff integral_indicator) auto + using `integrable M D` by auto next - from integral_cmult(1)[OF int, of "ln b"] + from integrable_mult_left(1)[OF int, of "ln b"] show "integrable M (\x. D x * (ln b * log b (D x)))" by (simp add: ac_simps) next @@ -242,7 +232,7 @@ also have "\ = (\ x. ln b * (D x * log b (D x)) \M)" by (simp add: ac_simps) also have "\ = ln b * (\ x. D x * log b (D x) \M)" - using int by (rule integral_cmult) + using int by simp finally show ?thesis using b_gt_1 D by (subst KL_density) (auto simp: zero_less_mult_iff) qed @@ -260,7 +250,7 @@ qed then have "AE x in M. log b (real (RN_deriv M M x)) = 0" by (elim AE_mp) simp - from integral_cong_AE[OF this] + from integral_cong_AE[OF _ _ this] have "integral\<^sup>L M (entropy_density b M M) = 0" by (simp add: entropy_density_def comp_def) then show "KL_divergence b M M = 0" @@ -291,7 +281,7 @@ have "N = density M (RN_deriv M N)" using ac by (rule density_RN_deriv[symmetric]) also have "\ = density M D" - using borel_measurable_RN_deriv[OF ac] D by (auto intro!: density_cong) + using D by (auto intro!: density_cong) finally have N: "N = density M D" . from absolutely_continuous_AE[OF ac(2,1) D(2)] D b_gt_1 ac measurable_entropy_density @@ -299,7 +289,7 @@ by (intro integrable_cong_AE[THEN iffD2, OF _ _ _ int]) (auto simp: N entropy_density_def) with D b_gt_1 have "integrable M (\x. D x * log b (D x))" - by (subst integral_density(2)[symmetric]) (auto simp: N[symmetric] comp_def) + by (subst integrable_real_density[symmetric]) (auto simp: N[symmetric] comp_def) with `prob_space N` D show ?thesis unfolding N by (intro KL_eq_0_iff_eq) auto @@ -335,7 +325,7 @@ using f g by (auto simp: AE_density) show "integrable (density M f) (\x. g x / f x * log b (g x / f x))" using `1 < b` f g ac - by (subst integral_density) + by (subst integrable_density) (auto intro!: integrable_cong_AE[THEN iffD2, OF _ _ _ int] measurable_If) qed also have "\ = KL_divergence b (density M f) (density M g)" @@ -414,7 +404,7 @@ "distributed M N X f \ g \ borel_measurable N \ integrable N (\x. f x * g x) \ integrable M (\x. g (X x))" by (auto simp: distributed_real_AE - distributed_distr_eq_density[symmetric] integral_density[symmetric] integrable_distr_eq) + distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq) lemma distributed_transform_integrable: assumes Px: "distributed M N X Px" @@ -431,8 +421,9 @@ finally show ?thesis . qed -lemma integrable_cong_AE_imp: "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" - using integrable_cong_AE by blast +lemma integrable_cong_AE_imp: + "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" + using integrable_cong_AE[of f M g] by (auto simp: eq_commute) lemma (in information_space) finite_entropy_integrable: "finite_entropy S X Px \ integrable S (\x. Px x * log b (Px x))" @@ -485,7 +476,7 @@ show ac: "absolutely_continuous P Q" by (simp add: absolutely_continuous_def) then have ed: "entropy_density b P Q \ borel_measurable P" - by (rule P.measurable_entropy_density) simp + by simp have "AE x in P. 1 = RN_deriv P Q x" proof (rule P.RN_deriv_unique) @@ -494,13 +485,15 @@ qed auto then have ae_0: "AE x in P. entropy_density b P Q x = 0" by eventually_elim (auto simp: entropy_density_def) - then have "integrable P (entropy_density b P Q) \ integrable Q (\x. 0)" + then have "integrable P (entropy_density b P Q) \ integrable Q (\x. 0::real)" using ed unfolding `Q = P` by (intro integrable_cong_AE) auto then show "integrable Q (entropy_density b P Q)" by simp - show "mutual_information b S T X Y = 0" + from ae_0 have "mutual_information b S T X Y = (\x. 0 \P)" unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] `Q = P` - using ae_0 by (simp cong: integral_cong_AE) } + by (intro integral_cong_AE) auto + then show "mutual_information b S T X Y = 0" + by simp } { assume ac: "absolutely_continuous P Q" assume int: "integrable Q (entropy_density b P Q)" @@ -722,8 +715,8 @@ lemma (in information_space) fixes Pxy :: "'b \ 'c \ real" and Px :: "'b \ real" and Py :: "'c \ real" assumes "sigma_finite_measure S" "sigma_finite_measure T" - assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py" - assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" + assumes Px[measurable]: "distributed M S X Px" and Py[measurable]: "distributed M T Y Py" + assumes Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y" shows mutual_information_eq_0: "mutual_information b S T X Y = 0" proof - @@ -743,7 +736,7 @@ ultimately have "AE x in S \\<^sub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0" by eventually_elim simp then have "(\x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \(S \\<^sub>M T)) = (\x. 0 \(S \\<^sub>M T))" - by (rule integral_cong_AE) + by (intro integral_cong_AE) auto then show ?thesis by (subst mutual_information_distr[OF assms(1-5)]) simp qed @@ -810,7 +803,7 @@ lemma (in information_space) fixes X :: "'a \ 'b" - assumes X: "distributed M MX X f" + assumes X[measurable]: "distributed M MX X f" shows entropy_distr: "entropy b MX X = - (\x. f x * log b (f x) \MX)" (is ?eq) proof - note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X] @@ -825,18 +818,15 @@ apply auto done - have int_eq: "- (\ x. log b (f x) \distr M MX X) = - (\ x. f x * log b (f x) \MX)" + have int_eq: "(\ x. f x * log b (f x) \MX) = (\ x. log b (f x) \distr M MX X)" unfolding distributed_distr_eq_density[OF X] using D by (subst integral_density) (auto simp: borel_measurable_ereal_iff) show ?eq - unfolding entropy_def KL_divergence_def entropy_density_def comp_def - apply (subst integral_cong_AE) - apply (rule ae_eq) - apply (rule int_eq) - done + unfolding entropy_def KL_divergence_def entropy_density_def comp_def int_eq neg_equal_iff_equal + using ae_eq by (intro integral_cong_AE) auto qed lemma (in prob_space) distributed_imp_emeasure_nonzero: @@ -861,7 +851,7 @@ lemma (in information_space) entropy_le: fixes Px :: "'b \ real" and MX :: "'b measure" - assumes X: "distributed M MX X Px" + assumes X[measurable]: "distributed M MX X Px" and fin: "emeasure MX {x \ space MX. Px x \ 0} \ \" and int: "integrable MX (\x. - Px x * log b (Px x))" shows "entropy b MX X \ log b (measure MX {x \ space MX. Px x \ 0})" @@ -873,7 +863,7 @@ have " - log b (measure MX {x \ space MX. Px x \ 0}) = - log b (\ x. indicator {x \ space MX. Px x \ 0} x \MX)" using Px fin - by (subst integral_indicator) (auto simp: measure_def borel_measurable_ereal_iff) + by (subst integral_real_indicator) (auto simp: measure_def borel_measurable_ereal_iff) also have "- log b (\ x. indicator {x \ space MX. Px x \ 0} x \MX) = - log b (\ x. 1 / Px x \distr M MX X)" unfolding distributed_distr_eq_density[OF X] using Px apply (intro arg_cong[where f="log b"] arg_cong[where f=uminus]) @@ -889,7 +879,7 @@ by (intro positive_integral_cong) (auto split: split_max) then show "integrable (distr M MX X) (\x. 1 / Px x)" unfolding distributed_distr_eq_density[OF X] using Px - by (auto simp: positive_integral_density integrable_def borel_measurable_ereal_iff fin positive_integral_max_0 + by (auto simp: positive_integral_density real_integrable_def borel_measurable_ereal_iff fin positive_integral_max_0 cong: positive_integral_cong) have "integrable MX (\x. Px x * log b (1 / Px x)) = integrable MX (\x. - Px x * log b (Px x))" @@ -900,7 +890,7 @@ then show "integrable (distr M MX X) (\x. - log b (1 / Px x))" unfolding distributed_distr_eq_density[OF X] using Px int - by (subst integral_density) (auto simp: borel_measurable_ereal_iff) + by (subst integrable_real_density) (auto simp: borel_measurable_ereal_iff) qed (auto simp: minus_log_convex[OF b_gt_1]) also have "\ = (\ x. log b (Px x) \distr M MX X)" unfolding distributed_distr_eq_density[OF X] using Px @@ -940,11 +930,11 @@ have eq: "(\ x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \MX) = (\ x. (- log b (measure MX A) / measure MX A) * indicator A x \MX)" using measure_nonneg[of MX A] uniform_distributed_params[OF X] - by (auto intro!: integral_cong split: split_indicator simp: log_divide_eq) + by (intro integral_cong) (auto split: split_indicator simp: log_divide_eq) show "- (\ x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \MX) = log b (measure MX A)" unfolding eq using uniform_distributed_params[OF X] - by (subst lebesgue_integral_cmult) (auto simp: measure_def) + by (subst integral_mult_right) (auto simp: measure_def) qed lemma (in information_space) entropy_simple_distributed: @@ -1068,7 +1058,7 @@ unfolding conditional_mutual_information_def apply (subst mi_eq) apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz]) - apply (subst integral_diff(2)[symmetric]) + apply (subst integral_diff[symmetric]) apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff) done @@ -1104,7 +1094,7 @@ apply auto done also have "\ = (\\<^sup>+(y, z). \\<^sup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \S \T \\<^sub>M P)" - by (subst STP.positive_integral_snd_measurable[symmetric]) (auto simp add: split_beta') + by (subst STP.positive_integral_snd[symmetric]) (auto simp add: split_beta') also have "\ = (\\<^sup>+x. ereal (Pyz x) * 1 \T \\<^sub>M P)" apply (rule positive_integral_cong_AE) using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space @@ -1123,7 +1113,7 @@ also have "\ < \" by simp finally have fin: "(\\<^sup>+ x. ?f x \?P) \ \" by simp - have pos: "(\\<^sup>+ x. ?f x \?P) \ 0" + have pos: "(\\<^sup>+x. ?f x \?P) \ 0" apply (subst positive_integral_density) apply simp apply (rule distributed_AE[OF Pxyz]) @@ -1131,7 +1121,7 @@ apply (simp add: split_beta') proof let ?g = "\x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))" - assume "(\\<^sup>+ x. ?g x \(S \\<^sub>M T \\<^sub>M P)) = 0" + assume "(\\<^sup>+x. ?g x \(S \\<^sub>M T \\<^sub>M P)) = 0" then have "AE x in S \\<^sub>M T \\<^sub>M P. ?g x \ 0" by (intro positive_integral_0_iff_AE[THEN iffD1]) auto then have "AE x in S \\<^sub>M T \\<^sub>M P. Pxyz x = 0" @@ -1154,19 +1144,32 @@ done have I3: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" - apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]]) + apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integrable_diff[OF I1 I2]]) using ae apply (auto simp: split_beta') done have "- log b 1 \ - log b (integral\<^sup>L ?P ?f)" proof (intro le_imp_neg_le log_le[OF b_gt_1]) - show "0 < integral\<^sup>L ?P ?f" - using neg pos fin positive_integral_positive[of ?P ?f] - by (cases "(\\<^sup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def less_le split_beta') - show "integral\<^sup>L ?P ?f \ 1" - using neg le1 fin positive_integral_positive[of ?P ?f] - by (cases "(\\<^sup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def) + have If: "integrable ?P ?f" + unfolding real_integrable_def + proof (intro conjI) + from neg show "(\\<^sup>+ x. - ?f x \?P) \ \" + by simp + from fin show "(\\<^sup>+ x. ?f x \?P) \ \" + by simp + qed simp + then have "(\\<^sup>+ x. ?f x \?P) = (\x. ?f x \?P)" + apply (rule positive_integral_eq_integral) + apply (subst AE_density) + apply simp + using ae5 ae6 ae7 ae8 + apply eventually_elim + apply auto + done + with positive_integral_positive[of ?P ?f] pos le1 + show "0 < (\x. ?f x \?P)" "(\x. ?f x \?P) \ 1" + by (simp_all add: one_ereal_def) qed also have "- log b (integral\<^sup>L ?P ?f) \ (\ x. - log b (?f x) \?P)" proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) @@ -1175,10 +1178,10 @@ using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] by eventually_elim (auto) show "integrable ?P ?f" - unfolding integrable_def + unfolding real_integrable_def using fin neg by (auto simp: split_beta') show "integrable ?P (\x. - log b (?f x))" - apply (subst integral_density) + apply (subst integrable_real_density) apply simp apply (auto intro!: distributed_real_AE[OF Pxyz]) [] apply simp @@ -1192,13 +1195,12 @@ qed (auto simp: b_gt_1 minus_log_convex) also have "\ = conditional_mutual_information b S T P X Y Z" unfolding `?eq` - apply (subst integral_density) + apply (subst integral_real_density) apply simp apply (auto intro!: distributed_real_AE[OF Pxyz]) [] apply simp apply (intro integral_cong_AE) using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] - apply eventually_elim apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps) done finally show ?nonneg @@ -1316,7 +1318,7 @@ unfolding conditional_mutual_information_def apply (subst mi_eq) apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz]) - apply (subst integral_diff(2)[symmetric]) + apply (subst integral_diff[symmetric]) apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff) done @@ -1349,8 +1351,7 @@ apply auto done also have "\ = (\\<^sup>+(y, z). \\<^sup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \S \T \\<^sub>M P)" - by (subst STP.positive_integral_snd_measurable[symmetric]) - (auto simp add: split_beta') + by (subst STP.positive_integral_snd[symmetric]) (auto simp add: split_beta') also have "\ = (\\<^sup>+x. ereal (Pyz x) * 1 \T \\<^sub>M P)" apply (rule positive_integral_cong_AE) using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space @@ -1399,19 +1400,32 @@ done have I3: "integrable (S \\<^sub>M T \\<^sub>M P) (\(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))" - apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]]) + apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integrable_diff[OF I1 I2]]) using ae apply (auto simp: split_beta') done have "- log b 1 \ - log b (integral\<^sup>L ?P ?f)" proof (intro le_imp_neg_le log_le[OF b_gt_1]) - show "0 < integral\<^sup>L ?P ?f" - using neg pos fin positive_integral_positive[of ?P ?f] - by (cases "(\\<^sup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def less_le split_beta') - show "integral\<^sup>L ?P ?f \ 1" - using neg le1 fin positive_integral_positive[of ?P ?f] - by (cases "(\\<^sup>+ x. ?f x \?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def) + have If: "integrable ?P ?f" + unfolding real_integrable_def + proof (intro conjI) + from neg show "(\\<^sup>+ x. - ?f x \?P) \ \" + by simp + from fin show "(\\<^sup>+ x. ?f x \?P) \ \" + by simp + qed simp + then have "(\\<^sup>+ x. ?f x \?P) = (\x. ?f x \?P)" + apply (rule positive_integral_eq_integral) + apply (subst AE_density) + apply simp + using ae5 ae6 ae7 ae8 + apply eventually_elim + apply auto + done + with positive_integral_positive[of ?P ?f] pos le1 + show "0 < (\x. ?f x \?P)" "(\x. ?f x \?P) \ 1" + by (simp_all add: one_ereal_def) qed also have "- log b (integral\<^sup>L ?P ?f) \ (\ x. - log b (?f x) \?P)" proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"]) @@ -1420,10 +1434,10 @@ using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] by eventually_elim (auto) show "integrable ?P ?f" - unfolding integrable_def + unfolding real_integrable_def using fin neg by (auto simp: split_beta') show "integrable ?P (\x. - log b (?f x))" - apply (subst integral_density) + apply (subst integrable_real_density) apply simp apply (auto intro!: distributed_real_AE[OF Pxyz]) [] apply simp @@ -1437,13 +1451,12 @@ qed (auto simp: b_gt_1 minus_log_convex) also have "\ = conditional_mutual_information b S T P X Y Z" unfolding `?eq` - apply (subst integral_density) + apply (subst integral_real_density) apply simp apply (auto intro!: distributed_real_AE[OF Pxyz]) [] apply simp apply (intro integral_cong_AE) using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE] - apply eventually_elim apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps) done finally show ?nonneg @@ -1532,7 +1545,7 @@ "\(X | Y) \ conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y" lemma (in information_space) conditional_entropy_generic_eq: - fixes Px :: "'b \ real" and Py :: "'c \ real" + fixes Pxy :: "_ \ real" and Py :: "'c \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" assumes Py[measurable]: "distributed M T Y Py" assumes Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" @@ -1552,19 +1565,15 @@ unfolding AE_density[OF distributed_borel_measurable, OF Pxy] unfolding distributed_distr_eq_density[OF Py] apply (rule ST.AE_pair_measure) - apply (auto intro!: sets.sets_Collect borel_measurable_eq measurable_compose[OF _ distributed_real_measurable[OF Py]] - distributed_real_measurable[OF Pxy] distributed_real_AE[OF Py] - borel_measurable_real_of_ereal measurable_compose[OF _ borel_measurable_RN_deriv_density]) + apply auto using distributed_RN_deriv[OF Py] apply auto done ultimately have "conditional_entropy b S T X Y = - (\x. Pxy x * log b (Pxy x / Py (snd x)) \(S \\<^sub>M T))" unfolding conditional_entropy_def neg_equal_iff_equal - apply (subst integral_density(1)[symmetric]) - apply (auto simp: distributed_real_measurable[OF Pxy] distributed_real_AE[OF Pxy] - measurable_compose[OF _ distributed_real_measurable[OF Py]] - distributed_distr_eq_density[OF Pxy] + apply (subst integral_real_density[symmetric]) + apply (auto simp: distributed_real_AE[OF Pxy] distributed_distr_eq_density[OF Pxy] intro!: integral_cong_AE) done then show ?thesis by (simp add: split_beta') @@ -1573,8 +1582,8 @@ lemma (in information_space) conditional_entropy_eq_entropy: fixes Px :: "'b \ real" and Py :: "'c \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" - assumes Py: "distributed M T Y Py" - assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" + assumes Py[measurable]: "distributed M T Y Py" + assumes Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" assumes I1: "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Pxy x))" assumes I2: "integrable (S \\<^sub>M T) (\x. Pxy x * log b (Py (snd x)))" shows "conditional_entropy b S T X Y = entropy b (S \\<^sub>M T) (\x. (X x, Y x)) - entropy b T Y" @@ -1606,7 +1615,6 @@ unfolding conditional_entropy_generic_eq[OF S T Py Pxy] neg_equal_iff_equal apply (intro integral_cong_AE) using ae - apply eventually_elim apply auto done also have "\ = - (\x. Pxy x * log b (Pxy x) \(S \\<^sub>M T)) - - (\x. Pxy x * log b (Py (snd x)) \(S \\<^sub>M T))" @@ -1702,8 +1710,8 @@ lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr: fixes Px :: "'b \ real" and Py :: "'c \ real" and Pxy :: "('b \ 'c) \ real" assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" - assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py" - assumes Pxy: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" + assumes Px[measurable]: "distributed M S X Px" and Py[measurable]: "distributed M T Y Py" + assumes Pxy[measurable]: "distributed M (S \\<^sub>M T) (\x. (X x, Y x)) Pxy" assumes Ix: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Px (fst x)))" assumes Iy: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Py (snd x)))" assumes Ixy: "integrable(S \\<^sub>M T) (\x. Pxy x * log b (Pxy x))" @@ -1758,13 +1766,13 @@ have "entropy b S X + entropy b T Y - entropy b (S \\<^sub>M T) (\x. (X x, Y x)) = integral\<^sup>L (S \\<^sub>M T) ?f" unfolding X Y XY apply (subst integral_diff) - apply (intro integral_diff Ixy Ix Iy)+ + apply (intro integrable_diff Ixy Ix Iy)+ apply (subst integral_diff) - apply (intro integral_diff Ixy Ix Iy)+ + apply (intro Ixy Ix Iy)+ apply (simp add: field_simps) done also have "\ = integral\<^sup>L (S \\<^sub>M T) ?g" - using `AE x in _. ?f x = ?g x` by (rule integral_cong_AE) + using `AE x in _. ?f x = ?g x` by (intro integral_cong_AE) auto also have "\ = mutual_information b S T X Y" by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric]) finally show ?thesis .. diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Mon May 19 11:27:02 2014 +0200 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2900 +0,0 @@ -(* Title: HOL/Probability/Lebesgue_Integration.thy - Author: Johannes Hölzl, TU München - Author: Armin Heller, TU München -*) - -header {*Lebesgue Integration*} - -theory Lebesgue_Integration - imports Measure_Space Borel_Space -begin - -lemma indicator_less_ereal[simp]: - "indicator A x \ (indicator B x::ereal) \ (x \ A \ x \ B)" - by (simp add: indicator_def not_le) - -section "Simple function" - -text {* - -Our simple functions are not restricted to positive real numbers. Instead -they are just functions with a finite range and are measurable when singleton -sets are measurable. - -*} - -definition "simple_function M g \ - finite (g ` space M) \ - (\x \ g ` space M. g -` {x} \ space M \ sets M)" - -lemma simple_functionD: - assumes "simple_function M g" - shows "finite (g ` space M)" and "g -` X \ space M \ sets M" -proof - - show "finite (g ` space M)" - using assms unfolding simple_function_def by auto - have "g -` X \ space M = g -` (X \ g`space M) \ space M" by auto - also have "\ = (\x\X \ g`space M. g-`{x} \ space M)" by auto - finally show "g -` X \ space M \ sets M" using assms - by (auto simp del: UN_simps simp: simple_function_def) -qed - -lemma measurable_simple_function[measurable_dest]: - "simple_function M f \ f \ measurable M (count_space UNIV)" - unfolding simple_function_def measurable_def -proof safe - fix A assume "finite (f ` space M)" "\x\f ` space M. f -` {x} \ space M \ sets M" - then have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) \ sets M" - by (intro sets.finite_UN) auto - also have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) = f -` A \ space M" - by (auto split: split_if_asm) - finally show "f -` A \ space M \ sets M" . -qed simp - -lemma borel_measurable_simple_function: - "simple_function M f \ f \ borel_measurable M" - by (auto dest!: measurable_simple_function simp: measurable_def) - -lemma simple_function_measurable2[intro]: - assumes "simple_function M f" "simple_function M g" - shows "f -` A \ g -` B \ space M \ sets M" -proof - - have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" - by auto - then show ?thesis using assms[THEN simple_functionD(2)] by auto -qed - -lemma simple_function_indicator_representation: - fixes f ::"'a \ ereal" - assumes f: "simple_function M f" and x: "x \ space M" - shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" - (is "?l = ?r") -proof - - have "?r = (\y \ f ` space M. - (if y = f x then y * indicator (f -` {y} \ space M) x else 0))" - by (auto intro!: setsum_cong2) - also have "... = f x * indicator (f -` {f x} \ space M) x" - using assms by (auto dest: simple_functionD simp: setsum_delta) - also have "... = f x" using x by (auto simp: indicator_def) - finally show ?thesis by auto -qed - -lemma simple_function_notspace: - "simple_function M (\x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h") -proof - - have "?h ` space M \ {0}" unfolding indicator_def by auto - hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) - have "?h -` {0} \ space M = space M" by auto - thus ?thesis unfolding simple_function_def by auto -qed - -lemma simple_function_cong: - assumes "\t. t \ space M \ f t = g t" - shows "simple_function M f \ simple_function M g" -proof - - have "f ` space M = g ` space M" - "\x. f -` {x} \ space M = g -` {x} \ space M" - using assms by (auto intro!: image_eqI) - thus ?thesis unfolding simple_function_def using assms by simp -qed - -lemma simple_function_cong_algebra: - assumes "sets N = sets M" "space N = space M" - shows "simple_function M f \ simple_function N f" - unfolding simple_function_def assms .. - -lemma simple_function_borel_measurable: - fixes f :: "'a \ 'x::{t2_space}" - assumes "f \ borel_measurable M" and "finite (f ` space M)" - shows "simple_function M f" - using assms unfolding simple_function_def - by (auto intro: borel_measurable_vimage) - -lemma simple_function_eq_measurable: - fixes f :: "'a \ ereal" - shows "simple_function M f \ finite (f`space M) \ f \ measurable M (count_space UNIV)" - using simple_function_borel_measurable[of f] measurable_simple_function[of M f] - by (fastforce simp: simple_function_def) - -lemma simple_function_const[intro, simp]: - "simple_function M (\x. c)" - by (auto intro: finite_subset simp: simple_function_def) -lemma simple_function_compose[intro, simp]: - assumes "simple_function M f" - shows "simple_function M (g \ f)" - unfolding simple_function_def -proof safe - show "finite ((g \ f) ` space M)" - using assms unfolding simple_function_def by (auto simp: image_comp [symmetric]) -next - fix x assume "x \ space M" - let ?G = "g -` {g (f x)} \ (f`space M)" - have *: "(g \ f) -` {(g \ f) x} \ space M = - (\x\?G. f -` {x} \ space M)" by auto - show "(g \ f) -` {(g \ f) x} \ space M \ sets M" - using assms unfolding simple_function_def * - by (rule_tac sets.finite_UN) auto -qed - -lemma simple_function_indicator[intro, simp]: - assumes "A \ sets M" - shows "simple_function M (indicator A)" -proof - - have "indicator A ` space M \ {0, 1}" (is "?S \ _") - by (auto simp: indicator_def) - hence "finite ?S" by (rule finite_subset) simp - moreover have "- A \ space M = space M - A" by auto - ultimately show ?thesis unfolding simple_function_def - using assms by (auto simp: indicator_def [abs_def]) -qed - -lemma simple_function_Pair[intro, simp]: - assumes "simple_function M f" - assumes "simple_function M g" - shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") - unfolding simple_function_def -proof safe - show "finite (?p ` space M)" - using assms unfolding simple_function_def - by (rule_tac finite_subset[of _ "f`space M \ g`space M"]) auto -next - fix x assume "x \ space M" - have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = - (f -` {f x} \ space M) \ (g -` {g x} \ space M)" - by auto - with `x \ space M` show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" - using assms unfolding simple_function_def by auto -qed - -lemma simple_function_compose1: - assumes "simple_function M f" - shows "simple_function M (\x. g (f x))" - using simple_function_compose[OF assms, of g] - by (simp add: comp_def) - -lemma simple_function_compose2: - assumes "simple_function M f" and "simple_function M g" - shows "simple_function M (\x. h (f x) (g x))" -proof - - have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" - using assms by auto - thus ?thesis by (simp_all add: comp_def) -qed - -lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] - and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] - and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] - and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] - and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] - and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] - and simple_function_max[intro, simp] = simple_function_compose2[where h=max] - -lemma simple_function_setsum[intro, simp]: - assumes "\i. i \ P \ simple_function M (f i)" - shows "simple_function M (\x. \i\P. f i x)" -proof cases - assume "finite P" from this assms show ?thesis by induct auto -qed auto - -lemma simple_function_ereal[intro, simp]: - fixes f g :: "'a \ real" assumes sf: "simple_function M f" - shows "simple_function M (\x. ereal (f x))" - by (auto intro!: simple_function_compose1[OF sf]) - -lemma simple_function_real_of_nat[intro, simp]: - fixes f g :: "'a \ nat" assumes sf: "simple_function M f" - shows "simple_function M (\x. real (f x))" - by (auto intro!: simple_function_compose1[OF sf]) - -lemma borel_measurable_implies_simple_function_sequence: - fixes u :: "'a \ ereal" - assumes u: "u \ borel_measurable M" - shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ - (\x. (SUP i. f i x) = max 0 (u x)) \ (\i x. 0 \ f i x)" -proof - - def f \ "\x i. if real i \ u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)" - { fix x j have "f x j \ j * 2 ^ j" unfolding f_def - proof (split split_if, intro conjI impI) - assume "\ real j \ u x" - then have "natfloor (real (u x) * 2 ^ j) \ natfloor (j * 2 ^ j)" - by (cases "u x") (auto intro!: natfloor_mono) - moreover have "real (natfloor (j * 2 ^ j)) \ j * 2^j" - by (intro real_natfloor_le) auto - ultimately show "natfloor (real (u x) * 2 ^ j) \ j * 2 ^ j" - unfolding real_of_nat_le_iff by auto - qed auto } - note f_upper = this - - have real_f: - "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))" - unfolding f_def by auto - - let ?g = "\j x. real (f x j) / 2^j :: ereal" - show ?thesis - proof (intro exI[of _ ?g] conjI allI ballI) - fix i - have "simple_function M (\x. real (f x i))" - proof (intro simple_function_borel_measurable) - show "(\x. real (f x i)) \ borel_measurable M" - using u by (auto simp: real_f) - have "(\x. real (f x i))`space M \ real`{..i*2^i}" - using f_upper[of _ i] by auto - then show "finite ((\x. real (f x i))`space M)" - by (rule finite_subset) auto - qed - then show "simple_function M (?g i)" - by (auto intro: simple_function_ereal simple_function_div) - next - show "incseq ?g" - proof (intro incseq_ereal incseq_SucI le_funI) - fix x and i :: nat - have "f x i * 2 \ f x (Suc i)" unfolding f_def - proof ((split split_if)+, intro conjI impI) - assume "ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" - then show "i * 2 ^ i * 2 \ natfloor (real (u x) * 2 ^ Suc i)" - by (cases "u x") (auto intro!: le_natfloor) - next - assume "\ ereal (real i) \ u x" "ereal (real (Suc i)) \ u x" - then show "natfloor (real (u x) * 2 ^ i) * 2 \ Suc i * 2 ^ Suc i" - by (cases "u x") auto - next - assume "\ ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" - have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2" - by simp - also have "\ \ natfloor (real (u x) * 2 ^ i * 2)" - proof cases - assume "0 \ u x" then show ?thesis - by (intro le_mult_natfloor) - next - assume "\ 0 \ u x" then show ?thesis - by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg) - qed - also have "\ = natfloor (real (u x) * 2 ^ Suc i)" - by (simp add: ac_simps) - finally show "natfloor (real (u x) * 2 ^ i) * 2 \ natfloor (real (u x) * 2 ^ Suc i)" . - qed simp - then show "?g i x \ ?g (Suc i) x" - by (auto simp: field_simps) - qed - next - fix x show "(SUP i. ?g i x) = max 0 (u x)" - proof (rule SUP_eqI) - fix i show "?g i x \ max 0 (u x)" unfolding max_def f_def - by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg - mult_nonpos_nonneg) - next - fix y assume *: "\i. i \ UNIV \ ?g i x \ y" - have "\i. 0 \ ?g i x" by auto - from order_trans[OF this *] have "0 \ y" by simp - show "max 0 (u x) \ y" - proof (cases y) - case (real r) - with * have *: "\i. f x i \ r * 2^i" by (auto simp: divide_le_eq) - from reals_Archimedean2[of r] * have "u x \ \" by (auto simp: f_def) (metis less_le_not_le) - then have "\p. max 0 (u x) = ereal p \ 0 \ p" by (cases "u x") (auto simp: max_def) - then guess p .. note ux = this - obtain m :: nat where m: "p < real m" using reals_Archimedean2 .. - have "p \ r" - proof (rule ccontr) - assume "\ p \ r" - with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] - obtain N where "\n\N. r * 2^n < p * 2^n - 1" by (auto simp: field_simps) - then have "r * 2^max N m < p * 2^max N m - 1" by simp - moreover - have "real (natfloor (p * 2 ^ max N m)) \ r * 2 ^ max N m" - using *[of "max N m"] m unfolding real_f using ux - by (cases "0 \ u x") (simp_all add: max_def split: split_if_asm) - then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" - by (metis real_natfloor_gt_diff_one less_le_trans) - ultimately show False by auto - qed - then show "max 0 (u x) \ y" using real ux by simp - qed (insert `0 \ y`, auto) - qed - qed auto -qed - -lemma borel_measurable_implies_simple_function_sequence': - fixes u :: "'a \ ereal" - assumes u: "u \ borel_measurable M" - obtains f where "\i. simple_function M (f i)" "incseq f" "\i. \ \ range (f i)" - "\x. (SUP i. f i x) = max 0 (u x)" "\i x. 0 \ f i x" - using borel_measurable_implies_simple_function_sequence[OF u] by auto - -lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]: - fixes u :: "'a \ ereal" - assumes u: "simple_function M u" - assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" - assumes set: "\A. A \ sets M \ P (indicator A)" - assumes mult: "\u c. P u \ P (\x. c * u x)" - assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" - shows "P u" -proof (rule cong) - from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" - proof eventually_elim - fix x assume x: "x \ space M" - from simple_function_indicator_representation[OF u x] - show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. - qed -next - from u have "finite (u ` space M)" - unfolding simple_function_def by auto - then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" - proof induct - case empty show ?case - using set[of "{}"] by (simp add: indicator_def[abs_def]) - qed (auto intro!: add mult set simple_functionD u) -next - show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" - apply (subst simple_function_cong) - apply (rule simple_function_indicator_representation[symmetric]) - apply (auto intro: u) - done -qed fact - -lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: - fixes u :: "'a \ ereal" - assumes u: "simple_function M u" and nn: "\x. 0 \ u x" - assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" - assumes set: "\A. A \ sets M \ P (indicator A)" - assumes mult: "\u c. 0 \ c \ simple_function M u \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" - assumes add: "\u v. simple_function M u \ (\x. 0 \ u x) \ P u \ simple_function M v \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" - shows "P u" -proof - - show ?thesis - proof (rule cong) - fix x assume x: "x \ space M" - from simple_function_indicator_representation[OF u x] - show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. - next - show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" - apply (subst simple_function_cong) - apply (rule simple_function_indicator_representation[symmetric]) - apply (auto intro: u) - done - next - from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" - unfolding simple_function_def by auto - then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" - proof induct - case empty show ?case - using set[of "{}"] by (simp add: indicator_def[abs_def]) - qed (auto intro!: add mult set simple_functionD u setsum_nonneg - simple_function_setsum) - qed fact -qed - -lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: - fixes u :: "'a \ ereal" - assumes u: "u \ borel_measurable M" "\x. 0 \ u x" - assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" - assumes set: "\A. A \ sets M \ P (indicator A)" - assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" - assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" - assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i. P (U i)) \ incseq U \ P (SUP i. U i)" - shows "P u" - using u -proof (induct rule: borel_measurable_implies_simple_function_sequence') - fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and - sup: "\x. (SUP i. U i x) = max 0 (u x)" and nn: "\i x. 0 \ U i x" - have u_eq: "u = (SUP i. U i)" - using nn u sup by (auto simp: max_def) - - from U have "\i. U i \ borel_measurable M" - by (simp add: borel_measurable_simple_function) - - show "P u" - unfolding u_eq - proof (rule seq) - fix i show "P (U i)" - using `simple_function M (U i)` nn - by (induct rule: simple_function_induct_nn) - (auto intro: set mult add cong dest!: borel_measurable_simple_function) - qed fact+ -qed - -lemma simple_function_If_set: - assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" - shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") -proof - - def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" - show ?thesis unfolding simple_function_def - proof safe - have "?IF ` space M \ f ` space M \ g ` space M" by auto - from finite_subset[OF this] assms - show "finite (?IF ` space M)" unfolding simple_function_def by auto - next - fix x assume "x \ space M" - then have *: "?IF -` {?IF x} \ space M = (if x \ A - then ((F (f x) \ (A \ space M)) \ (G (f x) - (G (f x) \ (A \ space M)))) - else ((F (g x) \ (A \ space M)) \ (G (g x) - (G (g x) \ (A \ space M)))))" - using sets.sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) - have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" - unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto - show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto - qed -qed - -lemma simple_function_If: - assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" - shows "simple_function M (\x. if P x then f x else g x)" -proof - - have "{x\space M. P x} = {x. P x} \ space M" by auto - with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp -qed - -lemma simple_function_subalgebra: - assumes "simple_function N f" - and N_subalgebra: "sets N \ sets M" "space N = space M" - shows "simple_function M f" - using assms unfolding simple_function_def by auto - -lemma simple_function_comp: - assumes T: "T \ measurable M M'" - and f: "simple_function M' f" - shows "simple_function M (\x. f (T x))" -proof (intro simple_function_def[THEN iffD2] conjI ballI) - have "(\x. f (T x)) ` space M \ f ` space M'" - using T unfolding measurable_def by auto - then show "finite ((\x. f (T x)) ` space M)" - using f unfolding simple_function_def by (auto intro: finite_subset) - fix i assume i: "i \ (\x. f (T x)) ` space M" - then have "i \ f ` space M'" - using T unfolding measurable_def by auto - then have "f -` {i} \ space M' \ sets M'" - using f unfolding simple_function_def by auto - then have "T -` (f -` {i} \ space M') \ space M \ sets M" - using T unfolding measurable_def by auto - also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" - using T unfolding measurable_def by auto - finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . -qed - -section "Simple integral" - -definition simple_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>S") where - "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" - -syntax - "_simple_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>S _. _ \_" [60,61] 110) - -translations - "\\<^sup>S x. f \M" == "CONST simple_integral M (%x. f)" - -lemma simple_integral_cong: - assumes "\t. t \ space M \ f t = g t" - shows "integral\<^sup>S M f = integral\<^sup>S M g" -proof - - have "f ` space M = g ` space M" - "\x. f -` {x} \ space M = g -` {x} \ space M" - using assms by (auto intro!: image_eqI) - thus ?thesis unfolding simple_integral_def by simp -qed - -lemma simple_integral_const[simp]: - "(\\<^sup>Sx. c \M) = c * (emeasure M) (space M)" -proof (cases "space M = {}") - case True thus ?thesis unfolding simple_integral_def by simp -next - case False hence "(\x. c) ` space M = {c}" by auto - thus ?thesis unfolding simple_integral_def by simp -qed - -lemma simple_function_partition: - assumes f: "simple_function M f" and g: "simple_function M g" - assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" - assumes v: "\x. x \ space M \ f x = v (g x)" - shows "integral\<^sup>S M f = (\y\g ` space M. v y * emeasure M {x\space M. g x = y})" - (is "_ = ?r") -proof - - from f g have [simp]: "finite (f`space M)" "finite (g`space M)" - by (auto simp: simple_function_def) - from f g have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" - by (auto intro: measurable_simple_function) - - { fix y assume "y \ space M" - then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" - by (auto cong: sub simp: v[symmetric]) } - note eq = this - - have "integral\<^sup>S M f = - (\y\f`space M. y * (\z\g`space M. - if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" - unfolding simple_integral_def - proof (safe intro!: setsum_cong ereal_left_mult_cong) - fix y assume y: "y \ space M" "f y \ 0" - have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = - {z. \x\space M. f y = f x \ z = g x}" - by auto - have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = - f -` {f y} \ space M" - by (auto simp: eq_commute cong: sub rev_conj_cong) - have "finite (g`space M)" by simp - then have "finite {z. \x\space M. f y = f x \ z = g x}" - by (rule rev_finite_subset) auto - then show "emeasure M (f -` {f y} \ space M) = - (\z\g ` space M. if \x\space M. f y = f x \ z = g x then emeasure M {x \ space M. g x = z} else 0)" - apply (simp add: setsum_cases) - apply (subst setsum_emeasure) - apply (auto simp: disjoint_family_on_def eq) - done - qed - also have "\ = (\y\f`space M. (\z\g`space M. - if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" - by (auto intro!: setsum_cong simp: setsum_ereal_right_distrib emeasure_nonneg) - also have "\ = ?r" - by (subst setsum_commute) - (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq) - finally show "integral\<^sup>S M f = ?r" . -qed - -lemma simple_integral_add[simp]: - assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" - shows "(\\<^sup>Sx. f x + g x \M) = integral\<^sup>S M f + integral\<^sup>S M g" -proof - - have "(\\<^sup>Sx. f x + g x \M) = - (\y\(\x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\space M. (f x, g x) = y})" - by (intro simple_function_partition) (auto intro: f g) - also have "\ = (\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) + - (\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y})" - using assms(2,4) by (auto intro!: setsum_cong ereal_left_distrib simp: setsum_addf[symmetric]) - also have "(\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. f x \M)" - by (intro simple_function_partition[symmetric]) (auto intro: f g) - also have "(\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. g x \M)" - by (intro simple_function_partition[symmetric]) (auto intro: f g) - finally show ?thesis . -qed - -lemma simple_integral_setsum[simp]: - assumes "\i x. i \ P \ 0 \ f i x" - assumes "\i. i \ P \ simple_function M (f i)" - shows "(\\<^sup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^sup>S M (f i))" -proof cases - assume "finite P" - from this assms show ?thesis - by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) -qed auto - -lemma simple_integral_mult[simp]: - assumes f: "simple_function M f" "\x. 0 \ f x" - shows "(\\<^sup>Sx. c * f x \M) = c * integral\<^sup>S M f" -proof - - have "(\\<^sup>Sx. c * f x \M) = (\y\f ` space M. (c * y) * emeasure M {x\space M. f x = y})" - using f by (intro simple_function_partition) auto - also have "\ = c * integral\<^sup>S M f" - using f unfolding simple_integral_def - by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult_assoc Int_def conj_commute) - finally show ?thesis . -qed - -lemma simple_integral_mono_AE: - assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" - and mono: "AE x in M. f x \ g x" - shows "integral\<^sup>S M f \ integral\<^sup>S M g" -proof - - let ?\ = "\P. emeasure M {x\space M. P x}" - have "integral\<^sup>S M f = (\y\(\x. (f x, g x))`space M. fst y * ?\ (\x. (f x, g x) = y))" - using f g by (intro simple_function_partition) auto - also have "\ \ (\y\(\x. (f x, g x))`space M. snd y * ?\ (\x. (f x, g x) = y))" - proof (clarsimp intro!: setsum_mono) - fix x assume "x \ space M" - let ?M = "?\ (\y. f y = f x \ g y = g x)" - show "f x * ?M \ g x * ?M" - proof cases - assume "?M \ 0" - then have "0 < ?M" - by (simp add: less_le emeasure_nonneg) - also have "\ \ ?\ (\y. f x \ g x)" - using mono by (intro emeasure_mono_AE) auto - finally have "\ \ f x \ g x" - by (intro notI) auto - then show ?thesis - by (intro ereal_mult_right_mono) auto - qed simp - qed - also have "\ = integral\<^sup>S M g" - using f g by (intro simple_function_partition[symmetric]) auto - finally show ?thesis . -qed - -lemma simple_integral_mono: - assumes "simple_function M f" and "simple_function M g" - and mono: "\ x. x \ space M \ f x \ g x" - shows "integral\<^sup>S M f \ integral\<^sup>S M g" - using assms by (intro simple_integral_mono_AE) auto - -lemma simple_integral_cong_AE: - assumes "simple_function M f" and "simple_function M g" - and "AE x in M. f x = g x" - shows "integral\<^sup>S M f = integral\<^sup>S M g" - using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) - -lemma simple_integral_cong': - assumes sf: "simple_function M f" "simple_function M g" - and mea: "(emeasure M) {x\space M. f x \ g x} = 0" - shows "integral\<^sup>S M f = integral\<^sup>S M g" -proof (intro simple_integral_cong_AE sf AE_I) - show "(emeasure M) {x\space M. f x \ g x} = 0" by fact - show "{x \ space M. f x \ g x} \ sets M" - using sf[THEN borel_measurable_simple_function] by auto -qed simp - -lemma simple_integral_indicator: - assumes A: "A \ sets M" - assumes f: "simple_function M f" - shows "(\\<^sup>Sx. f x * indicator A x \M) = - (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" -proof - - have eq: "(\x. (f x, indicator A x)) ` space M \ {x. snd x = 1} = (\x. (f x, 1::ereal))`A" - using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: split_if_asm) - have eq2: "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" - by (auto simp: image_iff) - - have "(\\<^sup>Sx. f x * indicator A x \M) = - (\y\(\x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\space M. (f x, indicator A x) = y})" - using assms by (intro simple_function_partition) auto - also have "\ = (\y\(\x. (f x, indicator A x::ereal))`space M. - if snd y = 1 then fst y * emeasure M (f -` {fst y} \ space M \ A) else 0)" - by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum_cong) - also have "\ = (\y\(\x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \ space M \ A))" - using assms by (subst setsum_cases) (auto intro!: simple_functionD(1) simp: eq) - also have "\ = (\y\fst`(\x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \ space M \ A))" - by (subst setsum_reindex[where f=fst]) (auto simp: inj_on_def) - also have "\ = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" - using A[THEN sets.sets_into_space] - by (intro setsum_mono_zero_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) - finally show ?thesis . -qed - -lemma simple_integral_indicator_only[simp]: - assumes "A \ sets M" - shows "integral\<^sup>S M (indicator A) = emeasure M A" - using simple_integral_indicator[OF assms, of "\x. 1"] sets.sets_into_space[OF assms] - by (simp_all add: image_constant_conv Int_absorb1 split: split_if_asm) - -lemma simple_integral_null_set: - assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" - shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" -proof - - have "AE x in M. indicator N x = (0 :: ereal)" - using `N \ null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) - then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" - using assms apply (intro simple_integral_cong_AE) by auto - then show ?thesis by simp -qed - -lemma simple_integral_cong_AE_mult_indicator: - assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" - shows "integral\<^sup>S M f = (\\<^sup>Sx. f x * indicator S x \M)" - using assms by (intro simple_integral_cong_AE) auto - -lemma simple_integral_cmult_indicator: - assumes A: "A \ sets M" - shows "(\\<^sup>Sx. c * indicator A x \M) = c * emeasure M A" - using simple_integral_mult[OF simple_function_indicator[OF A]] - unfolding simple_integral_indicator_only[OF A] by simp - -lemma simple_integral_positive: - assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" - shows "0 \ integral\<^sup>S M f" -proof - - have "integral\<^sup>S M (\x. 0) \ integral\<^sup>S M f" - using simple_integral_mono_AE[OF _ f ae] by auto - then show ?thesis by simp -qed - -section "Continuous positive integration" - -definition positive_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>P") where - "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^sup>S M g)" - -syntax - "_positive_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>+ _. _ \_" [60,61] 110) - -translations - "\\<^sup>+ x. f \M" == "CONST positive_integral M (%x. f)" - -lemma positive_integral_positive: - "0 \ integral\<^sup>P M f" - by (auto intro!: SUP_upper2[of "\x. 0"] simp: positive_integral_def le_fun_def) - -lemma positive_integral_not_MInfty[simp]: "integral\<^sup>P M f \ -\" - using positive_integral_positive[of M f] by auto - -lemma positive_integral_def_finite: - "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f \ range g \ {0 ..< \}}. integral\<^sup>S M g)" - (is "_ = SUPREMUM ?A ?f") - unfolding positive_integral_def -proof (safe intro!: antisym SUP_least) - fix g assume g: "simple_function M g" "g \ max 0 \ f" - let ?G = "{x \ space M. \ g x \ \}" - note gM = g(1)[THEN borel_measurable_simple_function] - have \_G_pos: "0 \ (emeasure M) ?G" using gM by auto - let ?g = "\y x. if g x = \ then y else max 0 (g x)" - from g gM have g_in_A: "\y. 0 \ y \ y \ \ \ ?g y \ ?A" - apply (safe intro!: simple_function_max simple_function_If) - apply (force simp: max_def le_fun_def split: split_if_asm)+ - done - show "integral\<^sup>S M g \ SUPREMUM ?A ?f" - proof cases - have g0: "?g 0 \ ?A" by (intro g_in_A) auto - assume "(emeasure M) ?G = 0" - with gM have "AE x in M. x \ ?G" - by (auto simp add: AE_iff_null intro!: null_setsI) - with gM g show ?thesis - by (intro SUP_upper2[OF g0] simple_integral_mono_AE) - (auto simp: max_def intro!: simple_function_If) - next - assume \_G: "(emeasure M) ?G \ 0" - have "SUPREMUM ?A (integral\<^sup>S M) = \" - proof (intro SUP_PInfty) - fix n :: nat - let ?y = "ereal (real n) / (if (emeasure M) ?G = \ then 1 else (emeasure M) ?G)" - have "0 \ ?y" "?y \ \" using \_G \_G_pos by (auto simp: ereal_divide_eq) - then have "?g ?y \ ?A" by (rule g_in_A) - have "real n \ ?y * (emeasure M) ?G" - using \_G \_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) - also have "\ = (\\<^sup>Sx. ?y * indicator ?G x \M)" - using `0 \ ?y` `?g ?y \ ?A` gM - by (subst simple_integral_cmult_indicator) auto - also have "\ \ integral\<^sup>S M (?g ?y)" using `?g ?y \ ?A` gM - by (intro simple_integral_mono) auto - finally show "\i\?A. real n \ integral\<^sup>S M i" - using `?g ?y \ ?A` by blast - qed - then show ?thesis by simp - qed -qed (auto intro: SUP_upper) - -lemma positive_integral_mono_AE: - assumes ae: "AE x in M. u x \ v x" shows "integral\<^sup>P M u \ integral\<^sup>P M v" - unfolding positive_integral_def -proof (safe intro!: SUP_mono) - fix n assume n: "simple_function M n" "n \ max 0 \ u" - from ae[THEN AE_E] guess N . note N = this - then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) - let ?n = "\x. n x * indicator (space M - N) x" - have "AE x in M. n x \ ?n x" "simple_function M ?n" - using n N ae_N by auto - moreover - { fix x have "?n x \ max 0 (v x)" - proof cases - assume x: "x \ space M - N" - with N have "u x \ v x" by auto - with n(2)[THEN le_funD, of x] x show ?thesis - by (auto simp: max_def split: split_if_asm) - qed simp } - then have "?n \ max 0 \ v" by (auto simp: le_funI) - moreover have "integral\<^sup>S M n \ integral\<^sup>S M ?n" - using ae_N N n by (auto intro!: simple_integral_mono_AE) - ultimately show "\m\{g. simple_function M g \ g \ max 0 \ v}. integral\<^sup>S M n \ integral\<^sup>S M m" - by force -qed - -lemma positive_integral_mono: - "(\x. x \ space M \ u x \ v x) \ integral\<^sup>P M u \ integral\<^sup>P M v" - by (auto intro: positive_integral_mono_AE) - -lemma positive_integral_cong_AE: - "AE x in M. u x = v x \ integral\<^sup>P M u = integral\<^sup>P M v" - by (auto simp: eq_iff intro!: positive_integral_mono_AE) - -lemma positive_integral_cong: - "(\x. x \ space M \ u x = v x) \ integral\<^sup>P M u = integral\<^sup>P M v" - by (auto intro: positive_integral_cong_AE) - -lemma positive_integral_eq_simple_integral: - assumes f: "simple_function M f" "\x. 0 \ f x" shows "integral\<^sup>P M f = integral\<^sup>S M f" -proof - - let ?f = "\x. f x * indicator (space M) x" - have f': "simple_function M ?f" using f by auto - with f(2) have [simp]: "max 0 \ ?f = ?f" - by (auto simp: fun_eq_iff max_def split: split_indicator) - have "integral\<^sup>P M ?f \ integral\<^sup>S M ?f" using f' - by (force intro!: SUP_least simple_integral_mono simp: le_fun_def positive_integral_def) - moreover have "integral\<^sup>S M ?f \ integral\<^sup>P M ?f" - unfolding positive_integral_def - using f' by (auto intro!: SUP_upper) - ultimately show ?thesis - by (simp cong: positive_integral_cong simple_integral_cong) -qed - -lemma positive_integral_eq_simple_integral_AE: - assumes f: "simple_function M f" "AE x in M. 0 \ f x" shows "integral\<^sup>P M f = integral\<^sup>S M f" -proof - - have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) - with f have "integral\<^sup>P M f = integral\<^sup>S M (\x. max 0 (f x))" - by (simp cong: positive_integral_cong_AE simple_integral_cong_AE - add: positive_integral_eq_simple_integral) - with assms show ?thesis - by (auto intro!: simple_integral_cong_AE split: split_max) -qed - -lemma positive_integral_SUP_approx: - assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" - and u: "simple_function M u" "u \ (SUP i. f i)" "u`space M \ {0..<\}" - shows "integral\<^sup>S M u \ (SUP i. integral\<^sup>P M (f i))" (is "_ \ ?S") -proof (rule ereal_le_mult_one_interval) - have "0 \ (SUP i. integral\<^sup>P M (f i))" - using f(3) by (auto intro!: SUP_upper2 positive_integral_positive) - then show "(SUP i. integral\<^sup>P M (f i)) \ -\" by auto - have u_range: "\x. x \ space M \ 0 \ u x \ u x \ \" - using u(3) by auto - fix a :: ereal assume "0 < a" "a < 1" - hence "a \ 0" by auto - let ?B = "\i. {x \ space M. a * u x \ f i x}" - have B: "\i. ?B i \ sets M" - using f `simple_function M u`[THEN borel_measurable_simple_function] by auto - - let ?uB = "\i x. u x * indicator (?B i) x" - - { fix i have "?B i \ ?B (Suc i)" - proof safe - fix i x assume "a * u x \ f i x" - also have "\ \ f (Suc i) x" - using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto - finally show "a * u x \ f (Suc i) x" . - qed } - note B_mono = this - - note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B] - - let ?B' = "\i n. (u -` {i} \ space M) \ ?B n" - have measure_conv: "\i. (emeasure M) (u -` {i} \ space M) = (SUP n. (emeasure M) (?B' i n))" - proof - - fix i - have 1: "range (?B' i) \ sets M" using B_u by auto - have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI) - have "(\n. ?B' i n) = u -` {i} \ space M" - proof safe - fix x i assume x: "x \ space M" - show "x \ (\i. ?B' (u x) i)" - proof cases - assume "u x = 0" thus ?thesis using `x \ space M` f(3) by simp - next - assume "u x \ 0" - with `a < 1` u_range[OF `x \ space M`] - have "a * u x < 1 * u x" - by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) - also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) - finally obtain i where "a * u x < f i x" unfolding SUP_def - by (auto simp add: less_SUP_iff) - hence "a * u x \ f i x" by auto - thus ?thesis using `x \ space M` by auto - qed - qed - then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp - qed - - have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" - unfolding simple_integral_indicator[OF B `simple_function M u`] - proof (subst SUP_ereal_setsum, safe) - fix x n assume "x \ space M" - with u_range show "incseq (\i. u x * (emeasure M) (?B' (u x) i))" "\i. 0 \ u x * (emeasure M) (?B' (u x) i)" - using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) - next - show "integral\<^sup>S M u = (\i\u ` space M. SUP n. i * (emeasure M) (?B' i n))" - using measure_conv u_range B_u unfolding simple_integral_def - by (auto intro!: setsum_cong SUP_ereal_cmult [symmetric]) - qed - moreover - have "a * (SUP i. integral\<^sup>S M (?uB i)) \ ?S" - apply (subst SUP_ereal_cmult [symmetric]) - proof (safe intro!: SUP_mono bexI) - fix i - have "a * integral\<^sup>S M (?uB i) = (\\<^sup>Sx. a * ?uB i x \M)" - using B `simple_function M u` u_range - by (subst simple_integral_mult) (auto split: split_indicator) - also have "\ \ integral\<^sup>P M (f i)" - proof - - have *: "simple_function M (\x. a * ?uB i x)" using B `0 < a` u(1) by auto - show ?thesis using f(3) * u_range `0 < a` - by (subst positive_integral_eq_simple_integral[symmetric]) - (auto intro!: positive_integral_mono split: split_indicator) - qed - finally show "a * integral\<^sup>S M (?uB i) \ integral\<^sup>P M (f i)" - by auto - next - fix i show "0 \ \\<^sup>S x. ?uB i x \M" using B `0 < a` u(1) u_range - by (intro simple_integral_positive) (auto split: split_indicator) - qed (insert `0 < a`, auto) - ultimately show "a * integral\<^sup>S M u \ ?S" by simp -qed - -lemma incseq_positive_integral: - assumes "incseq f" shows "incseq (\i. integral\<^sup>P M (f i))" -proof - - have "\i x. f i x \ f (Suc i) x" - using assms by (auto dest!: incseq_SucD simp: le_fun_def) - then show ?thesis - by (auto intro!: incseq_SucI positive_integral_mono) -qed - -text {* Beppo-Levi monotone convergence theorem *} -lemma positive_integral_monotone_convergence_SUP: - assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" - shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" -proof (rule antisym) - show "(SUP j. integral\<^sup>P M (f j)) \ (\\<^sup>+ x. (SUP i. f i x) \M)" - by (auto intro!: SUP_least SUP_upper positive_integral_mono) -next - show "(\\<^sup>+ x. (SUP i. f i x) \M) \ (SUP j. integral\<^sup>P M (f j))" - unfolding positive_integral_def_finite[of _ "\x. SUP i. f i x"] - proof (safe intro!: SUP_least) - fix g assume g: "simple_function M g" - and *: "g \ max 0 \ (\x. SUP i. f i x)" "range g \ {0..<\}" - then have "\x. 0 \ (SUP i. f i x)" and g': "g`space M \ {0..<\}" - using f by (auto intro!: SUP_upper2) - with * show "integral\<^sup>S M g \ (SUP j. integral\<^sup>P M (f j))" - by (intro positive_integral_SUP_approx[OF f g _ g']) - (auto simp: le_fun_def max_def) - qed -qed - -lemma positive_integral_monotone_convergence_SUP_AE: - assumes f: "\i. AE x in M. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" - shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" -proof - - from f have "AE x in M. \i. f i x \ f (Suc i) x \ 0 \ f i x" - by (simp add: AE_all_countable) - from this[THEN AE_E] guess N . note N = this - let ?f = "\i x. if x \ space M - N then f i x else 0" - have f_eq: "AE x in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) - then have "(\\<^sup>+ x. (SUP i. f i x) \M) = (\\<^sup>+ x. (SUP i. ?f i x) \M)" - by (auto intro!: positive_integral_cong_AE) - also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \M))" - proof (rule positive_integral_monotone_convergence_SUP) - show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) - { fix i show "(\x. if x \ space M - N then f i x else 0) \ borel_measurable M" - using f N(3) by (intro measurable_If_set) auto - fix x show "0 \ ?f i x" - using N(1) by auto } - qed - also have "\ = (SUP i. (\\<^sup>+ x. f i x \M))" - using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] positive_integral_cong_AE ext) - finally show ?thesis . -qed - -lemma positive_integral_monotone_convergence_SUP_AE_incseq: - assumes f: "incseq f" "\i. AE x in M. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" - shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" - using f[unfolded incseq_Suc_iff le_fun_def] - by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel]) - auto - -lemma positive_integral_monotone_convergence_simple: - assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" - shows "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" - using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1) - f(3)[THEN borel_measurable_simple_function] f(2)] - by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPREMUM UNIV"] ext) - -lemma positive_integral_max_0: - "(\\<^sup>+x. max 0 (f x) \M) = integral\<^sup>P M f" - by (simp add: le_fun_def positive_integral_def) - -lemma positive_integral_cong_pos: - assumes "\x. x \ space M \ f x \ 0 \ g x \ 0 \ f x = g x" - shows "integral\<^sup>P M f = integral\<^sup>P M g" -proof - - have "integral\<^sup>P M (\x. max 0 (f x)) = integral\<^sup>P M (\x. max 0 (g x))" - proof (intro positive_integral_cong) - fix x assume "x \ space M" - from assms[OF this] show "max 0 (f x) = max 0 (g x)" - by (auto split: split_max) - qed - then show ?thesis by (simp add: positive_integral_max_0) -qed - -lemma SUP_simple_integral_sequences: - assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" - and g: "incseq g" "\i x. 0 \ g i x" "\i. simple_function M (g i)" - and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" - shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" - (is "SUPREMUM _ ?F = SUPREMUM _ ?G") -proof - - have "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" - using f by (rule positive_integral_monotone_convergence_simple) - also have "\ = (\\<^sup>+x. (SUP i. g i x) \M)" - unfolding eq[THEN positive_integral_cong_AE] .. - also have "\ = (SUP i. ?G i)" - using g by (rule positive_integral_monotone_convergence_simple[symmetric]) - finally show ?thesis by simp -qed - -lemma positive_integral_const[simp]: - "0 \ c \ (\\<^sup>+ x. c \M) = c * (emeasure M) (space M)" - by (subst positive_integral_eq_simple_integral) auto - -lemma positive_integral_linear: - assumes f: "f \ borel_measurable M" "\x. 0 \ f x" and "0 \ a" - and g: "g \ borel_measurable M" "\x. 0 \ g x" - shows "(\\<^sup>+ x. a * f x + g x \M) = a * integral\<^sup>P M f + integral\<^sup>P M g" - (is "integral\<^sup>P M ?L = _") -proof - - from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . - note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this - from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . - note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this - let ?L' = "\i x. a * u i x + v i x" - - have "?L \ borel_measurable M" using assms by auto - from borel_measurable_implies_simple_function_sequence'[OF this] guess l . - note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this - - have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" - using u v `0 \ a` - by (auto simp: incseq_Suc_iff le_fun_def - intro!: add_mono ereal_mult_left_mono simple_integral_mono) - have pos: "\i. 0 \ integral\<^sup>S M (u i)" "\i. 0 \ integral\<^sup>S M (v i)" "\i. 0 \ a * integral\<^sup>S M (u i)" - using u v `0 \ a` by (auto simp: simple_integral_positive) - { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \ -\" "integral\<^sup>S M (v i) \ -\" - by (auto split: split_if_asm) } - note not_MInf = this - - have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" - proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) - show "incseq ?L'" "\i x. 0 \ ?L' i x" "\i. simple_function M (?L' i)" - using u v `0 \ a` unfolding incseq_Suc_iff le_fun_def - by (auto intro!: add_mono ereal_mult_left_mono) - { fix x - { fix i have "a * u i x \ -\" "v i x \ -\" "u i x \ -\" using `0 \ a` u(6)[of i x] v(6)[of i x] - by auto } - then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" - using `0 \ a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] - by (subst SUP_ereal_cmult [symmetric, OF u(6) `0 \ a`]) - (auto intro!: SUP_ereal_add - simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } - then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" - unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) - by (intro AE_I2) (auto split: split_max) - qed - also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" - using u(2, 6) v(2, 6) `0 \ a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) - finally have "(\\<^sup>+ x. max 0 (a * f x + g x) \M) = a * (\\<^sup>+x. max 0 (f x) \M) + (\\<^sup>+x. max 0 (g x) \M)" - unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] - unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] - apply (subst SUP_ereal_cmult [symmetric, OF pos(1) `0 \ a`]) - apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) . - then show ?thesis by (simp add: positive_integral_max_0) -qed - -lemma positive_integral_cmult: - assumes f: "f \ borel_measurable M" "0 \ c" - shows "(\\<^sup>+ x. c * f x \M) = c * integral\<^sup>P M f" -proof - - have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` - by (auto split: split_max simp: ereal_zero_le_0_iff) - have "(\\<^sup>+ x. c * f x \M) = (\\<^sup>+ x. c * max 0 (f x) \M)" - by (simp add: positive_integral_max_0) - then show ?thesis - using positive_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" _ "\x. 0"] f - by (auto simp: positive_integral_max_0) -qed - -lemma positive_integral_multc: - assumes "f \ borel_measurable M" "0 \ c" - shows "(\\<^sup>+ x. f x * c \M) = integral\<^sup>P M f * c" - unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp - -lemma positive_integral_indicator[simp]: - "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" - by (subst positive_integral_eq_simple_integral) - (auto simp: simple_integral_indicator) - -lemma positive_integral_cmult_indicator: - "0 \ c \ A \ sets M \ (\\<^sup>+ x. c * indicator A x \M) = c * (emeasure M) A" - by (subst positive_integral_eq_simple_integral) - (auto simp: simple_function_indicator simple_integral_indicator) - -lemma positive_integral_indicator': - assumes [measurable]: "A \ space M \ sets M" - shows "(\\<^sup>+ x. indicator A x \M) = emeasure M (A \ space M)" -proof - - have "(\\<^sup>+ x. indicator A x \M) = (\\<^sup>+ x. indicator (A \ space M) x \M)" - by (intro positive_integral_cong) (simp split: split_indicator) - also have "\ = emeasure M (A \ space M)" - by simp - finally show ?thesis . -qed - -lemma positive_integral_add: - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" - and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" - shows "(\\<^sup>+ x. f x + g x \M) = integral\<^sup>P M f + integral\<^sup>P M g" -proof - - have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" - using assms by (auto split: split_max) - have "(\\<^sup>+ x. f x + g x \M) = (\\<^sup>+ x. max 0 (f x + g x) \M)" - by (simp add: positive_integral_max_0) - also have "\ = (\\<^sup>+ x. max 0 (f x) + max 0 (g x) \M)" - unfolding ae[THEN positive_integral_cong_AE] .. - also have "\ = (\\<^sup>+ x. max 0 (f x) \M) + (\\<^sup>+ x. max 0 (g x) \M)" - using positive_integral_linear[of "\x. max 0 (f x)" _ 1 "\x. max 0 (g x)"] f g - by auto - finally show ?thesis - by (simp add: positive_integral_max_0) -qed - -lemma positive_integral_setsum: - assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x in M. 0 \ f i x" - shows "(\\<^sup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^sup>P M (f i))" -proof cases - assume f: "finite P" - from assms have "AE x in M. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto - from f this assms(1) show ?thesis - proof induct - case (insert i P) - then have "f i \ borel_measurable M" "AE x in M. 0 \ f i x" - "(\x. \i\P. f i x) \ borel_measurable M" "AE x in M. 0 \ (\i\P. f i x)" - by (auto intro!: setsum_nonneg) - from positive_integral_add[OF this] - show ?case using insert by auto - qed simp -qed simp - -lemma positive_integral_Markov_inequality: - assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" and "A \ sets M" and c: "0 \ c" - shows "(emeasure M) ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^sup>+ x. u x * indicator A x \M)" - (is "(emeasure M) ?A \ _ * ?PI") -proof - - have "?A \ sets M" - using `A \ sets M` u by auto - hence "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" - using positive_integral_indicator by simp - also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u c - by (auto intro!: positive_integral_mono_AE - simp: indicator_def ereal_zero_le_0_iff) - also have "\ = c * (\\<^sup>+ x. u x * indicator A x \M)" - using assms - by (auto intro!: positive_integral_cmult simp: ereal_zero_le_0_iff) - finally show ?thesis . -qed - -lemma positive_integral_noteq_infinite: - assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" - and "integral\<^sup>P M g \ \" - shows "AE x in M. g x \ \" -proof (rule ccontr) - assume c: "\ (AE x in M. g x \ \)" - have "(emeasure M) {x\space M. g x = \} \ 0" - using c g by (auto simp add: AE_iff_null) - moreover have "0 \ (emeasure M) {x\space M. g x = \}" using g by (auto intro: measurable_sets) - ultimately have "0 < (emeasure M) {x\space M. g x = \}" by auto - then have "\ = \ * (emeasure M) {x\space M. g x = \}" by auto - also have "\ \ (\\<^sup>+x. \ * indicator {x\space M. g x = \} x \M)" - using g by (subst positive_integral_cmult_indicator) auto - also have "\ \ integral\<^sup>P M g" - using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def) - finally show False using `integral\<^sup>P M g \ \` by auto -qed - -lemma positive_integral_PInf: - assumes f: "f \ borel_measurable M" - and not_Inf: "integral\<^sup>P M f \ \" - shows "(emeasure M) (f -` {\} \ space M) = 0" -proof - - have "\ * (emeasure M) (f -` {\} \ space M) = (\\<^sup>+ x. \ * indicator (f -` {\} \ space M) x \M)" - using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets) - also have "\ \ integral\<^sup>P M (\x. max 0 (f x))" - by (auto intro!: positive_integral_mono simp: indicator_def max_def) - finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^sup>P M f" - by (simp add: positive_integral_max_0) - moreover have "0 \ (emeasure M) (f -` {\} \ space M)" - by (rule emeasure_nonneg) - ultimately show ?thesis - using assms by (auto split: split_if_asm) -qed - -lemma positive_integral_PInf_AE: - assumes "f \ borel_measurable M" "integral\<^sup>P M f \ \" shows "AE x in M. f x \ \" -proof (rule AE_I) - show "(emeasure M) (f -` {\} \ space M) = 0" - by (rule positive_integral_PInf[OF assms]) - show "f -` {\} \ space M \ sets M" - using assms by (auto intro: borel_measurable_vimage) -qed auto - -lemma simple_integral_PInf: - assumes "simple_function M f" "\x. 0 \ f x" - and "integral\<^sup>S M f \ \" - shows "(emeasure M) (f -` {\} \ space M) = 0" -proof (rule positive_integral_PInf) - show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) - show "integral\<^sup>P M f \ \" - using assms by (simp add: positive_integral_eq_simple_integral) -qed - -lemma positive_integral_diff: - assumes f: "f \ borel_measurable M" - and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" - and fin: "integral\<^sup>P M g \ \" - and mono: "AE x in M. g x \ f x" - shows "(\\<^sup>+ x. f x - g x \M) = integral\<^sup>P M f - integral\<^sup>P M g" -proof - - have diff: "(\x. f x - g x) \ borel_measurable M" "AE x in M. 0 \ f x - g x" - using assms by (auto intro: ereal_diff_positive) - have pos_f: "AE x in M. 0 \ f x" using mono g by auto - { fix a b :: ereal assume "0 \ a" "a \ \" "0 \ b" "a \ b" then have "b - a + a = b" - by (cases rule: ereal2_cases[of a b]) auto } - note * = this - then have "AE x in M. f x = f x - g x + g x" - using mono positive_integral_noteq_infinite[OF g fin] assms by auto - then have **: "integral\<^sup>P M f = (\\<^sup>+x. f x - g x \M) + integral\<^sup>P M g" - unfolding positive_integral_add[OF diff g, symmetric] - by (rule positive_integral_cong_AE) - show ?thesis unfolding ** - using fin positive_integral_positive[of M g] - by (cases rule: ereal2_cases[of "\\<^sup>+ x. f x - g x \M" "integral\<^sup>P M g"]) auto -qed - -lemma positive_integral_suminf: - assumes f: "\i. f i \ borel_measurable M" "\i. AE x in M. 0 \ f i x" - shows "(\\<^sup>+ x. (\i. f i x) \M) = (\i. integral\<^sup>P M (f i))" -proof - - have all_pos: "AE x in M. \i. 0 \ f i x" - using assms by (auto simp: AE_all_countable) - have "(\i. integral\<^sup>P M (f i)) = (SUP n. \iP M (f i))" - using positive_integral_positive by (rule suminf_ereal_eq_SUP) - also have "\ = (SUP n. \\<^sup>+x. (\iM)" - unfolding positive_integral_setsum[OF f] .. - also have "\ = \\<^sup>+x. (SUP n. \iM" using f all_pos - by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) - (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3) - also have "\ = \\<^sup>+x. (\i. f i x) \M" using all_pos - by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUP) - finally show ?thesis by simp -qed - -text {* Fatou's lemma: convergence theorem on limes inferior *} -lemma positive_integral_lim_INF: - fixes u :: "nat \ 'a \ ereal" - assumes u: "\i. u i \ borel_measurable M" "\i. AE x in M. 0 \ u i x" - shows "(\\<^sup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>P M (u n))" -proof - - have pos: "AE x in M. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) - have "(\\<^sup>+ x. liminf (\n. u n x) \M) = - (SUP n. \\<^sup>+ x. (INF i:{n..}. u i x) \M)" - unfolding liminf_SUP_INF using pos u - by (intro positive_integral_monotone_convergence_SUP_AE) - (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) - also have "\ \ liminf (\n. integral\<^sup>P M (u n))" - unfolding liminf_SUP_INF - by (auto intro!: SUP_mono exI INF_greatest positive_integral_mono INF_lower) - finally show ?thesis . -qed - -lemma positive_integral_null_set: - assumes "N \ null_sets M" shows "(\\<^sup>+ x. u x * indicator N x \M) = 0" -proof - - have "(\\<^sup>+ x. u x * indicator N x \M) = (\\<^sup>+ x. 0 \M)" - proof (intro positive_integral_cong_AE AE_I) - show "{x \ space M. u x * indicator N x \ 0} \ N" - by (auto simp: indicator_def) - show "(emeasure M) N = 0" "N \ sets M" - using assms by auto - qed - then show ?thesis by simp -qed - -lemma positive_integral_0_iff: - assumes u: "u \ borel_measurable M" and pos: "AE x in M. 0 \ u x" - shows "integral\<^sup>P M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" - (is "_ \ (emeasure M) ?A = 0") -proof - - have u_eq: "(\\<^sup>+ x. u x * indicator ?A x \M) = integral\<^sup>P M u" - by (auto intro!: positive_integral_cong simp: indicator_def) - show ?thesis - proof - assume "(emeasure M) ?A = 0" - with positive_integral_null_set[of ?A M u] u - show "integral\<^sup>P M u = 0" by (simp add: u_eq null_sets_def) - next - { fix r :: ereal and n :: nat assume gt_1: "1 \ real n * r" - then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) - then have "0 \ r" by (auto simp add: ereal_zero_less_0_iff) } - note gt_1 = this - assume *: "integral\<^sup>P M u = 0" - let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" - have "0 = (SUP n. (emeasure M) (?M n \ ?A))" - proof - - { fix n :: nat - from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] - have "(emeasure M) (?M n \ ?A) \ 0" unfolding u_eq * using u by simp - moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto - ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } - thus ?thesis by simp - qed - also have "\ = (emeasure M) (\n. ?M n \ ?A)" - proof (safe intro!: SUP_emeasure_incseq) - fix n show "?M n \ ?A \ sets M" - using u by (auto intro!: sets.Int) - next - show "incseq (\n. {x \ space M. 1 \ real n * u x} \ {x \ space M. u x \ 0})" - proof (safe intro!: incseq_SucI) - fix n :: nat and x - assume *: "1 \ real n * u x" - also from gt_1[OF *] have "real n * u x \ real (Suc n) * u x" - using `0 \ u x` by (auto intro!: ereal_mult_right_mono) - finally show "1 \ real (Suc n) * u x" by auto - qed - qed - also have "\ = (emeasure M) {x\space M. 0 < u x}" - proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1) - fix x assume "0 < u x" and [simp, intro]: "x \ space M" - show "x \ (\n. ?M n \ ?A)" - proof (cases "u x") - case (real r) with `0 < u x` have "0 < r" by auto - obtain j :: nat where "1 / r \ real j" using real_arch_simple .. - hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using `0 < r` by auto - hence "1 \ real j * r" using real `0 < r` by auto - thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) - qed (insert `0 < u x`, auto) - qed auto - finally have "(emeasure M) {x\space M. 0 < u x} = 0" by simp - moreover - from pos have "AE x in M. \ (u x < 0)" by auto - then have "(emeasure M) {x\space M. u x < 0} = 0" - using AE_iff_null[of M] u by auto - moreover have "(emeasure M) {x\space M. u x \ 0} = (emeasure M) {x\space M. u x < 0} + (emeasure M) {x\space M. 0 < u x}" - using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) - ultimately show "(emeasure M) ?A = 0" by simp - qed -qed - -lemma positive_integral_0_iff_AE: - assumes u: "u \ borel_measurable M" - shows "integral\<^sup>P M u = 0 \ (AE x in M. u x \ 0)" -proof - - have sets: "{x\space M. max 0 (u x) \ 0} \ sets M" - using u by auto - from positive_integral_0_iff[of "\x. max 0 (u x)"] - have "integral\<^sup>P M u = 0 \ (AE x in M. max 0 (u x) = 0)" - unfolding positive_integral_max_0 - using AE_iff_null[OF sets] u by auto - also have "\ \ (AE x in M. u x \ 0)" by (auto split: split_max) - finally show ?thesis . -qed - -lemma AE_iff_positive_integral: - "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>P M (indicator {x. \ P x}) = 0" - by (subst positive_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def - sets.sets_Collect_neg indicator_def[abs_def] measurable_If) - -lemma positive_integral_const_If: - "(\\<^sup>+x. a \M) = (if 0 \ a then a * (emeasure M) (space M) else 0)" - by (auto intro!: positive_integral_0_iff_AE[THEN iffD2]) - -lemma positive_integral_subalgebra: - assumes f: "f \ borel_measurable N" "\x. 0 \ f x" - and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" - shows "integral\<^sup>P N f = integral\<^sup>P M f" -proof - - have [simp]: "\f :: 'a \ ereal. f \ borel_measurable N \ f \ borel_measurable M" - using N by (auto simp: measurable_def) - have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" - using N by (auto simp add: eventually_ae_filter null_sets_def) - have [simp]: "\A. A \ sets N \ A \ sets M" - using N by auto - from f show ?thesis - apply induct - apply (simp_all add: positive_integral_add positive_integral_cmult positive_integral_monotone_convergence_SUP N) - apply (auto intro!: positive_integral_cong cong: positive_integral_cong simp: N(2)[symmetric]) - done -qed - -lemma positive_integral_nat_function: - fixes f :: "'a \ nat" - assumes "f \ measurable M (count_space UNIV)" - shows "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\t. emeasure M {x\space M. t < f x})" -proof - - def F \ "\i. {x\space M. i < f x}" - with assms have [measurable]: "\i. F i \ sets M" - by auto - - { fix x assume "x \ space M" - have "(\i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" - using sums_If_finite[of "\i. i < f x" "\_. 1::real"] by simp - then have "(\i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))" - unfolding sums_ereal . - moreover have "\i. ereal (if i < f x then 1 else 0) = indicator (F i) x" - using `x \ space M` by (simp add: one_ereal_def F_def) - ultimately have "ereal(of_nat(f x)) = (\i. indicator (F i) x)" - by (simp add: sums_iff) } - then have "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" - by (simp cong: positive_integral_cong) - also have "\ = (\i. emeasure M (F i))" - by (simp add: positive_integral_suminf) - finally show ?thesis - by (simp add: F_def) -qed - -section "Lebesgue Integral" - -definition integrable :: "'a measure \ ('a \ real) \ bool" where - "integrable M f \ f \ borel_measurable M \ - (\\<^sup>+ x. ereal (f x) \M) \ \ \ (\\<^sup>+ x. ereal (- f x) \M) \ \" - -lemma borel_measurable_integrable[measurable_dest]: - "integrable M f \ f \ borel_measurable M" - by (auto simp: integrable_def) - -lemma integrableD[dest]: - assumes "integrable M f" - shows "f \ borel_measurable M" "(\\<^sup>+ x. ereal (f x) \M) \ \" "(\\<^sup>+ x. ereal (- f x) \M) \ \" - using assms unfolding integrable_def by auto - -definition lebesgue_integral :: "'a measure \ ('a \ real) \ real" ("integral\<^sup>L") where - "integral\<^sup>L M f = real ((\\<^sup>+ x. ereal (f x) \M)) - real ((\\<^sup>+ x. ereal (- f x) \M))" - -syntax - "_lebesgue_integral" :: "pttrn \ real \ 'a measure \ real" ("\ _. _ \_" [60,61] 110) - -translations - "\ x. f \M" == "CONST lebesgue_integral M (%x. f)" - -lemma integrableE: - assumes "integrable M f" - obtains r q where - "(\\<^sup>+x. ereal (f x)\M) = ereal r" - "(\\<^sup>+x. ereal (-f x)\M) = ereal q" - "f \ borel_measurable M" "integral\<^sup>L M f = r - q" - using assms unfolding integrable_def lebesgue_integral_def - using positive_integral_positive[of M "\x. ereal (f x)"] - using positive_integral_positive[of M "\x. ereal (-f x)"] - by (cases rule: ereal2_cases[of "(\\<^sup>+x. ereal (-f x)\M)" "(\\<^sup>+x. ereal (f x)\M)"]) auto - -lemma integral_cong: - assumes "\x. x \ space M \ f x = g x" - shows "integral\<^sup>L M f = integral\<^sup>L M g" - using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def) - -lemma integral_cong_AE: - assumes cong: "AE x in M. f x = g x" - shows "integral\<^sup>L M f = integral\<^sup>L M g" -proof - - have *: "AE x in M. ereal (f x) = ereal (g x)" - "AE x in M. ereal (- f x) = ereal (- g x)" using cong by auto - show ?thesis - unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def .. -qed - -lemma integrable_cong_AE: - assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" - assumes "AE x in M. f x = g x" - shows "integrable M f = integrable M g" -proof - - have "(\\<^sup>+ x. ereal (f x) \M) = (\\<^sup>+ x. ereal (g x) \M)" - "(\\<^sup>+ x. ereal (- f x) \M) = (\\<^sup>+ x. ereal (- g x) \M)" - using assms by (auto intro!: positive_integral_cong_AE) - with assms show ?thesis - by (auto simp: integrable_def) -qed - -lemma integrable_cong: - "(\x. x \ space M \ f x = g x) \ integrable M f \ integrable M g" - by (simp cong: positive_integral_cong measurable_cong add: integrable_def) - -lemma integral_mono_AE: - assumes fg: "integrable M f" "integrable M g" - and mono: "AE t in M. f t \ g t" - shows "integral\<^sup>L M f \ integral\<^sup>L M g" -proof - - have "AE x in M. ereal (f x) \ ereal (g x)" - using mono by auto - moreover have "AE x in M. ereal (- g x) \ ereal (- f x)" - using mono by auto - ultimately show ?thesis using fg - by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono - simp: positive_integral_positive lebesgue_integral_def algebra_simps) -qed - -lemma integral_mono: - assumes "integrable M f" "integrable M g" "\t. t \ space M \ f t \ g t" - shows "integral\<^sup>L M f \ integral\<^sup>L M g" - using assms by (auto intro: integral_mono_AE) - -lemma positive_integral_eq_integral: - assumes f: "integrable M f" - assumes nonneg: "AE x in M. 0 \ f x" - shows "(\\<^sup>+ x. ereal (f x) \M) = integral\<^sup>L M f" -proof - - have "(\\<^sup>+ x. max 0 (ereal (- f x)) \M) = (\\<^sup>+ x. 0 \M)" - using nonneg by (intro positive_integral_cong_AE) (auto split: split_max) - with f positive_integral_positive show ?thesis - by (cases "\\<^sup>+ x. ereal (f x) \M") - (auto simp add: lebesgue_integral_def positive_integral_max_0 integrable_def) -qed - -lemma integral_eq_positive_integral: - assumes f: "\x. 0 \ f x" - shows "integral\<^sup>L M f = real (\\<^sup>+ x. ereal (f x) \M)" -proof - - { fix x have "max 0 (ereal (- f x)) = 0" using f[of x] by (simp split: split_max) } - then have "0 = (\\<^sup>+ x. max 0 (ereal (- f x)) \M)" by simp - also have "\ = (\\<^sup>+ x. ereal (- f x) \M)" unfolding positive_integral_max_0 .. - finally show ?thesis - unfolding lebesgue_integral_def by simp -qed - -lemma integral_minus[intro, simp]: - assumes "integrable M f" - shows "integrable M (\x. - f x)" "(\x. - f x \M) = - integral\<^sup>L M f" - using assms by (auto simp: integrable_def lebesgue_integral_def) - -lemma integral_minus_iff[simp]: - "integrable M (\x. - f x) \ integrable M f" -proof - assume "integrable M (\x. - f x)" - then have "integrable M (\x. - (- f x))" - by (rule integral_minus) - then show "integrable M f" by simp -qed (rule integral_minus) - -lemma integral_of_positive_diff: - assumes integrable: "integrable M u" "integrable M v" - and f_def: "\x. f x = u x - v x" and pos: "\x. 0 \ u x" "\x. 0 \ v x" - shows "integrable M f" and "integral\<^sup>L M f = integral\<^sup>L M u - integral\<^sup>L M v" -proof - - let ?f = "\x. max 0 (ereal (f x))" - let ?mf = "\x. max 0 (ereal (- f x))" - let ?u = "\x. max 0 (ereal (u x))" - let ?v = "\x. max 0 (ereal (v x))" - - from borel_measurable_diff[of u M v] integrable - have f_borel: "?f \ borel_measurable M" and - mf_borel: "?mf \ borel_measurable M" and - v_borel: "?v \ borel_measurable M" and - u_borel: "?u \ borel_measurable M" and - "f \ borel_measurable M" - by (auto simp: f_def[symmetric] integrable_def) - - have "(\\<^sup>+ x. ereal (u x - v x) \M) \ integral\<^sup>P M ?u" - using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0) - moreover have "(\\<^sup>+ x. ereal (v x - u x) \M) \ integral\<^sup>P M ?v" - using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0) - ultimately show f: "integrable M f" - using `integrable M u` `integrable M v` `f \ borel_measurable M` - by (auto simp: integrable_def f_def positive_integral_max_0) - - have "\x. ?u x + ?mf x = ?v x + ?f x" - unfolding f_def using pos by (simp split: split_max) - then have "(\\<^sup>+ x. ?u x + ?mf x \M) = (\\<^sup>+ x. ?v x + ?f x \M)" by simp - then have "real (integral\<^sup>P M ?u + integral\<^sup>P M ?mf) = - real (integral\<^sup>P M ?v + integral\<^sup>P M ?f)" - using positive_integral_add[OF u_borel _ mf_borel] - using positive_integral_add[OF v_borel _ f_borel] - by auto - then show "integral\<^sup>L M f = integral\<^sup>L M u - integral\<^sup>L M v" - unfolding positive_integral_max_0 - unfolding pos[THEN integral_eq_positive_integral] - using integrable f by (auto elim!: integrableE) -qed - -lemma integral_linear: - assumes "integrable M f" "integrable M g" and "0 \ a" - shows "integrable M (\t. a * f t + g t)" - and "(\ t. a * f t + g t \M) = a * integral\<^sup>L M f + integral\<^sup>L M g" (is ?EQ) -proof - - let ?f = "\x. max 0 (ereal (f x))" - let ?g = "\x. max 0 (ereal (g x))" - let ?mf = "\x. max 0 (ereal (- f x))" - let ?mg = "\x. max 0 (ereal (- g x))" - let ?p = "\t. max 0 (a * f t) + max 0 (g t)" - let ?n = "\t. max 0 (- (a * f t)) + max 0 (- g t)" - - from assms have linear: - "(\\<^sup>+ x. ereal a * ?f x + ?g x \M) = ereal a * integral\<^sup>P M ?f + integral\<^sup>P M ?g" - "(\\<^sup>+ x. ereal a * ?mf x + ?mg x \M) = ereal a * integral\<^sup>P M ?mf + integral\<^sup>P M ?mg" - by (auto intro!: positive_integral_linear simp: integrable_def) - - have *: "(\\<^sup>+x. ereal (- ?p x) \M) = 0" "(\\<^sup>+x. ereal (- ?n x) \M) = 0" - using `0 \ a` assms by (auto simp: positive_integral_0_iff_AE integrable_def) - have **: "\x. ereal a * ?f x + ?g x = max 0 (ereal (?p x))" - "\x. ereal a * ?mf x + ?mg x = max 0 (ereal (?n x))" - using `0 \ a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff) - - have "integrable M ?p" "integrable M ?n" - "\t. a * f t + g t = ?p t - ?n t" "\t. 0 \ ?p t" "\t. 0 \ ?n t" - using linear assms unfolding integrable_def ** * - by (auto simp: positive_integral_max_0) - note diff = integral_of_positive_diff[OF this] - - show "integrable M (\t. a * f t + g t)" by (rule diff) - from assms linear show ?EQ - unfolding diff(2) ** positive_integral_max_0 - unfolding lebesgue_integral_def * - by (auto elim!: integrableE simp: field_simps) -qed - -lemma integral_add[simp, intro]: - assumes "integrable M f" "integrable M g" - shows "integrable M (\t. f t + g t)" - and "(\ t. f t + g t \M) = integral\<^sup>L M f + integral\<^sup>L M g" - using assms integral_linear[where a=1] by auto - -lemma integral_zero[simp, intro]: - shows "integrable M (\x. 0)" "(\ x.0 \M) = 0" - unfolding integrable_def lebesgue_integral_def - by auto - -lemma lebesgue_integral_uminus: - "(\x. - f x \M) = - integral\<^sup>L M f" - unfolding lebesgue_integral_def by simp - -lemma lebesgue_integral_cmult_nonneg: - assumes f: "f \ borel_measurable M" and "0 \ c" - shows "(\x. c * f x \M) = c * integral\<^sup>L M f" -proof - - { have "real (ereal c * integral\<^sup>P M (\x. max 0 (ereal (f x)))) = - real (integral\<^sup>P M (\x. ereal c * max 0 (ereal (f x))))" - using f `0 \ c` by (subst positive_integral_cmult) auto - also have "\ = real (integral\<^sup>P M (\x. max 0 (ereal (c * f x))))" - using `0 \ c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def zero_le_mult_iff) - finally have "real (integral\<^sup>P M (\x. ereal (c * f x))) = c * real (integral\<^sup>P M (\x. ereal (f x)))" - by (simp add: positive_integral_max_0) } - moreover - { have "real (ereal c * integral\<^sup>P M (\x. max 0 (ereal (- f x)))) = - real (integral\<^sup>P M (\x. ereal c * max 0 (ereal (- f x))))" - using f `0 \ c` by (subst positive_integral_cmult) auto - also have "\ = real (integral\<^sup>P M (\x. max 0 (ereal (- c * f x))))" - using `0 \ c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def mult_le_0_iff) - finally have "real (integral\<^sup>P M (\x. ereal (- c * f x))) = c * real (integral\<^sup>P M (\x. ereal (- f x)))" - by (simp add: positive_integral_max_0) } - ultimately show ?thesis - by (simp add: lebesgue_integral_def field_simps) -qed - -lemma lebesgue_integral_cmult: - assumes f: "f \ borel_measurable M" - shows "(\x. c * f x \M) = c * integral\<^sup>L M f" -proof (cases rule: linorder_le_cases) - assume "0 \ c" with f show ?thesis by (rule lebesgue_integral_cmult_nonneg) -next - assume "c \ 0" - with lebesgue_integral_cmult_nonneg[OF f, of "-c"] - show ?thesis - by (simp add: lebesgue_integral_def) -qed - -lemma lebesgue_integral_multc: - "f \ borel_measurable M \ (\x. f x * c \M) = integral\<^sup>L M f * c" - using lebesgue_integral_cmult[of f M c] by (simp add: ac_simps) - -lemma integral_multc: - "integrable M f \ (\ x. f x * c \M) = integral\<^sup>L M f * c" - by (simp add: lebesgue_integral_multc) - -lemma integral_cmult[simp, intro]: - assumes "integrable M f" - shows "integrable M (\t. a * f t)" (is ?P) - and "(\ t. a * f t \M) = a * integral\<^sup>L M f" (is ?I) -proof - - have "integrable M (\t. a * f t) \ (\ t. a * f t \M) = a * integral\<^sup>L M f" - proof (cases rule: le_cases) - assume "0 \ a" show ?thesis - using integral_linear[OF assms integral_zero(1) `0 \ a`] - by simp - next - assume "a \ 0" hence "0 \ - a" by auto - have *: "\t. - a * t + 0 = (-a) * t" by simp - show ?thesis using integral_linear[OF assms integral_zero(1) `0 \ - a`] - integral_minus(1)[of M "\t. - a * f t"] - unfolding * integral_zero by simp - qed - thus ?P ?I by auto -qed - -lemma integral_diff[simp, intro]: - assumes f: "integrable M f" and g: "integrable M g" - shows "integrable M (\t. f t - g t)" - and "(\ t. f t - g t \M) = integral\<^sup>L M f - integral\<^sup>L M g" - using integral_add[OF f integral_minus(1)[OF g]] - unfolding integral_minus(2)[OF g] - by auto - -lemma integral_indicator[simp, intro]: - assumes "A \ sets M" and "(emeasure M) A \ \" - shows "integral\<^sup>L M (indicator A) = real (emeasure M A)" (is ?int) - and "integrable M (indicator A)" (is ?able) -proof - - from `A \ sets M` have *: - "\x. ereal (indicator A x) = indicator A x" - "(\\<^sup>+x. ereal (- indicator A x) \M) = 0" - by (auto split: split_indicator simp: positive_integral_0_iff_AE one_ereal_def) - show ?int ?able - using assms unfolding lebesgue_integral_def integrable_def - by (auto simp: *) -qed - -lemma integral_cmul_indicator: - assumes "A \ sets M" and "c \ 0 \ (emeasure M) A \ \" - shows "integrable M (\x. c * indicator A x)" (is ?P) - and "(\x. c * indicator A x \M) = c * real ((emeasure M) A)" (is ?I) -proof - - show ?P - proof (cases "c = 0") - case False with assms show ?thesis by simp - qed simp - - show ?I - proof (cases "c = 0") - case False with assms show ?thesis by simp - qed simp -qed - -lemma integral_setsum[simp, intro]: - assumes "\n. n \ S \ integrable M (f n)" - shows "(\x. (\ i \ S. f i x) \M) = (\ i \ S. integral\<^sup>L M (f i))" (is "?int S") - and "integrable M (\x. \ i \ S. f i x)" (is "?I S") -proof - - have "?int S \ ?I S" - proof (cases "finite S") - assume "finite S" - from this assms show ?thesis by (induct S) simp_all - qed simp - thus "?int S" and "?I S" by auto -qed - -lemma integrable_bound: - assumes "integrable M f" and f: "AE x in M. \g x\ \ f x" - assumes borel: "g \ borel_measurable M" - shows "integrable M g" -proof - - have "(\\<^sup>+ x. ereal (g x) \M) \ (\\<^sup>+ x. ereal \g x\ \M)" - by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^sup>+ x. ereal (f x) \M)" - using f by (auto intro!: positive_integral_mono_AE) - also have "\ < \" - using `integrable M f` unfolding integrable_def by auto - finally have pos: "(\\<^sup>+ x. ereal (g x) \M) < \" . - - have "(\\<^sup>+ x. ereal (- g x) \M) \ (\\<^sup>+ x. ereal (\g x\) \M)" - by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^sup>+ x. ereal (f x) \M)" - using f by (auto intro!: positive_integral_mono_AE) - also have "\ < \" - using `integrable M f` unfolding integrable_def by auto - finally have neg: "(\\<^sup>+ x. ereal (- g x) \M) < \" . - - from neg pos borel show ?thesis - unfolding integrable_def by auto -qed - -lemma integrable_abs: - assumes f[measurable]: "integrable M f" - shows "integrable M (\ x. \f x\)" -proof - - from assms have *: "(\\<^sup>+x. ereal (- \f x\)\M) = 0" - "\x. ereal \f x\ = max 0 (ereal (f x)) + max 0 (ereal (- f x))" - by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max) - with assms show ?thesis - by (simp add: positive_integral_add positive_integral_max_0 integrable_def) -qed - -lemma integral_subalgebra: - assumes borel: "f \ borel_measurable N" - and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" - shows "integrable N f \ integrable M f" (is ?P) - and "integral\<^sup>L N f = integral\<^sup>L M f" (is ?I) -proof - - have "(\\<^sup>+ x. max 0 (ereal (f x)) \N) = (\\<^sup>+ x. max 0 (ereal (f x)) \M)" - "(\\<^sup>+ x. max 0 (ereal (- f x)) \N) = (\\<^sup>+ x. max 0 (ereal (- f x)) \M)" - using borel by (auto intro!: positive_integral_subalgebra N) - moreover have "f \ borel_measurable M \ f \ borel_measurable N" - using assms unfolding measurable_def by auto - ultimately show ?P ?I - by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0) -qed - -lemma lebesgue_integral_nonneg: - assumes ae: "(AE x in M. 0 \ f x)" shows "0 \ integral\<^sup>L M f" -proof - - have "(\\<^sup>+x. max 0 (ereal (- f x)) \M) = (\\<^sup>+x. 0 \M)" - using ae by (intro positive_integral_cong_AE) (auto simp: max_def) - then show ?thesis - by (auto simp: lebesgue_integral_def positive_integral_max_0 - intro!: real_of_ereal_pos positive_integral_positive) -qed - -lemma integrable_abs_iff: - "f \ borel_measurable M \ integrable M (\ x. \f x\) \ integrable M f" - by (auto intro!: integrable_bound[where g=f] integrable_abs) - -lemma integrable_max: - assumes int: "integrable M f" "integrable M g" - shows "integrable M (\ x. max (f x) (g x))" -proof (rule integrable_bound) - show "integrable M (\x. \f x\ + \g x\)" - using int by (simp add: integrable_abs) - show "(\x. max (f x) (g x)) \ borel_measurable M" - using int unfolding integrable_def by auto -qed auto - -lemma integrable_min: - assumes int: "integrable M f" "integrable M g" - shows "integrable M (\ x. min (f x) (g x))" -proof (rule integrable_bound) - show "integrable M (\x. \f x\ + \g x\)" - using int by (simp add: integrable_abs) - show "(\x. min (f x) (g x)) \ borel_measurable M" - using int unfolding integrable_def by auto -qed auto - -lemma integral_triangle_inequality: - assumes "integrable M f" - shows "\integral\<^sup>L M f\ \ (\x. \f x\ \M)" -proof - - have "\integral\<^sup>L M f\ = max (integral\<^sup>L M f) (- integral\<^sup>L M f)" by auto - also have "\ \ (\x. \f x\ \M)" - using assms integral_minus(2)[of M f, symmetric] - by (auto intro!: integral_mono integrable_abs simp del: integral_minus) - finally show ?thesis . -qed - -lemma integrable_nonneg: - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" "(\\<^sup>+ x. f x \M) \ \" - shows "integrable M f" - unfolding integrable_def -proof (intro conjI f) - have "(\\<^sup>+ x. ereal (- f x) \M) = 0" - using f by (subst positive_integral_0_iff_AE) auto - then show "(\\<^sup>+ x. ereal (- f x) \M) \ \" by simp -qed - -lemma integral_positive: - assumes "integrable M f" "\x. x \ space M \ 0 \ f x" - shows "0 \ integral\<^sup>L M f" -proof - - have "0 = (\x. 0 \M)" by auto - also have "\ \ integral\<^sup>L M f" - using assms by (rule integral_mono[OF integral_zero(1)]) - finally show ?thesis . -qed - -lemma integral_monotone_convergence_pos: - assumes i: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" - and pos: "\i. AE x in M. 0 \ f i x" - and lim: "AE x in M. (\i. f i x) ----> u x" - and ilim: "(\i. integral\<^sup>L M (f i)) ----> x" - and u: "u \ borel_measurable M" - shows "integrable M u" - and "integral\<^sup>L M u = x" -proof - - have "(\\<^sup>+ x. ereal (u x) \M) = (SUP n. (\\<^sup>+ x. ereal (f n x) \M))" - proof (subst positive_integral_monotone_convergence_SUP_AE[symmetric]) - fix i - from mono pos show "AE x in M. ereal (f i x) \ ereal (f (Suc i) x) \ 0 \ ereal (f i x)" - by eventually_elim (auto simp: mono_def) - show "(\x. ereal (f i x)) \ borel_measurable M" - using i by auto - next - show "(\\<^sup>+ x. ereal (u x) \M) = \\<^sup>+ x. (SUP i. ereal (f i x)) \M" - apply (rule positive_integral_cong_AE) - using lim mono - by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2]) - qed - also have "\ = ereal x" - using mono i unfolding positive_integral_eq_integral[OF i pos] - by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim) - finally have "(\\<^sup>+ x. ereal (u x) \M) = ereal x" . - moreover have "(\\<^sup>+ x. ereal (- u x) \M) = 0" - proof (subst positive_integral_0_iff_AE) - show "(\x. ereal (- u x)) \ borel_measurable M" - using u by auto - from mono pos[of 0] lim show "AE x in M. ereal (- u x) \ 0" - proof eventually_elim - fix x assume "mono (\n. f n x)" "0 \ f 0 x" "(\i. f i x) ----> u x" - then show "ereal (- u x) \ 0" - using incseq_le[of "\n. f n x" "u x" 0] by (simp add: mono_def incseq_def) - qed - qed - ultimately show "integrable M u" "integral\<^sup>L M u = x" - by (auto simp: integrable_def lebesgue_integral_def u) -qed - -lemma integral_monotone_convergence: - assumes f: "\i. integrable M (f i)" and mono: "AE x in M. mono (\n. f n x)" - and lim: "AE x in M. (\i. f i x) ----> u x" - and ilim: "(\i. integral\<^sup>L M (f i)) ----> x" - and u: "u \ borel_measurable M" - shows "integrable M u" - and "integral\<^sup>L M u = x" -proof - - have 1: "\i. integrable M (\x. f i x - f 0 x)" - using f by auto - have 2: "AE x in M. mono (\n. f n x - f 0 x)" - using mono by (auto simp: mono_def le_fun_def) - have 3: "\n. AE x in M. 0 \ f n x - f 0 x" - using mono by (auto simp: field_simps mono_def le_fun_def) - have 4: "AE x in M. (\i. f i x - f 0 x) ----> u x - f 0 x" - using lim by (auto intro!: tendsto_diff) - have 5: "(\i. (\x. f i x - f 0 x \M)) ----> x - integral\<^sup>L M (f 0)" - using f ilim by (auto intro!: tendsto_diff) - have 6: "(\x. u x - f 0 x) \ borel_measurable M" - using f[of 0] u by auto - note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5 6] - have "integrable M (\x. (u x - f 0 x) + f 0 x)" - using diff(1) f by (rule integral_add(1)) - with diff(2) f show "integrable M u" "integral\<^sup>L M u = x" - by auto -qed - -lemma integral_0_iff: - assumes "integrable M f" - shows "(\x. \f x\ \M) = 0 \ (emeasure M) {x\space M. f x \ 0} = 0" -proof - - have *: "(\\<^sup>+x. ereal (- \f x\) \M) = 0" - using assms by (auto simp: positive_integral_0_iff_AE integrable_def) - have "integrable M (\x. \f x\)" using assms by (rule integrable_abs) - hence "(\x. ereal (\f x\)) \ borel_measurable M" - "(\\<^sup>+ x. ereal \f x\ \M) \ \" unfolding integrable_def by auto - from positive_integral_0_iff[OF this(1)] this(2) - show ?thesis unfolding lebesgue_integral_def * - using positive_integral_positive[of M "\x. ereal \f x\"] - by (auto simp add: real_of_ereal_eq_0) -qed - -lemma integral_real: - "AE x in M. \f x\ \ \ \ (\x. real (f x) \M) = real (integral\<^sup>P M f) - real (integral\<^sup>P M (\x. - f x))" - using assms unfolding lebesgue_integral_def - by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real) - -lemma (in finite_measure) lebesgue_integral_const[simp]: - shows "integrable M (\x. a)" - and "(\x. a \M) = a * measure M (space M)" -proof - - { fix a :: real assume "0 \ a" - then have "(\\<^sup>+ x. ereal a \M) = ereal a * (emeasure M) (space M)" - by (subst positive_integral_const) auto - moreover - from `0 \ a` have "(\\<^sup>+ x. ereal (-a) \M) = 0" - by (subst positive_integral_0_iff_AE) auto - ultimately have "integrable M (\x. a)" by (auto simp: integrable_def) } - note * = this - show "integrable M (\x. a)" - proof cases - assume "0 \ a" with * show ?thesis . - next - assume "\ 0 \ a" - then have "0 \ -a" by auto - from *[OF this] show ?thesis by simp - qed - show "(\x. a \M) = a * measure M (space M)" - by (simp add: lebesgue_integral_def positive_integral_const_If emeasure_eq_measure) -qed - -lemma (in finite_measure) integrable_const_bound: - assumes "AE x in M. \f x\ \ B" and "f \ borel_measurable M" shows "integrable M f" - by (auto intro: integrable_bound[where f="\x. B"] lebesgue_integral_const assms) - -lemma (in finite_measure) integral_less_AE: - assumes int: "integrable M X" "integrable M Y" - assumes A: "(emeasure M) A \ 0" "A \ sets M" "AE x in M. x \ A \ X x \ Y x" - assumes gt: "AE x in M. X x \ Y x" - shows "integral\<^sup>L M X < integral\<^sup>L M Y" -proof - - have "integral\<^sup>L M X \ integral\<^sup>L M Y" - using gt int by (intro integral_mono_AE) auto - moreover - have "integral\<^sup>L M X \ integral\<^sup>L M Y" - proof - assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y" - have "integral\<^sup>L M (\x. \Y x - X x\) = integral\<^sup>L M (\x. Y x - X x)" - using gt by (intro integral_cong_AE) auto - also have "\ = 0" - using eq int by simp - finally have "(emeasure M) {x \ space M. Y x - X x \ 0} = 0" - using int by (simp add: integral_0_iff) - moreover - have "(\\<^sup>+x. indicator A x \M) \ (\\<^sup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" - using A by (intro positive_integral_mono_AE) auto - then have "(emeasure M) A \ (emeasure M) {x \ space M. Y x - X x \ 0}" - using int A by (simp add: integrable_def) - ultimately have "emeasure M A = 0" - using emeasure_nonneg[of M A] by simp - with `(emeasure M) A \ 0` show False by auto - qed - ultimately show ?thesis by auto -qed - -lemma (in finite_measure) integral_less_AE_space: - assumes int: "integrable M X" "integrable M Y" - assumes gt: "AE x in M. X x < Y x" "(emeasure M) (space M) \ 0" - shows "integral\<^sup>L M X < integral\<^sup>L M Y" - using gt by (intro integral_less_AE[OF int, where A="space M"]) auto - -lemma integral_dominated_convergence: - assumes u[measurable]: "\i. integrable M (u i)" and bound: "\j. AE x in M. \u j x\ \ w x" - and w[measurable]: "integrable M w" - and u': "AE x in M. (\i. u i x) ----> u' x" - and [measurable]: "u' \ borel_measurable M" - shows "integrable M u'" - and "(\i. (\x. \u i x - u' x\ \M)) ----> 0" (is "?lim_diff") - and "(\i. integral\<^sup>L M (u i)) ----> integral\<^sup>L M u'" (is ?lim) -proof - - have all_bound: "AE x in M. \j. \u j x\ \ w x" - using bound by (auto simp: AE_all_countable) - with u' have u'_bound: "AE x in M. \u' x\ \ w x" - by eventually_elim (auto intro: LIMSEQ_le_const2 tendsto_rabs) - - from bound[of 0] have w_pos: "AE x in M. 0 \ w x" - by eventually_elim auto - - show "integrable M u'" - by (rule integrable_bound) fact+ - - let ?diff = "\n x. 2 * w x - \u n x - u' x\" - have diff: "\n. integrable M (\x. \u n x - u' x\)" - using w u `integrable M u'` by (auto intro!: integrable_abs) - - from u'_bound all_bound - have diff_less_2w: "AE x in M. \j. \u j x - u' x\ \ 2 * w x" - proof (eventually_elim, intro allI) - fix x j assume *: "\u' x\ \ w x" "\j. \u j x\ \ w x" - then have "\u j x - u' x\ \ \u j x\ + \u' x\" by auto - also have "\ \ w x + w x" - using * by (intro add_mono) auto - finally show "\u j x - u' x\ \ 2 * w x" by simp - qed - - have PI_diff: "\n. (\\<^sup>+ x. ereal (?diff n x) \M) = - (\\<^sup>+ x. ereal (2 * w x) \M) - (\\<^sup>+ x. ereal \u n x - u' x\ \M)" - using diff w diff_less_2w w_pos - by (subst positive_integral_diff[symmetric]) - (auto simp: integrable_def intro!: positive_integral_cong_AE) - - have "integrable M (\x. 2 * w x)" - using w by auto - hence I2w_fin: "(\\<^sup>+ x. ereal (2 * w x) \M) \ \" and - borel_2w: "(\x. ereal (2 * w x)) \ borel_measurable M" - unfolding integrable_def by auto - - have "limsup (\n. \\<^sup>+ x. ereal \u n x - u' x\ \M) = 0" (is "limsup ?f = 0") - proof cases - assume eq_0: "(\\<^sup>+ x. max 0 (ereal (2 * w x)) \M) = 0" (is "?wx = 0") - { fix n - have "?f n \ ?wx" (is "integral\<^sup>P M ?f' \ _") - using diff_less_2w unfolding positive_integral_max_0 - by (intro positive_integral_mono_AE) auto - then have "?f n = 0" - using positive_integral_positive[of M ?f'] eq_0 by auto } - then show ?thesis by (simp add: Limsup_const) - next - assume neq_0: "(\\<^sup>+ x. max 0 (ereal (2 * w x)) \M) \ 0" (is "?wx \ 0") - have "0 = limsup (\n. 0 :: ereal)" by (simp add: Limsup_const) - also have "\ \ limsup (\n. \\<^sup>+ x. ereal \u n x - u' x\ \M)" - by (simp add: Limsup_mono positive_integral_positive) - finally have pos: "0 \ limsup (\n. \\<^sup>+ x. ereal \u n x - u' x\ \M)" . - have "?wx = (\\<^sup>+ x. liminf (\n. max 0 (ereal (?diff n x))) \M)" - using u' - proof (intro positive_integral_cong_AE, eventually_elim) - fix x assume u': "(\i. u i x) ----> u' x" - show "max 0 (ereal (2 * w x)) = liminf (\n. max 0 (ereal (?diff n x)))" - unfolding ereal_max_0 - proof (rule lim_imp_Liminf[symmetric], unfold lim_ereal) - have "(\i. ?diff i x) ----> 2 * w x - \u' x - u' x\" - using u' by (safe intro!: tendsto_intros) - then show "(\i. max 0 (?diff i x)) ----> max 0 (2 * w x)" - by (auto intro!: tendsto_max) - qed (rule trivial_limit_sequentially) - qed - also have "\ \ liminf (\n. \\<^sup>+ x. max 0 (ereal (?diff n x)) \M)" - using w u unfolding integrable_def - by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF) - also have "\ = (\\<^sup>+ x. ereal (2 * w x) \M) - - limsup (\n. \\<^sup>+ x. ereal \u n x - u' x\ \M)" - unfolding PI_diff positive_integral_max_0 - using positive_integral_positive[of M "\x. ereal (2 * w x)"] - by (subst liminf_ereal_cminus) auto - finally show ?thesis - using neq_0 I2w_fin positive_integral_positive[of M "\x. ereal (2 * w x)"] pos - unfolding positive_integral_max_0 - by (cases rule: ereal2_cases[of "\\<^sup>+ x. ereal (2 * w x) \M" "limsup (\n. \\<^sup>+ x. ereal \u n x - u' x\ \M)"]) - auto - qed - - have "liminf ?f \ limsup ?f" - by (intro Liminf_le_Limsup trivial_limit_sequentially) - moreover - { have "0 = liminf (\n. 0 :: ereal)" by (simp add: Liminf_const) - also have "\ \ liminf ?f" - by (simp add: Liminf_mono positive_integral_positive) - finally have "0 \ liminf ?f" . } - ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0" - using `limsup ?f = 0` by auto - have "\n. (\\<^sup>+ x. ereal \u n x - u' x\ \M) = ereal (\x. \u n x - u' x\ \M)" - using diff positive_integral_positive[of M] - by (subst integral_eq_positive_integral[of _ M]) (auto simp: ereal_real integrable_def) - then show ?lim_diff - using Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq] - by simp - - show ?lim - proof (rule LIMSEQ_I) - fix r :: real assume "0 < r" - from LIMSEQ_D[OF `?lim_diff` this] - obtain N where N: "\n. N \ n \ (\x. \u n x - u' x\ \M) < r" - using diff by (auto simp: integral_positive) - - show "\N. \n\N. norm (integral\<^sup>L M (u n) - integral\<^sup>L M u') < r" - proof (safe intro!: exI[of _ N]) - fix n assume "N \ n" - have "\integral\<^sup>L M (u n) - integral\<^sup>L M u'\ = \(\x. u n x - u' x \M)\" - using u `integrable M u'` by auto - also have "\ \ (\x. \u n x - u' x\ \M)" using u `integrable M u'` - by (rule_tac integral_triangle_inequality) auto - also note N[OF `N \ n`] - finally show "norm (integral\<^sup>L M (u n) - integral\<^sup>L M u') < r" by simp - qed - qed -qed - -lemma integral_sums: - assumes integrable[measurable]: "\i. integrable M (f i)" - and summable: "\x. x \ space M \ summable (\i. \f i x\)" - and sums: "summable (\i. (\x. \f i x\ \M))" - shows "integrable M (\x. (\i. f i x))" (is "integrable M ?S") - and "(\i. integral\<^sup>L M (f i)) sums (\x. (\i. f i x) \M)" (is ?integral) -proof - - have "\x\space M. \w. (\i. \f i x\) sums w" - using summable unfolding summable_def by auto - from bchoice[OF this] - obtain w where w: "\x. x \ space M \ (\i. \f i x\) sums w x" by auto - then have w_borel: "w \ borel_measurable M" unfolding sums_def - by (rule borel_measurable_LIMSEQ) auto - - let ?w = "\y. if y \ space M then w y else 0" - - obtain x where abs_sum: "(\i. (\x. \f i x\ \M)) sums x" - using sums unfolding summable_def .. - - have 1: "\n. integrable M (\x. \ij. AE x in M. \\i \ ?w x" - using AE_space - proof eventually_elim - fix j x assume [simp]: "x \ space M" - have "\\i \ (\if i x\)" by (rule setsum_abs) - also have "\ \ w x" using w[of x] setsum_le_suminf[of "\i. \f i x\"] unfolding sums_iff by auto - finally show "\\i \ ?w x" by simp - qed - - have 3: "integrable M ?w" - proof (rule integral_monotone_convergence(1)) - let ?F = "\n y. (\if i y\)" - let ?w' = "\n y. if y \ space M then ?F n y else 0" - have "\n. integrable M (?F n)" - using integrable by (auto intro!: integrable_abs) - thus "\n. integrable M (?w' n)" by (simp cong: integrable_cong) - show "AE x in M. mono (\n. ?w' n x)" - by (auto simp: mono_def le_fun_def intro!: setsum_mono2) - show "AE x in M. (\n. ?w' n x) ----> ?w x" - using w by (simp_all add: tendsto_const sums_def) - have *: "\n. integral\<^sup>L M (?w' n) = (\i< n. (\x. \f i x\ \M))" - using integrable by (simp add: integrable_abs cong: integral_cong) - from abs_sum - show "(\i. integral\<^sup>L M (?w' i)) ----> x" unfolding * sums_def . - qed (simp add: w_borel measurable_If_set) - - from summable[THEN summable_rabs_cancel] - have 4: "AE x in M. (\n. \i (\i. f i x)" - by (auto intro: summable_LIMSEQ) - - note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 - borel_measurable_suminf[OF integrableD(1)[OF integrable]]] - - from int show "integrable M ?S" by simp - - show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF integrable] - using int(2) by simp -qed - -lemma integrable_mult_indicator: - "A \ sets M \ integrable M f \ integrable M (\x. f x * indicator A x)" - by (rule integrable_bound[where f="\x. \f x\"]) - (auto intro: integrable_abs split: split_indicator) - -lemma tendsto_integral_at_top: - fixes M :: "real measure" - assumes M: "sets M = sets borel" and f[measurable]: "integrable M f" - shows "((\y. \ x. f x * indicator {.. y} x \M) ---> \ x. f x \M) at_top" -proof - - have M_measure[simp]: "borel_measurable M = borel_measurable borel" - using M by (simp add: sets_eq_imp_space_eq measurable_def) - { fix f assume f: "integrable M f" "\x. 0 \ f x" - then have [measurable]: "f \ borel_measurable borel" - by (simp add: integrable_def) - have "((\y. \ x. f x * indicator {.. y} x \M) ---> \ x. f x \M) at_top" - proof (rule tendsto_at_topI_sequentially) - have "\j. AE x in M. \f x * indicator {.. j} x\ \ f x" - using f(2) by (intro AE_I2) (auto split: split_indicator) - have int: "\n. integrable M (\x. f x * indicator {.. n} x)" - by (rule integrable_mult_indicator) (auto simp: M f) - show "(\n. \ x. f x * indicator {..real n} x \M) ----> integral\<^sup>L M f" - proof (rule integral_dominated_convergence) - { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" - by (rule eventually_sequentiallyI[of "natceiling x"]) - (auto split: split_indicator simp: natceiling_le_eq) } - from filterlim_cong[OF refl refl this] - show "AE x in M. (\n. f x * indicator {..real n} x) ----> f x" - by (simp add: tendsto_const) - qed (fact+, simp) - show "mono (\y. \ x. f x * indicator {..y} x \M)" - by (intro monoI integral_mono int) (auto split: split_indicator intro: f) - qed } - note nonneg = this - let ?P = "\y. \ x. max 0 (f x) * indicator {..y} x \M" - let ?N = "\y. \ x. max 0 (- f x) * indicator {..y} x \M" - let ?p = "integral\<^sup>L M (\x. max 0 (f x))" - let ?n = "integral\<^sup>L M (\x. max 0 (- f x))" - have "(?P ---> ?p) at_top" "(?N ---> ?n) at_top" - by (auto intro!: nonneg integrable_max f) - note tendsto_diff[OF this] - also have "(\y. ?P y - ?N y) = (\y. \ x. f x * indicator {..y} x \M)" - by (subst integral_diff(2)[symmetric]) - (auto intro!: integrable_mult_indicator integrable_max f integral_cong ext - simp: M split: split_max) - also have "?p - ?n = integral\<^sup>L M f" - by (subst integral_diff(2)[symmetric]) - (auto intro!: integrable_max f integral_cong ext simp: M split: split_max) - finally show ?thesis . -qed - -lemma integral_monotone_convergence_at_top: - fixes M :: "real measure" - assumes M: "sets M = sets borel" - assumes nonneg: "AE x in M. 0 \ f x" - assumes borel: "f \ borel_measurable borel" - assumes int: "\y. integrable M (\x. f x * indicator {.. y} x)" - assumes conv: "((\y. \ x. f x * indicator {.. y} x \M) ---> x) at_top" - shows "integrable M f" "integral\<^sup>L M f = x" -proof - - from nonneg have "AE x in M. mono (\n::nat. f x * indicator {..real n} x)" - by (auto split: split_indicator intro!: monoI) - { fix x have "eventually (\n. f x * indicator {..real n} x = f x) sequentially" - by (rule eventually_sequentiallyI[of "natceiling x"]) - (auto split: split_indicator simp: natceiling_le_eq) } - from filterlim_cong[OF refl refl this] - have "AE x in M. (\i. f x * indicator {..real i} x) ----> f x" - by (simp add: tendsto_const) - have "(\i. \ x. f x * indicator {..real i} x \M) ----> x" - using conv filterlim_real_sequentially by (rule filterlim_compose) - have M_measure[simp]: "borel_measurable M = borel_measurable borel" - using M by (simp add: sets_eq_imp_space_eq measurable_def) - have "f \ borel_measurable M" - using borel by simp - show "integrable M f" - by (rule integral_monotone_convergence) fact+ - show "integral\<^sup>L M f = x" - by (rule integral_monotone_convergence) fact+ -qed - - -section "Lebesgue integration on countable spaces" - -lemma integral_on_countable: - assumes f: "f \ borel_measurable M" - and bij: "bij_betw enum S (f ` space M)" - and enum_zero: "enum ` (-S) \ {0}" - and fin: "\x. x \ 0 \ (emeasure M) (f -` {x} \ space M) \ \" - and abs_summable: "summable (\r. \enum r * real ((emeasure M) (f -` {enum r} \ space M))\)" - shows "integrable M f" - and "(\r. enum r * real ((emeasure M) (f -` {enum r} \ space M))) sums integral\<^sup>L M f" (is ?sums) -proof - - let ?A = "\r. f -` {enum r} \ space M" - let ?F = "\r x. enum r * indicator (?A r) x" - have enum_eq: "\r. enum r * real ((emeasure M) (?A r)) = integral\<^sup>L M (?F r)" - using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) - - { fix x assume "x \ space M" - hence "f x \ enum ` S" using bij unfolding bij_betw_def by auto - then obtain i where "i\S" "enum i = f x" by auto - have F: "\j. ?F j x = (if j = i then f x else 0)" - proof cases - fix j assume "j = i" - thus "?thesis j" using `x \ space M` `enum i = f x` by (simp add: indicator_def) - next - fix j assume "j \ i" - show "?thesis j" using bij `i \ S` `j \ i` `enum i = f x` enum_zero - by (cases "j \ S") (auto simp add: indicator_def bij_betw_def inj_on_def) - qed - hence F_abs: "\j. \if j = i then f x else 0\ = (if j = i then \f x\ else 0)" by auto - have "(\i. ?F i x) sums f x" - "(\i. \?F i x\) sums \f x\" - by (auto intro!: sums_single simp: F F_abs) } - note F_sums_f = this(1) and F_abs_sums_f = this(2) - - have int_f: "integral\<^sup>L M f = (\x. (\r. ?F r x) \M)" "integrable M f = integrable M (\x. \r. ?F r x)" - using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff) - - { fix r - have "(\x. \?F r x\ \M) = (\x. \enum r\ * indicator (?A r) x \M)" - by (auto simp: indicator_def intro!: integral_cong) - also have "\ = \enum r\ * real ((emeasure M) (?A r))" - using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) - finally have "(\x. \?F r x\ \M) = \enum r * real ((emeasure M) (?A r))\" - using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) } - note int_abs_F = this - - have 1: "\i. integrable M (\x. ?F i x)" - using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) - - have 2: "\x. x \ space M \ summable (\i. \?F i x\)" - using F_abs_sums_f unfolding sums_iff by auto - - from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] - show ?sums unfolding enum_eq int_f by simp - - from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] - show "integrable M f" unfolding int_f by simp -qed - -section {* Distributions *} - -lemma positive_integral_distr': - assumes T: "T \ measurable M M'" - and f: "f \ borel_measurable (distr M M' T)" "\x. 0 \ f x" - shows "integral\<^sup>P (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" - using f -proof induct - case (cong f g) - with T show ?case - apply (subst positive_integral_cong[of _ f g]) - apply simp - apply (subst positive_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) - apply (simp add: measurable_def Pi_iff) - apply simp - done -next - case (set A) - then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" - by (auto simp: indicator_def) - from set T show ?case - by (subst positive_integral_cong[OF eq]) - (auto simp add: emeasure_distr intro!: positive_integral_indicator[symmetric] measurable_sets) -qed (simp_all add: measurable_compose[OF T] T positive_integral_cmult positive_integral_add - positive_integral_monotone_convergence_SUP le_fun_def incseq_def) - -lemma positive_integral_distr: - "T \ measurable M M' \ f \ borel_measurable M' \ integral\<^sup>P (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" - by (subst (1 2) positive_integral_max_0[symmetric]) - (simp add: positive_integral_distr') - -lemma integral_distr: - "T \ measurable M M' \ f \ borel_measurable M' \ integral\<^sup>L (distr M M' T) f = (\ x. f (T x) \M)" - unfolding lebesgue_integral_def - by (subst (1 2) positive_integral_distr) auto - -lemma integrable_distr_eq: - "T \ measurable M M' \ f \ borel_measurable M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" - unfolding integrable_def - by (subst (1 2) positive_integral_distr) (auto simp: comp_def) - -lemma integrable_distr: - "T \ measurable M M' \ integrable (distr M M' T) f \ integrable M (\x. f (T x))" - by (subst integrable_distr_eq[symmetric]) auto - -section {* Lebesgue integration on @{const count_space} *} - -lemma simple_function_count_space[simp]: - "simple_function (count_space A) f \ finite (f ` A)" - unfolding simple_function_def by simp - -lemma positive_integral_count_space: - assumes A: "finite {a\A. 0 < f a}" - shows "integral\<^sup>P (count_space A) f = (\a|a\A \ 0 < f a. f a)" -proof - - have *: "(\\<^sup>+x. max 0 (f x) \count_space A) = - (\\<^sup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" - by (auto intro!: positive_integral_cong - simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less) - also have "\ = (\a|a\A \ 0 < f a. \\<^sup>+ x. f a * indicator {a} x \count_space A)" - by (subst positive_integral_setsum) - (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) - also have "\ = (\a|a\A \ 0 < f a. f a)" - by (auto intro!: setsum_cong simp: positive_integral_cmult_indicator one_ereal_def[symmetric]) - finally show ?thesis by (simp add: positive_integral_max_0) -qed - -lemma integrable_count_space: - "finite X \ integrable (count_space X) f" - by (auto simp: positive_integral_count_space integrable_def) - -lemma positive_integral_count_space_finite: - "finite A \ (\\<^sup>+x. f x \count_space A) = (\a\A. max 0 (f a))" - by (subst positive_integral_max_0[symmetric]) - (auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le) - -lemma lebesgue_integral_count_space_finite_support: - assumes f: "finite {a\A. f a \ 0}" shows "(\x. f x \count_space A) = (\a | a \ A \ f a \ 0. f a)" -proof - - have *: "\r::real. 0 < max 0 r \ 0 < r" "\x. max 0 (ereal x) = ereal (max 0 x)" - "\a. a \ A \ 0 < f a \ max 0 (f a) = f a" - "\a. a \ A \ f a < 0 \ max 0 (- f a) = - f a" - "{a \ A. f a \ 0} = {a \ A. 0 < f a} \ {a \ A. f a < 0}" - "({a \ A. 0 < f a} \ {a \ A. f a < 0}) = {}" - by (auto split: split_max) - have "finite {a \ A. 0 < f a}" "finite {a \ A. f a < 0}" - by (auto intro: finite_subset[OF _ f]) - then show ?thesis - unfolding lebesgue_integral_def - apply (subst (1 2) positive_integral_max_0[symmetric]) - apply (subst (1 2) positive_integral_count_space) - apply (auto simp add: * setsum_negf setsum_Un - simp del: ereal_max) - done -qed - -lemma lebesgue_integral_count_space_finite: - "finite A \ (\x. f x \count_space A) = (\a\A. f a)" - apply (auto intro!: setsum_mono_zero_left - simp: positive_integral_count_space_finite lebesgue_integral_def) - apply (subst (1 2) setsum_real_of_ereal[symmetric]) - apply (auto simp: max_def setsum_subtractf[symmetric] intro!: setsum_cong) - done - -lemma emeasure_UN_countable: - assumes sets: "\i. i \ I \ X i \ sets M" and I: "countable I" - assumes disj: "disjoint_family_on X I" - shows "emeasure M (UNION I X) = (\\<^sup>+i. emeasure M (X i) \count_space I)" -proof cases - assume "finite I" with sets disj show ?thesis - by (subst setsum_emeasure[symmetric]) - (auto intro!: setsum_cong simp add: max_def subset_eq positive_integral_count_space_finite emeasure_nonneg) -next - assume f: "\ finite I" - then have [intro]: "I \ {}" by auto - from from_nat_into_inj_infinite[OF I f] from_nat_into[OF this] disj - have disj2: "disjoint_family (\i. X (from_nat_into I i))" - unfolding disjoint_family_on_def by metis - - from f have "bij_betw (from_nat_into I) UNIV I" - using bij_betw_from_nat_into[OF I] by simp - then have "(\i\I. X i) = (\i. (X \ from_nat_into I) i)" - unfolding SUP_def image_comp [symmetric] by (simp add: bij_betw_def) - then have "emeasure M (UNION I X) = emeasure M (\i. X (from_nat_into I i))" - by simp - also have "\ = (\i. emeasure M (X (from_nat_into I i)))" - by (intro suminf_emeasure[symmetric] disj disj2) (auto intro!: sets from_nat_into[OF `I \ {}`]) - also have "\ = (\n. \\<^sup>+i. emeasure M (X i) * indicator {from_nat_into I n} i \count_space I)" - proof (intro arg_cong[where f=suminf] ext) - fix i - have eq: "{a \ I. 0 < emeasure M (X a) * indicator {from_nat_into I i} a} - = (if 0 < emeasure M (X (from_nat_into I i)) then {from_nat_into I i} else {})" - using ereal_0_less_1 - by (auto simp: ereal_zero_less_0_iff indicator_def from_nat_into `I \ {}` simp del: ereal_0_less_1) - have "(\\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I) = - (if 0 < emeasure M (X (from_nat_into I i)) then emeasure M (X (from_nat_into I i)) else 0)" - by (subst positive_integral_count_space) (simp_all add: eq) - also have "\ = emeasure M (X (from_nat_into I i))" - by (simp add: less_le emeasure_nonneg) - finally show "emeasure M (X (from_nat_into I i)) = - \\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I" .. - qed - also have "\ = (\\<^sup>+i. emeasure M (X i) \count_space I)" - apply (subst positive_integral_suminf[symmetric]) - apply (auto simp: emeasure_nonneg intro!: positive_integral_cong) - proof - - fix x assume "x \ I" - then have "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = (\i\{to_nat_on I x}. emeasure M (X x) * indicator {from_nat_into I i} x)" - by (intro suminf_finite) (auto simp: indicator_def I f) - also have "\ = emeasure M (X x)" - by (simp add: I f `x\I`) - finally show "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = emeasure M (X x)" . - qed - finally show ?thesis . -qed - -section {* Measures with Restricted Space *} - -lemma positive_integral_restrict_space: - assumes \: "\ \ sets M" and f: "f \ borel_measurable M" "\x. 0 \ f x" "\x. x \ space M - \ \ f x = 0" - shows "positive_integral (restrict_space M \) f = positive_integral M f" -using f proof (induct rule: borel_measurable_induct) - case (cong f g) then show ?case - using positive_integral_cong[of M f g] positive_integral_cong[of "restrict_space M \" f g] - sets.sets_into_space[OF `\ \ sets M`] - by (simp add: subset_eq space_restrict_space) -next - case (set A) - then have "A \ \" - unfolding indicator_eq_0_iff by (auto dest: sets.sets_into_space) - with set `\ \ sets M` sets.sets_into_space[OF `\ \ sets M`] show ?case - by (subst positive_integral_indicator') - (auto simp add: sets_restrict_space_iff space_restrict_space - emeasure_restrict_space Int_absorb2 - dest: sets.sets_into_space) -next - case (mult f c) then show ?case - by (cases "c = 0") (simp_all add: measurable_restrict_space1 \ positive_integral_cmult) -next - case (add f g) then show ?case - by (simp add: measurable_restrict_space1 \ positive_integral_add ereal_add_nonneg_eq_0_iff) -next - case (seq F) then show ?case - by (auto simp add: SUP_eq_iff measurable_restrict_space1 \ positive_integral_monotone_convergence_SUP) -qed - -section {* Measure spaces with an associated density *} - -definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where - "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" - -lemma - shows sets_density[simp]: "sets (density M f) = sets M" - and space_density[simp]: "space (density M f) = space M" - by (auto simp: density_def) - -(* FIXME: add conversion to simplify space, sets and measurable *) -lemma space_density_imp[measurable_dest]: - "\x M f. x \ space (density M f) \ x \ space M" by auto - -lemma - shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" - and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" - and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" - unfolding measurable_def simple_function_def by simp_all - -lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ - (AE x in M. f x = f' x) \ density M f = density M f'" - unfolding density_def by (auto intro!: measure_of_eq positive_integral_cong_AE sets.space_closed) - -lemma density_max_0: "density M f = density M (\x. max 0 (f x))" -proof - - have "\x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" - by (auto simp: indicator_def) - then show ?thesis - unfolding density_def by (simp add: positive_integral_max_0) -qed - -lemma density_ereal_max_0: "density M (\x. ereal (f x)) = density M (\x. ereal (max 0 (f x)))" - by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) - -lemma emeasure_density: - assumes f[measurable]: "f \ borel_measurable M" and A[measurable]: "A \ sets M" - shows "emeasure (density M f) A = (\\<^sup>+ x. f x * indicator A x \M)" - (is "_ = ?\ A") - unfolding density_def -proof (rule emeasure_measure_of_sigma) - show "sigma_algebra (space M) (sets M)" .. - show "positive (sets M) ?\" - using f by (auto simp: positive_def intro!: positive_integral_positive) - have \_eq: "?\ = (\A. \\<^sup>+ x. max 0 (f x) * indicator A x \M)" (is "?\ = ?\'") - apply (subst positive_integral_max_0[symmetric]) - apply (intro ext positive_integral_cong_AE AE_I2) - apply (auto simp: indicator_def) - done - show "countably_additive (sets M) ?\" - unfolding \_eq - proof (intro countably_additiveI) - fix A :: "nat \ 'a set" assume "range A \ sets M" - then have "\i. A i \ sets M" by auto - then have *: "\i. (\x. max 0 (f x) * indicator (A i) x) \ borel_measurable M" - by (auto simp: set_eq_iff) - assume disj: "disjoint_family A" - have "(\n. ?\' (A n)) = (\\<^sup>+ x. (\n. max 0 (f x) * indicator (A n) x) \M)" - using f * by (simp add: positive_integral_suminf) - also have "\ = (\\<^sup>+ x. max 0 (f x) * (\n. indicator (A n) x) \M)" using f - by (auto intro!: suminf_cmult_ereal positive_integral_cong_AE) - also have "\ = (\\<^sup>+ x. max 0 (f x) * indicator (\n. A n) x \M)" - unfolding suminf_indicator[OF disj] .. - finally show "(\n. ?\' (A n)) = ?\' (\x. A x)" by simp - qed -qed fact - -lemma null_sets_density_iff: - assumes f: "f \ borel_measurable M" - shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x \ 0)" -proof - - { assume "A \ sets M" - have eq: "(\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. max 0 (f x) * indicator A x \M)" - apply (subst positive_integral_max_0[symmetric]) - apply (intro positive_integral_cong) - apply (auto simp: indicator_def) - done - have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ - emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" - unfolding eq - using f `A \ sets M` - by (intro positive_integral_0_iff) auto - also have "\ \ (AE x in M. max 0 (f x) * indicator A x = 0)" - using f `A \ sets M` - by (intro AE_iff_measurable[OF _ refl, symmetric]) auto - also have "(AE x in M. max 0 (f x) * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" - by (auto simp add: indicator_def max_def split: split_if_asm) - finally have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } - with f show ?thesis - by (simp add: null_sets_def emeasure_density cong: conj_cong) -qed - -lemma AE_density: - assumes f: "f \ borel_measurable M" - shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" -proof - assume "AE x in density M f. P x" - with f obtain N where "{x \ space M. \ P x} \ N" "N \ sets M" and ae: "AE x in M. x \ N \ f x \ 0" - by (auto simp: eventually_ae_filter null_sets_density_iff) - then have "AE x in M. x \ N \ P x" by auto - with ae show "AE x in M. 0 < f x \ P x" - by (rule eventually_elim2) auto -next - fix N assume ae: "AE x in M. 0 < f x \ P x" - then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" - by (auto simp: eventually_ae_filter) - then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. \ 0 < f x}" - "N \ {x\space M. \ 0 < f x} \ sets M" and ae2: "AE x in M. x \ N" - using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in) - show "AE x in density M f. P x" - using ae2 - unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] - by (intro exI[of _ "N \ {x\space M. \ 0 < f x}"] conjI *) - (auto elim: eventually_elim2) -qed - -lemma positive_integral_density': - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" - assumes g: "g \ borel_measurable M" "\x. 0 \ g x" - shows "integral\<^sup>P (density M f) g = (\\<^sup>+ x. f x * g x \M)" -using g proof induct - case (cong u v) - then show ?case - apply (subst positive_integral_cong[OF cong(3)]) - apply (simp_all cong: positive_integral_cong) - done -next - case (set A) then show ?case - by (simp add: emeasure_density f) -next - case (mult u c) - moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) - ultimately show ?case - using f by (simp add: positive_integral_cmult) -next - case (add u v) - then have "\x. f x * (v x + u x) = f x * v x + f x * u x" - by (simp add: ereal_right_distrib) - with add f show ?case - by (auto simp add: positive_integral_add ereal_zero_le_0_iff intro!: positive_integral_add[symmetric]) -next - case (seq U) - from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" - by eventually_elim (simp add: SUP_ereal_cmult seq) - from seq f show ?case - apply (simp add: positive_integral_monotone_convergence_SUP) - apply (subst positive_integral_cong_AE[OF eq]) - apply (subst positive_integral_monotone_convergence_SUP_AE) - apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) - done -qed - -lemma positive_integral_density: - "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ - integral\<^sup>P (density M f) g' = (\\<^sup>+ x. f x * g' x \M)" - by (subst (1 2) positive_integral_max_0[symmetric]) - (auto intro!: positive_integral_cong_AE - simp: measurable_If max_def ereal_zero_le_0_iff positive_integral_density') - -lemma integral_density: - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" - and g: "g \ borel_measurable M" - shows "integral\<^sup>L (density M f) g = (\ x. f x * g x \M)" - and "integrable (density M f) g \ integrable M (\x. f x * g x)" - unfolding lebesgue_integral_def integrable_def using f g - by (auto simp: positive_integral_density) - -lemma emeasure_restricted: - assumes S: "S \ sets M" and X: "X \ sets M" - shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" -proof - - have "emeasure (density M (indicator S)) X = (\\<^sup>+x. indicator S x * indicator X x \M)" - using S X by (simp add: emeasure_density) - also have "\ = (\\<^sup>+x. indicator (S \ X) x \M)" - by (auto intro!: positive_integral_cong simp: indicator_def) - also have "\ = emeasure M (S \ X)" - using S X by (simp add: sets.Int) - finally show ?thesis . -qed - -lemma measure_restricted: - "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" - by (simp add: emeasure_restricted measure_def) - -lemma (in finite_measure) finite_measure_restricted: - "S \ sets M \ finite_measure (density M (indicator S))" - by default (simp add: emeasure_restricted) - -lemma emeasure_density_const: - "A \ sets M \ 0 \ c \ emeasure (density M (\_. c)) A = c * emeasure M A" - by (auto simp: positive_integral_cmult_indicator emeasure_density) - -lemma measure_density_const: - "A \ sets M \ 0 < c \ c \ \ \ measure (density M (\_. c)) A = real c * measure M A" - by (auto simp: emeasure_density_const measure_def) - -lemma density_density_eq: - "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. 0 \ f x \ - density (density M f) g = density M (\x. f x * g x)" - by (auto intro!: measure_eqI simp: emeasure_density positive_integral_density ac_simps) - -lemma distr_density_distr: - assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" - and inv: "\x\space M. T' (T x) = x" - assumes f: "f \ borel_measurable M'" - shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") -proof (rule measure_eqI) - fix A assume A: "A \ sets ?R" - { fix x assume "x \ space M" - with sets.sets_into_space[OF A] - have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ereal)" - using T inv by (auto simp: indicator_def measurable_space) } - with A T T' f show "emeasure ?R A = emeasure ?L A" - by (simp add: measurable_comp emeasure_density emeasure_distr - positive_integral_distr measurable_sets cong: positive_integral_cong) -qed simp - -lemma density_density_divide: - fixes f g :: "'a \ real" - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" - assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" - assumes ac: "AE x in M. f x = 0 \ g x = 0" - shows "density (density M f) (\x. g x / f x) = density M g" -proof - - have "density M g = density M (\x. f x * (g x / f x))" - using f g ac by (auto intro!: density_cong measurable_If) - then show ?thesis - using f g by (subst density_density_eq) auto -qed - -section {* Point measure *} - -definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where - "point_measure A f = density (count_space A) f" - -lemma - shows space_point_measure: "space (point_measure A f) = A" - and sets_point_measure: "sets (point_measure A f) = Pow A" - by (auto simp: point_measure_def) - -lemma measurable_point_measure_eq1[simp]: - "g \ measurable (point_measure A f) M \ g \ A \ space M" - unfolding point_measure_def by simp - -lemma measurable_point_measure_eq2_finite[simp]: - "finite A \ - g \ measurable M (point_measure A f) \ - (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" - unfolding point_measure_def by (simp add: measurable_count_space_eq2) - -lemma simple_function_point_measure[simp]: - "simple_function (point_measure A f) g \ finite (g ` A)" - by (simp add: point_measure_def) - -lemma emeasure_point_measure: - assumes A: "finite {a\X. 0 < f a}" "X \ A" - shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" -proof - - have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" - using `X \ A` by auto - with A show ?thesis - by (simp add: emeasure_density positive_integral_count_space ereal_zero_le_0_iff - point_measure_def indicator_def) -qed - -lemma emeasure_point_measure_finite: - "finite A \ (\i. i \ A \ 0 \ f i) \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" - by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) - -lemma emeasure_point_measure_finite2: - "X \ A \ finite X \ (\i. i \ X \ 0 \ f i) \ emeasure (point_measure A f) X = (\a\X. f a)" - by (subst emeasure_point_measure) - (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) - -lemma null_sets_point_measure_iff: - "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x \ 0)" - by (auto simp: AE_count_space null_sets_density_iff point_measure_def) - -lemma AE_point_measure: - "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" - unfolding point_measure_def - by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) - -lemma positive_integral_point_measure: - "finite {a\A. 0 < f a \ 0 < g a} \ - integral\<^sup>P (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" - unfolding point_measure_def - apply (subst density_max_0) - apply (subst positive_integral_density) - apply (simp_all add: AE_count_space positive_integral_density) - apply (subst positive_integral_count_space ) - apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff) - apply (rule finite_subset) - prefer 2 - apply assumption - apply auto - done - -lemma positive_integral_point_measure_finite: - "finite A \ (\a. a \ A \ 0 \ f a) \ (\a. a \ A \ 0 \ g a) \ - integral\<^sup>P (point_measure A f) g = (\a\A. f a * g a)" - by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le) - -lemma lebesgue_integral_point_measure_finite: - "finite A \ (\a. a \ A \ 0 \ f a) \ integral\<^sup>L (point_measure A f) g = (\a\A. f a * g a)" - by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def) - -lemma integrable_point_measure_finite: - "finite A \ integrable (point_measure A (\x. ereal (f x))) g" - unfolding point_measure_def - apply (subst density_ereal_max_0) - apply (subst integral_density) - apply (auto simp: AE_count_space integrable_count_space) - done - -section {* Uniform measure *} - -definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" - -lemma - shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M" - and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" - by (auto simp: uniform_measure_def) - -lemma emeasure_uniform_measure[simp]: - assumes A: "A \ sets M" and B: "B \ sets M" - shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" -proof - - from A B have "emeasure (uniform_measure M A) B = (\\<^sup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" - by (auto simp add: uniform_measure_def emeasure_density split: split_indicator - intro!: positive_integral_cong) - also have "\ = emeasure M (A \ B) / emeasure M A" - using A B - by (subst positive_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg) - finally show ?thesis . -qed - -lemma measure_uniform_measure[simp]: - assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" - shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" - using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A - by (cases "emeasure M A" "emeasure M (A \ B)" rule: ereal2_cases) (simp_all add: measure_def) - -section {* Uniform count measure *} - -definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" - -lemma - shows space_uniform_count_measure: "space (uniform_count_measure A) = A" - and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" - unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) - -lemma emeasure_uniform_count_measure: - "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" - by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) - -lemma measure_uniform_count_measure: - "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" - by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) - -end diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon May 19 12:04:45 2014 +0200 @@ -6,7 +6,7 @@ header {* Lebsegue measure *} theory Lebesgue_Measure - imports Finite_Product_Measure + imports Finite_Product_Measure Bochner_Integration begin lemma absolutely_integrable_on_indicator[simp]: @@ -483,10 +483,14 @@ using sets.sigma_sets_eq[of borel] by (auto simp add: lborel_def measurable_def[abs_def]) +(* TODO: switch these rules! *) lemma emeasure_lborel[simp]: "A \ sets borel \ emeasure lborel A = emeasure lebesgue A" by (rule emeasure_measure_of[OF lborel_def]) (auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure) +lemma measure_lborel[simp]: "A \ sets borel \ measure lborel A = measure lebesgue A" + unfolding measure_def by simp + interpretation lborel: sigma_finite_measure lborel proof (default, intro conjI exI[of _ "\n. cube n"]) show "range cube \ sets lborel" by (auto intro: borel_closed) @@ -527,7 +531,9 @@ by (auto simp: emeasure_eq) } qed -lemma lebesgue_real_affine: + +(* GENEREALIZE to euclidean_spaces *) +lemma lborel_real_affine: fixes c :: real assumes "c \ 0" shows "lborel = density (distr lborel borel (\x. t + c * x)) (\_. \c\)" (is "_ = ?D") proof (rule lborel_eqI) @@ -551,200 +557,193 @@ qed qed simp -lemma lebesgue_integral_real_affine: - fixes c :: real assumes c: "c \ 0" and f: "f \ borel_measurable borel" - shows "(\ x. f x \ lborel) = \c\ * (\ x. f (t + c * x) \lborel)" - by (subst lebesgue_real_affine[OF c, of t]) - (simp add: f integral_density integral_distr lebesgue_integral_cmult) +lemma positive_integral_real_affine: + fixes c :: real assumes [measurable]: "f \ borel_measurable borel" and c: "c \ 0" + shows "(\\<^sup>+x. f x \lborel) = \c\ * (\\<^sup>+x. f (t + c * x) \lborel)" + by (subst lborel_real_affine[OF c, of t]) + (simp add: positive_integral_density positive_integral_distr positive_integral_cmult) + +lemma lborel_integrable_real_affine: + fixes f :: "real \ _ :: {banach, second_countable_topology}" + assumes f: "integrable lborel f" + shows "c \ 0 \ integrable lborel (\x. f (t + c * x))" + using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded + by (subst (asm) positive_integral_real_affine[where c=c and t=t]) auto + +lemma lborel_integrable_real_affine_iff: + fixes f :: "real \ 'a :: {banach, second_countable_topology}" + shows "c \ 0 \ integrable lborel (\x. f (t + c * x)) \ integrable lborel f" + using + lborel_integrable_real_affine[of f c t] + lborel_integrable_real_affine[of "\x. f (t + c * x)" "1/c" "-t/c"] + by (auto simp add: field_simps) + +lemma lborel_integral_real_affine: + fixes f :: "real \ 'a :: {banach, second_countable_topology}" and c :: real + assumes c: "c \ 0" and f[measurable]: "integrable lborel f" + shows "(\x. f x \ lborel) = \c\ *\<^sub>R (\x. f (t + c * x) \lborel)" + using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t] + by (subst lborel_real_affine[OF c, of t]) (simp add: integral_density integral_distr) + +lemma divideR_right: + fixes x y :: "'a::real_normed_vector" + shows "r \ 0 \ y = x /\<^sub>R r \ r *\<^sub>R y = x" + using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp + +lemma integrable_on_cmult_iff2: + fixes c :: real + shows "(\x. c * f x) integrable_on s \ c = 0 \ f integrable_on s" + using integrable_cmul[of "\x. c * f x" s "1 / c"] integrable_cmul[of f s c] + by (cases "c = 0") auto + +lemma lborel_has_bochner_integral_real_affine_iff: + fixes x :: "'a :: {banach, second_countable_topology}" + shows "c \ 0 \ + has_bochner_integral lborel f x \ + has_bochner_integral lborel (\x. f (t + c * x)) (x /\<^sub>R \c\)" + unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff + by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong) subsection {* Lebesgue integrable implies Gauge integrable *} -lemma simple_function_has_integral: - fixes f::"'a::ordered_euclidean_space \ ereal" - assumes f:"simple_function lebesgue f" - and f':"range f \ {0..<\}" - and om:"\x. x \ range f \ emeasure lebesgue (f -` {x} \ UNIV) = \ \ x = 0" - shows "((\x. real (f x)) has_integral (real (integral\<^sup>S lebesgue f))) UNIV" - unfolding simple_integral_def space_lebesgue -proof (subst lebesgue_simple_function_indicator) - let ?M = "\x. emeasure lebesgue (f -` {x} \ UNIV)" - let ?F = "\x. indicator (f -` {x})" - { fix x y assume "y \ range f" - from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)" - by (cases rule: ereal2_cases[of y "?F y x"]) - (auto simp: indicator_def one_ereal_def split: split_if_asm) } - moreover - { fix x assume x: "x\range f" - have "x * ?M x = real x * real (?M x)" - proof cases - assume "x \ 0" with om[OF x] have "?M x \ \" by auto - with subsetD[OF f' x] f[THEN simple_functionD(2)] show ?thesis - by (cases rule: ereal2_cases[of x "?M x"]) auto - qed simp } - ultimately - have "((\x. real (\y\range f. y * ?F y x)) has_integral real (\x\range f. x * ?M x)) UNIV \ - ((\x. \y\range f. real y * ?F y x) has_integral (\x\range f. real x * real (?M x))) UNIV" - by simp - also have \ - proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral - real_of_ereal_pos emeasure_nonneg ballI) - show *: "finite (range f)" "\y. f -` {y} \ sets lebesgue" - using simple_functionD[OF f] by auto - fix y assume "real y \ 0" "y \ range f" - with * om[OF this(2)] show "emeasure lebesgue (f -` {y}) = ereal (real (?M y))" - by (auto simp: ereal_real) +lemma has_integral_scaleR_left: + "(f has_integral y) s \ ((\x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s" + using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) + +lemma has_integral_mult_left: + fixes c :: "_ :: {real_normed_algebra}" + shows "(f has_integral y) s \ ((\x. f x * c) has_integral (y * c)) s" + using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) + +(* GENERALIZE Integration.dominated_convergence, then generalize the following theorems *) +lemma has_integral_dominated_convergence: + fixes f :: "nat \ 'n::ordered_euclidean_space \ real" + assumes "\k. (f k has_integral y k) s" "h integrable_on s" + "\k. \x\s. norm (f k x) \ h x" "\x\s. (\k. f k x) ----> g x" + and x: "y ----> x" + shows "(g has_integral x) s" +proof - + have int_f: "\k. (f k) integrable_on s" + using assms by (auto simp: integrable_on_def) + have "(g has_integral (integral s g)) s" + by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+ + moreover have "integral s g = x" + proof (rule LIMSEQ_unique) + show "(\i. integral s (f i)) ----> x" + using integral_unique[OF assms(1)] x by simp + show "(\i. integral s (f i)) ----> integral s g" + by (intro dominated_convergence[OF int_f assms(2)]) fact+ qed - finally show "((\x. real (\y\range f. y * ?F y x)) has_integral real (\x\range f. x * ?M x)) UNIV" . -qed fact - -lemma simple_function_has_integral': - fixes f::"'a::ordered_euclidean_space \ ereal" - assumes f: "simple_function lebesgue f" "\x. 0 \ f x" - and i: "integral\<^sup>S lebesgue f \ \" - shows "((\x. real (f x)) has_integral (real (integral\<^sup>S lebesgue f))) UNIV" -proof - - let ?f = "\x. if x \ f -` {\} then 0 else f x" - note f(1)[THEN simple_functionD(2)] - then have [simp, intro]: "\X. f -` X \ sets lebesgue" by auto - have f': "simple_function lebesgue ?f" - using f by (intro simple_function_If_set) auto - have rng: "range ?f \ {0..<\}" using f by auto - have "AE x in lebesgue. f x = ?f x" - using simple_integral_PInf[OF f i] - by (intro AE_I[where N="f -` {\} \ space lebesgue"]) auto - from f(1) f' this have eq: "integral\<^sup>S lebesgue f = integral\<^sup>S lebesgue ?f" - by (rule simple_integral_cong_AE) - have real_eq: "\x. real (f x) = real (?f x)" by auto - - show ?thesis - unfolding eq real_eq - proof (rule simple_function_has_integral[OF f' rng]) - fix x assume x: "x \ range ?f" and inf: "emeasure lebesgue (?f -` {x} \ UNIV) = \" - have "x * emeasure lebesgue (?f -` {x} \ UNIV) = (\\<^sup>S y. x * indicator (?f -` {x}) y \lebesgue)" - using f'[THEN simple_functionD(2)] - by (simp add: simple_integral_cmult_indicator) - also have "\ \ integral\<^sup>S lebesgue f" - using f'[THEN simple_functionD(2)] f - by (intro simple_integral_mono simple_function_mult simple_function_indicator) - (auto split: split_indicator) - finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm) - qed + ultimately show ?thesis + by simp qed lemma positive_integral_has_integral: - fixes f :: "'a::ordered_euclidean_space \ ereal" - assumes f: "f \ borel_measurable lebesgue" "range f \ {0..<\}" "integral\<^sup>P lebesgue f \ \" - shows "((\x. real (f x)) has_integral (real (integral\<^sup>P lebesgue f))) UNIV" -proof - - from borel_measurable_implies_simple_function_sequence'[OF f(1)] - guess u . note u = this - have SUP_eq: "\x. (SUP i. u i x) = f x" - using u(4) f(2)[THEN subsetD] by (auto split: split_max) - let ?u = "\i x. real (u i x)" - note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric] - { fix i - note u_eq - also have "integral\<^sup>P lebesgue (u i) \ (\\<^sup>+x. max 0 (f x) \lebesgue)" - by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric]) - finally have "integral\<^sup>S lebesgue (u i) \ \" - unfolding positive_integral_max_0 using f by auto } - note u_fin = this - then have u_int: "\i. (?u i has_integral real (integral\<^sup>S lebesgue (u i))) UNIV" - by (rule simple_function_has_integral'[OF u(1,5)]) - have "\x. \r\0. f x = ereal r" - proof - fix x from f(2) have "0 \ f x" "f x \ \" by (auto simp: subset_eq) - then show "\r\0. f x = ereal r" by (cases "f x") auto - qed - from choice[OF this] obtain f' where f': "f = (\x. ereal (f' x))" "\x. 0 \ f' x" by auto - - have "\i. \r. \x. 0 \ r x \ u i x = ereal (r x)" - proof - fix i show "\r. \x. 0 \ r x \ u i x = ereal (r x)" - proof (intro choice allI) - fix i x have "u i x \ \" using u(3)[of i] by (auto simp: image_iff) metis - then show "\r\0. u i x = ereal r" using u(5)[of i x] by (cases "u i x") auto - qed - qed - from choice[OF this] obtain u' where - u': "u = (\i x. ereal (u' i x))" "\i x. 0 \ u' i x" by (auto simp: fun_eq_iff) + fixes f::"'a::ordered_euclidean_space \ real" + assumes f: "f \ borel_measurable lebesgue" "\x. 0 \ f x" "(\\<^sup>+x. f x \lebesgue) = ereal r" + shows "(f has_integral r) UNIV" +using f proof (induct arbitrary: r rule: borel_measurable_induct_real) + case (set A) then show ?case + by (auto simp add: ereal_indicator has_integral_iff_lmeasure) +next + case (mult g c) + then have "ereal c * (\\<^sup>+ x. g x \lebesgue) = ereal r" + by (subst positive_integral_cmult[symmetric]) auto + then obtain r' where "(c = 0 \ r = 0) \ ((\\<^sup>+ x. ereal (g x) \lebesgue) = ereal r' \ r = c * r')" + by (cases "\\<^sup>+ x. ereal (g x) \lebesgue") (auto split: split_if_asm) + with mult show ?case + by (auto intro!: has_integral_cmult_real) +next + case (add g h) + moreover + then have "(\\<^sup>+ x. h x + g x \lebesgue) = (\\<^sup>+ x. h x \lebesgue) + (\\<^sup>+ x. g x \lebesgue)" + unfolding plus_ereal.simps[symmetric] by (subst positive_integral_add) auto + with add obtain a b where "(\\<^sup>+ x. h x \lebesgue) = ereal a" "(\\<^sup>+ x. g x \lebesgue) = ereal b" "r = a + b" + by (cases "\\<^sup>+ x. h x \lebesgue" "\\<^sup>+ x. g x \lebesgue" rule: ereal2_cases) auto + ultimately show ?case + by (auto intro!: has_integral_add) +next + case (seq U) + note seq(1)[measurable] and f[measurable] - have convergent: "f' integrable_on UNIV \ - (\k. integral UNIV (u' k)) ----> integral UNIV f'" - proof (intro monotone_convergence_increasing allI ballI) - show int: "\k. (u' k) integrable_on UNIV" - using u_int unfolding integrable_on_def u' by auto - show "\k x. u' k x \ u' (Suc k) x" using u(2,3,5) - by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_ereal_positive_mono) - show "\x. (\k. u' k x) ----> f' x" - using SUP_eq u(2) - by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_Suc_iff le_fun_def) - show "bounded {integral UNIV (u' k)|k. True}" - proof (safe intro!: bounded_realI) - fix k - have "\integral UNIV (u' k)\ = integral UNIV (u' k)" - by (intro abs_of_nonneg integral_nonneg int ballI u') - also have "\ = real (integral\<^sup>S lebesgue (u k))" - using u_int[THEN integral_unique] by (simp add: u') - also have "\ = real (integral\<^sup>P lebesgue (u k))" - using positive_integral_eq_simple_integral[OF u(1,5)] by simp - also have "\ \ real (integral\<^sup>P lebesgue f)" using f - by (auto intro!: real_of_ereal_positive_mono positive_integral_positive - positive_integral_mono SUP_upper simp: SUP_eq[symmetric]) - finally show "\integral UNIV (u' k)\ \ real (integral\<^sup>P lebesgue f)" . - qed + { fix i x + have "U i x \ f x" + using seq(5) + apply (rule LIMSEQ_le_const) + using seq(4) + apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def) + done } + note U_le_f = this + + { fix i + have "(\\<^sup>+x. ereal (U i x) \lebesgue) \ (\\<^sup>+x. ereal (f x) \lebesgue)" + using U_le_f by (intro positive_integral_mono) simp + then obtain p where "(\\<^sup>+x. U i x \lebesgue) = ereal p" "p \ r" + using seq(6) by (cases "\\<^sup>+x. U i x \lebesgue") auto + moreover then have "0 \ p" + by (metis ereal_less_eq(5) positive_integral_positive) + moreover note seq + ultimately have "\p. (\\<^sup>+x. U i x \lebesgue) = ereal p \ 0 \ p \ p \ r \ (U i has_integral p) UNIV" + by auto } + then obtain p where p: "\i. (\\<^sup>+x. ereal (U i x) \lebesgue) = ereal (p i)" + and bnd: "\i. p i \ r" "\i. 0 \ p i" + and U_int: "\i.(U i has_integral (p i)) UNIV" by metis + + have int_eq: "\i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) + + have *: "f integrable_on UNIV \ (\k. integral UNIV (U k)) ----> integral UNIV f" + proof (rule monotone_convergence_increasing) + show "\k. U k integrable_on UNIV" using U_int by auto + show "\k. \x\UNIV. U k x \ U (Suc k) x" using `incseq U` by (auto simp: incseq_def le_fun_def) + then show "bounded {integral UNIV (U k) |k. True}" + using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) + show "\x\UNIV. (\k. U k x) ----> f x" + using seq by auto qed - - have "integral\<^sup>P lebesgue f = ereal (integral UNIV f')" - proof (rule tendsto_unique[OF trivial_limit_sequentially]) - have "(\i. integral\<^sup>S lebesgue (u i)) ----> (SUP i. integral\<^sup>P lebesgue (u i))" - unfolding u_eq by (intro LIMSEQ_SUP incseq_positive_integral u) - also note positive_integral_monotone_convergence_SUP - [OF u(2) borel_measurable_simple_function[OF u(1)] u(5), symmetric] - finally show "(\k. integral\<^sup>S lebesgue (u k)) ----> integral\<^sup>P lebesgue f" - unfolding SUP_eq . - - { fix k - have "0 \ integral\<^sup>S lebesgue (u k)" - using u by (auto intro!: simple_integral_positive) - then have "integral\<^sup>S lebesgue (u k) = ereal (real (integral\<^sup>S lebesgue (u k)))" - using u_fin by (auto simp: ereal_real) } - note * = this - show "(\k. integral\<^sup>S lebesgue (u k)) ----> ereal (integral UNIV f')" - using convergent using u_int[THEN integral_unique, symmetric] - by (subst *) (simp add: u') - qed - then show ?thesis using convergent by (simp add: f' integrable_integral) + moreover have "(\i. (\\<^sup>+x. U i x \lebesgue)) ----> (\\<^sup>+x. f x \lebesgue)" + using seq U_le_f by (intro positive_integral_dominated_convergence[where w=f]) auto + ultimately have "integral UNIV f = r" + by (auto simp add: int_eq p seq intro: LIMSEQ_unique) + with * show ?case + by (simp add: has_integral_integral) qed -lemma lebesgue_integral_has_integral: - fixes f :: "'a::ordered_euclidean_space \ real" +lemma has_integral_integrable_lebesgue_nonneg: + fixes f::"'a::ordered_euclidean_space \ real" + assumes f: "integrable lebesgue f" "\x. 0 \ f x" + shows "(f has_integral integral\<^sup>L lebesgue f) UNIV" +proof (rule positive_integral_has_integral) + show "(\\<^sup>+ x. ereal (f x) \lebesgue) = ereal (integral\<^sup>L lebesgue f)" + using f by (intro positive_integral_eq_integral) auto +qed (insert f, auto) + +lemma has_integral_lebesgue_integral_lebesgue: + fixes f::"'a::ordered_euclidean_space \ real" assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" -proof - - let ?n = "\x. real (ereal (max 0 (- f x)))" and ?p = "\x. real (ereal (max 0 (f x)))" - have *: "f = (\x. ?p x - ?n x)" by (auto simp del: ereal_max) - { fix f :: "'a \ real" have "(\\<^sup>+ x. ereal (f x) \lebesgue) = (\\<^sup>+ x. ereal (max 0 (f x)) \lebesgue)" - by (intro positive_integral_cong_pos) (auto split: split_max) } - note eq = this - show ?thesis - unfolding lebesgue_integral_def - apply (subst *) - apply (rule has_integral_sub) - unfolding eq[of f] eq[of "\x. - f x"] - apply (safe intro!: positive_integral_has_integral) - using integrableD[OF f] - by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0 split: split_max - intro!: measurable_If) +using f proof induct + case (base A c) then show ?case + by (auto intro!: has_integral_mult_left simp: has_integral_iff_lmeasure) + (simp add: emeasure_eq_ereal_measure) +next + case (add f g) then show ?case + by (auto intro!: has_integral_add) +next + case (lim f s) + show ?case + proof (rule has_integral_dominated_convergence) + show "\i. (s i has_integral integral\<^sup>L lebesgue (s i)) UNIV" by fact + show "(\x. norm (2 * f x)) integrable_on UNIV" + using lim by (intro has_integral_integrable[OF has_integral_integrable_lebesgue_nonneg]) auto + show "\k. \x\UNIV. norm (s k x) \ norm (2 * f x)" + using lim by (auto simp add: abs_mult) + show "\x\UNIV. (\k. s k x) ----> f x" + using lim by auto + show "(\k. integral\<^sup>L lebesgue (s k)) ----> integral\<^sup>L lebesgue f" + using lim by (intro integral_dominated_convergence(3)[where w="\x. 2 * norm (f x)"]) auto + qed qed -lemma lebesgue_simple_integral_eq_borel: - assumes f: "f \ borel_measurable borel" - shows "integral\<^sup>S lebesgue f = integral\<^sup>S lborel f" - using f[THEN measurable_sets] - by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_lborel[symmetric] - simp: simple_integral_def) - lemma lebesgue_positive_integral_eq_borel: assumes f: "f \ borel_measurable borel" shows "integral\<^sup>P lebesgue f = integral\<^sup>P lborel f" @@ -755,40 +754,70 @@ qed lemma lebesgue_integral_eq_borel: + fixes f :: "_ \ _::{banach, second_countable_topology}" assumes "f \ borel_measurable borel" shows "integrable lebesgue f \ integrable lborel f" (is ?P) and "integral\<^sup>L lebesgue f = integral\<^sup>L lborel f" (is ?I) proof - have "sets lborel \ sets lebesgue" by auto - from integral_subalgebra[of f lborel, OF _ this _ _] assms + from integral_subalgebra[of f lborel, OF _ this _ _] + integrable_subalgebra[of f lborel, OF _ this _ _] assms show ?P ?I by auto qed -lemma borel_integral_has_integral: +lemma has_integral_lebesgue_integral: fixes f::"'a::ordered_euclidean_space => real" assumes f:"integrable lborel f" shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" proof - have borel: "f \ borel_measurable borel" - using f unfolding integrable_def by auto + using f unfolding integrable_iff_bounded by auto from f show ?thesis - using lebesgue_integral_has_integral[of f] + using has_integral_lebesgue_integral_lebesgue[of f] unfolding lebesgue_integral_eq_borel[OF borel] by simp qed -lemma positive_integral_lebesgue_has_integral: +lemma positive_integral_bound_simple_function: + assumes bnd: "\x. x \ space M \ 0 \ f x" "\x. x \ space M \ f x < \" + assumes f[measurable]: "simple_function M f" + assumes supp: "emeasure M {x\space M. f x \ 0} < \" + shows "positive_integral M f < \" +proof cases + assume "space M = {}" + then have "positive_integral M f = (\\<^sup>+x. 0 \M)" + by (intro positive_integral_cong) auto + then show ?thesis by simp +next + assume "space M \ {}" + with simple_functionD(1)[OF f] bnd have bnd: "0 \ Max (f`space M) \ Max (f`space M) < \" + by (subst Max_less_iff) (auto simp: Max_ge_iff) + + have "positive_integral M f \ (\\<^sup>+x. Max (f`space M) * indicator {x\space M. f x \ 0} x \M)" + proof (rule positive_integral_mono) + fix x assume "x \ space M" + with f show "f x \ Max (f ` space M) * indicator {x \ space M. f x \ 0} x" + by (auto split: split_indicator intro!: Max_ge simple_functionD) + qed + also have "\ < \" + using bnd supp by (subst positive_integral_cmult) auto + finally show ?thesis . +qed + + +lemma fixes f :: "'a::ordered_euclidean_space \ real" assumes f_borel: "f \ borel_measurable lebesgue" and nonneg: "\x. 0 \ f x" assumes I: "(f has_integral I) UNIV" - shows "(\\<^sup>+x. f x \lebesgue) = I" + shows integrable_has_integral_lebesgue_nonneg: "integrable lebesgue f" + and integral_has_integral_lebesgue_nonneg: "integral\<^sup>L lebesgue f = I" proof - from f_borel have "(\x. ereal (f x)) \ borel_measurable lebesgue" by auto from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this - have "(\\<^sup>+ x. ereal (f x) \lebesgue) = (SUP i. integral\<^sup>S lebesgue (F i))" + have "(\\<^sup>+ x. ereal (f x) \lebesgue) = (SUP i. integral\<^sup>P lebesgue (F i))" using F - by (subst positive_integral_monotone_convergence_simple) - (simp_all add: positive_integral_max_0 simple_function_def) + by (subst positive_integral_monotone_convergence_SUP[symmetric]) + (simp_all add: positive_integral_max_0 borel_measurable_simple_function) also have "\ \ ereal I" proof (rule SUP_least) fix i :: nat @@ -835,58 +864,58 @@ by (simp add: integrable_on_cmult_iff) } note F_finite = lmeasure_finite[OF this] - have "((\x. real (F i x)) has_integral real (integral\<^sup>S lebesgue (F i))) UNIV" - proof (rule simple_function_has_integral[of "F i"]) - show "simple_function lebesgue (F i)" - using F(1) by (simp add: simple_function_def) - show "range (F i) \ {0..<\}" - using F(3,5)[of i] by (auto simp: image_iff) metis - next - fix x assume "x \ range (F i)" "emeasure lebesgue (F i -` {x} \ UNIV) = \" - with F_finite[of x] show "x = 0" by auto - qed - from this I have "real (integral\<^sup>S lebesgue (F i)) \ I" + have F_eq: "\x. F i x = ereal (norm (real (F i x)))" + using F(3,5) by (auto simp: fun_eq_iff ereal_real image_iff eq_commute) + have F_eq2: "\x. F i x = ereal (real (F i x))" + using F(3,5) by (auto simp: fun_eq_iff ereal_real image_iff eq_commute) + + have int: "integrable lebesgue (\x. real (F i x))" + unfolding integrable_iff_bounded + proof + have "(\\<^sup>+x. F i x \lebesgue) < \" + proof (rule positive_integral_bound_simple_function) + fix x::'a assume "x \ space lebesgue" then show "0 \ F i x" "F i x < \" + using F by (auto simp: image_iff eq_commute) + next + have eq: "{x \ space lebesgue. F i x \ 0} = (\x\F i ` space lebesgue - {0}. F i -` {x} \ space lebesgue)" + by auto + show "emeasure lebesgue {x \ space lebesgue. F i x \ 0} < \" + unfolding eq using simple_functionD[OF F(1)] + by (subst setsum_emeasure[symmetric]) + (auto simp: disjoint_family_on_def setsum_Pinfty F_finite) + qed fact + with F_eq show "(\\<^sup>+x. norm (real (F i x)) \lebesgue) < \" by simp + qed (insert F(1), auto intro!: borel_measurable_real_of_ereal dest: borel_measurable_simple_function) + then have "((\x. real (F i x)) has_integral integral\<^sup>L lebesgue (\x. real (F i x))) UNIV" + by (rule has_integral_lebesgue_integral_lebesgue) + from this I have "integral\<^sup>L lebesgue (\x. real (F i x)) \ I" by (rule has_integral_le) (intro ballI F_bound) - moreover - { fix x assume x: "x \ range (F i)" - with F(3,5)[of i] have "x = 0 \ (0 < x \ x \ \)" - by (auto simp: image_iff le_less) metis - with F_finite[OF _ x] x have "x * emeasure lebesgue (F i -` {x} \ UNIV) \ \" - by auto } - then have "integral\<^sup>S lebesgue (F i) \ \" - unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast - moreover have "0 \ integral\<^sup>S lebesgue (F i)" - using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def) - ultimately show "integral\<^sup>S lebesgue (F i) \ ereal I" - by (cases "integral\<^sup>S lebesgue (F i)") auto + moreover have "integral\<^sup>P lebesgue (F i) = integral\<^sup>L lebesgue (\x. real (F i x))" + using int F by (subst positive_integral_eq_integral[symmetric]) (auto simp: F_eq2[symmetric] real_of_ereal_pos) + ultimately show "integral\<^sup>P lebesgue (F i) \ ereal I" + by simp qed - also have "\ < \" by simp - finally have finite: "(\\<^sup>+ x. ereal (f x) \lebesgue) \ \" by simp - have borel: "(\x. ereal (f x)) \ borel_measurable lebesgue" - using f_borel by (auto intro: borel_measurable_lebesgueI) - from positive_integral_has_integral[OF borel _ finite] - have "(f has_integral real (\\<^sup>+ x. ereal (f x) \lebesgue)) UNIV" - using nonneg by (simp add: subset_eq) - with I have "I = real (\\<^sup>+ x. ereal (f x) \lebesgue)" - by (rule has_integral_unique) - with finite positive_integral_positive[of _ "\x. ereal (f x)"] show ?thesis - by (cases "\\<^sup>+ x. ereal (f x) \lebesgue") auto + finally show "integrable lebesgue f" + using f_borel by (auto simp: integrable_iff_bounded nonneg) + from has_integral_lebesgue_integral_lebesgue[OF this] I + show "integral\<^sup>L lebesgue f = I" + by (metis has_integral_unique) qed -lemma has_integral_iff_positive_integral_lebesgue: +lemma has_integral_iff_has_bochner_integral_lebesgue_nonneg: fixes f :: "'a::ordered_euclidean_space \ real" - assumes f: "f \ borel_measurable lebesgue" "\x. 0 \ f x" - shows "(f has_integral I) UNIV \ integral\<^sup>P lebesgue f = I" - using f positive_integral_lebesgue_has_integral[of f I] positive_integral_has_integral[of f] - by (auto simp: subset_eq) + shows "f \ borel_measurable lebesgue \ (\x. 0 \ f x) \ + (f has_integral I) UNIV \ has_bochner_integral lebesgue f I" + by (metis has_bochner_integral_iff has_integral_unique has_integral_lebesgue_integral_lebesgue + integrable_has_integral_lebesgue_nonneg) -lemma has_integral_iff_positive_integral_lborel: +lemma fixes f :: "'a::ordered_euclidean_space \ real" - assumes f: "f \ borel_measurable borel" "\x. 0 \ f x" - shows "(f has_integral I) UNIV \ integral\<^sup>P lborel f = I" - using assms - by (subst has_integral_iff_positive_integral_lebesgue) - (auto simp: borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel) + assumes "f \ borel_measurable borel" "\x. 0 \ f x" "(f has_integral I) UNIV" + shows integrable_has_integral_nonneg: "integrable lborel f" + and integral_has_integral_nonneg: "integral\<^sup>L lborel f = I" + by (metis assms borel_measurable_lebesgueI integrable_has_integral_lebesgue_nonneg lebesgue_integral_eq_borel(1)) + (metis assms borel_measurable_lebesgueI has_integral_lebesgue_integral has_integral_unique integrable_has_integral_lebesgue_nonneg lebesgue_integral_eq_borel(1)) subsection {* Equivalence between product spaces and euclidean spaces *} @@ -978,52 +1007,36 @@ by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f) lemma borel_fubini_integrable: - fixes f :: "'a::ordered_euclidean_space \ real" + fixes f :: "'a::ordered_euclidean_space \ _::{banach, second_countable_topology}" shows "integrable lborel f \ integrable (\\<^sub>M (i::'a)\Basis. lborel) (\x. f (p2e x))" - (is "_ \ integrable ?B ?f") -proof - assume *: "integrable lborel f" - then have f: "f \ borel_measurable borel" - by auto - with measurable_p2e have "f \ p2e \ borel_measurable ?B" - by (rule measurable_comp) - with * f show "integrable ?B ?f" - by (simp add: comp_def borel_fubini_positiv_integral integrable_def) -next - assume *: "integrable ?B ?f" - then have "?f \ e2p \ borel_measurable (borel::'a measure)" - by (auto intro!: measurable_e2p) - then have "f \ borel_measurable borel" - by (simp cong: measurable_cong) - with * show "integrable lborel f" - by (simp add: borel_fubini_positiv_integral integrable_def) -qed + unfolding integrable_iff_bounded +proof (intro conj_cong[symmetric]) + show "((\x. f (p2e x)) \ borel_measurable (Pi\<^sub>M Basis (\i. lborel))) = (f \ borel_measurable lborel)" + proof + assume "((\x. f (p2e x)) \ borel_measurable (Pi\<^sub>M Basis (\i. lborel)))" + then have "(\x. f (p2e (e2p x))) \ borel_measurable borel" + by measurable + then show "f \ borel_measurable lborel" + by simp + qed simp +qed (simp add: borel_fubini_positiv_integral) lemma borel_fubini: - fixes f :: "'a::ordered_euclidean_space \ real" - assumes f: "f \ borel_measurable borel" - shows "integral\<^sup>L lborel f = \x. f (p2e x) \((\\<^sub>M (i::'a)\Basis. lborel))" - using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def) + fixes f :: "'a::ordered_euclidean_space \ _::{banach, second_countable_topology}" + shows "f \ borel_measurable borel \ + integral\<^sup>L lborel f = \x. f (p2e x) \((\\<^sub>M (i::'a)\Basis. lborel))" + by (subst lborel_eq_lborel_space) (simp add: integral_distr) lemma integrable_on_borel_integrable: fixes f :: "'a::ordered_euclidean_space \ real" - assumes f_borel: "f \ borel_measurable borel" and nonneg: "\x. 0 \ f x" - assumes f: "f integrable_on UNIV" - shows "integrable lborel f" -proof - - have "(\\<^sup>+ x. ereal (f x) \lborel) \ \" - using has_integral_iff_positive_integral_lborel[OF f_borel nonneg] f - by (auto simp: integrable_on_def) - moreover have "(\\<^sup>+ x. ereal (- f x) \lborel) = 0" - using f_borel nonneg by (subst positive_integral_0_iff_AE) auto - ultimately show ?thesis - using f_borel by (auto simp: integrable_def) -qed + shows "f \ borel_measurable borel \ (\x. 0 \ f x) \ f integrable_on UNIV \ integrable lborel f" + by (metis borel_measurable_lebesgueI integrable_has_integral_nonneg integrable_on_def + lebesgue_integral_eq_borel(1)) subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *} lemma borel_integrable_atLeastAtMost: - fixes a b :: real + fixes f :: "real \ real" assumes f: "\x. a \ x \ x \ b \ isCont f x" shows "integrable lborel (\x. f x * indicator {a .. b} x)" (is "integrable _ ?f") proof cases @@ -1037,10 +1050,10 @@ show ?thesis proof (rule integrable_bound) show "integrable lborel (\x. max \M\ \L\ * indicator {a..b} x)" - by (rule integral_cmul_indicator) simp_all - show "AE x in lborel. \?f x\ \ max \M\ \L\ * indicator {a..b} x" + by (intro integrable_mult_right integrable_real_indicator) simp_all + show "AE x in lborel. norm (?f x) \ norm (max \M\ \L\ * indicator {a..b} x)" proof (rule AE_I2) - fix x show "\?f x\ \ max \M\ \L\ * indicator {a..b} x" + fix x show "norm (?f x) \ norm (max \M\ \L\ * indicator {a..b} x)" using bounds[of x] by (auto split: split_indicator) qed @@ -1071,7 +1084,7 @@ let ?f = "\x. f x * indicator {a .. b} x" have "(?f has_integral (\x. ?f x \lborel)) UNIV" using borel_integrable_atLeastAtMost[OF f] - by (rule borel_integral_has_integral) + by (rule has_integral_lebesgue_integral) moreover have "(f has_integral F b - F a) {a .. b}" by (intro fundamental_theorem_of_calculus) @@ -1091,10 +1104,12 @@ *} -lemma positive_integral_FTC_atLeastAtMost: +lemma + fixes f :: "real \ real" assumes f_borel: "f \ borel_measurable borel" assumes f: "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" and "a \ b" - shows "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" + shows integral_FTC_Icc_nonneg: "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" (is ?eq) + and integrable_FTC_Icc_nonneg: "integrable lborel (\x. f x * indicator {a .. b} x)" (is ?int) proof - have i: "(f has_integral F b - F a) {a..b}" by (intro fundamental_theorem_of_calculus) @@ -1104,12 +1119,24 @@ by (rule has_integral_eq[OF _ i]) auto have i: "((\x. f x * indicator {a..b} x) has_integral F b - F a) UNIV" by (rule has_integral_on_superset[OF _ _ i]) auto - then have "(\\<^sup>+x. ereal (f x * indicator {a .. b} x) \lborel) = F b - F a" - using f f_borel - by (subst has_integral_iff_positive_integral_lborel[symmetric]) (auto split: split_indicator) - also have "(\\<^sup>+x. ereal (f x * indicator {a .. b} x) \lborel) = (\\<^sup>+x. ereal (f x) * indicator {a .. b} x \lborel)" - by (auto intro!: positive_integral_cong simp: indicator_def) - finally show ?thesis by simp + from i f f_borel show ?eq + by (intro integral_has_integral_nonneg) (auto split: split_indicator) + from i f f_borel show ?int + by (intro integrable_has_integral_nonneg) (auto split: split_indicator) +qed + +lemma positive_integral_FTC_atLeastAtMost: + assumes "f \ borel_measurable borel" "\x. x \ {a..b} \ DERIV F x :> f x" "\x. x \ {a..b} \ 0 \ f x" "a \ b" + shows "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = F b - F a" +proof - + have "integrable lborel (\x. f x * indicator {a .. b} x)" + by (rule integrable_FTC_Icc_nonneg) fact+ + then have "(\\<^sup>+x. f x * indicator {a .. b} x \lborel) = (\x. f x * indicator {a .. b} x \lborel)" + using assms by (intro positive_integral_eq_integral) (auto simp: indicator_def) + also have "(\x. f x * indicator {a .. b} x \lborel) = F b - F a" + by (rule integral_FTC_Icc_nonneg) fact+ + finally show ?thesis + unfolding ereal_indicator[symmetric] by simp qed lemma positive_integral_FTC_atLeast: diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Measurable.thy --- a/src/HOL/Probability/Measurable.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Measurable.thy Mon May 19 12:04:45 2014 +0200 @@ -256,6 +256,76 @@ shows "(\x. LEAST i. P i x) \ measurable M (count_space UNIV)" unfolding measurable_def by (safe intro!: sets_Least) simp_all +lemma measurable_Max_nat[measurable (raw)]: + fixes P :: "nat \ 'a \ bool" + assumes [measurable]: "\i. Measurable.pred M (P i)" + shows "(\x. Max {i. P i x}) \ measurable M (count_space UNIV)" + unfolding measurable_count_space_eq2_countable +proof safe + fix n + + { fix x assume "\i. \n\i. P n x" + then have "infinite {i. P i x}" + unfolding infinite_nat_iff_unbounded_le by auto + then have "Max {i. P i x} = the None" + by (rule Max.infinite) } + note 1 = this + + { fix x i j assume "P i x" "\n\j. \ P n x" + then have "finite {i. P i x}" + by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded) + with `P i x` have "P (Max {i. P i x}) x" "i \ Max {i. P i x}" "finite {i. P i x}" + using Max_in[of "{i. P i x}"] by auto } + note 2 = this + + have "(\x. Max {i. P i x}) -` {n} \ space M = {x\space M. Max {i. P i x} = n}" + by auto + also have "\ = + {x\space M. if (\i. \n\i. P n x) then the None = n else + if (\i. P i x) then P n x \ (\i>n. \ P i x) + else Max {} = n}" + by (intro arg_cong[where f=Collect] ext conj_cong) + (auto simp add: 1 2 not_le[symmetric] intro!: Max_eqI) + also have "\ \ sets M" + by measurable + finally show "(\x. Max {i. P i x}) -` {n} \ space M \ sets M" . +qed simp + +lemma measurable_Min_nat[measurable (raw)]: + fixes P :: "nat \ 'a \ bool" + assumes [measurable]: "\i. Measurable.pred M (P i)" + shows "(\x. Min {i. P i x}) \ measurable M (count_space UNIV)" + unfolding measurable_count_space_eq2_countable +proof safe + fix n + + { fix x assume "\i. \n\i. P n x" + then have "infinite {i. P i x}" + unfolding infinite_nat_iff_unbounded_le by auto + then have "Min {i. P i x} = the None" + by (rule Min.infinite) } + note 1 = this + + { fix x i j assume "P i x" "\n\j. \ P n x" + then have "finite {i. P i x}" + by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded) + with `P i x` have "P (Min {i. P i x}) x" "Min {i. P i x} \ i" "finite {i. P i x}" + using Min_in[of "{i. P i x}"] by auto } + note 2 = this + + have "(\x. Min {i. P i x}) -` {n} \ space M = {x\space M. Min {i. P i x} = n}" + by auto + also have "\ = + {x\space M. if (\i. \n\i. P n x) then the None = n else + if (\i. P i x) then P n x \ (\i P i x) + else Min {} = n}" + by (intro arg_cong[where f=Collect] ext conj_cong) + (auto simp add: 1 2 not_le[symmetric] intro!: Min_eqI) + also have "\ \ sets M" + by measurable + finally show "(\x. Min {i. P i x}) -` {n} \ space M \ sets M" . +qed simp + lemma measurable_count_space_insert[measurable (raw)]: "s \ S \ A \ sets (count_space S) \ insert s A \ sets (count_space S)" by simp diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon May 19 12:04:45 2014 +0200 @@ -0,0 +1,2039 @@ +(* Title: HOL/Probability/Nonnegative_Lebesgue_Integration.thy + Author: Johannes Hölzl, TU München + Author: Armin Heller, TU München +*) + +header {* Lebesgue Integration for Nonnegative Functions *} + +theory Nonnegative_Lebesgue_Integration + imports Measure_Space Borel_Space +begin + +lemma indicator_less_ereal[simp]: + "indicator A x \ (indicator B x::ereal) \ (x \ A \ x \ B)" + by (simp add: indicator_def not_le) + +section "Simple function" + +text {* + +Our simple functions are not restricted to positive real numbers. Instead +they are just functions with a finite range and are measurable when singleton +sets are measurable. + +*} + +definition "simple_function M g \ + finite (g ` space M) \ + (\x \ g ` space M. g -` {x} \ space M \ sets M)" + +lemma simple_functionD: + assumes "simple_function M g" + shows "finite (g ` space M)" and "g -` X \ space M \ sets M" +proof - + show "finite (g ` space M)" + using assms unfolding simple_function_def by auto + have "g -` X \ space M = g -` (X \ g`space M) \ space M" by auto + also have "\ = (\x\X \ g`space M. g-`{x} \ space M)" by auto + finally show "g -` X \ space M \ sets M" using assms + by (auto simp del: UN_simps simp: simple_function_def) +qed + +lemma measurable_simple_function[measurable_dest]: + "simple_function M f \ f \ measurable M (count_space UNIV)" + unfolding simple_function_def measurable_def +proof safe + fix A assume "finite (f ` space M)" "\x\f ` space M. f -` {x} \ space M \ sets M" + then have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) \ sets M" + by (intro sets.finite_UN) auto + also have "(\x\f ` space M. if x \ A then f -` {x} \ space M else {}) = f -` A \ space M" + by (auto split: split_if_asm) + finally show "f -` A \ space M \ sets M" . +qed simp + +lemma borel_measurable_simple_function: + "simple_function M f \ f \ borel_measurable M" + by (auto dest!: measurable_simple_function simp: measurable_def) + +lemma simple_function_measurable2[intro]: + assumes "simple_function M f" "simple_function M g" + shows "f -` A \ g -` B \ space M \ sets M" +proof - + have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" + by auto + then show ?thesis using assms[THEN simple_functionD(2)] by auto +qed + +lemma simple_function_indicator_representation: + fixes f ::"'a \ ereal" + assumes f: "simple_function M f" and x: "x \ space M" + shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" + (is "?l = ?r") +proof - + have "?r = (\y \ f ` space M. + (if y = f x then y * indicator (f -` {y} \ space M) x else 0))" + by (auto intro!: setsum_cong2) + also have "... = f x * indicator (f -` {f x} \ space M) x" + using assms by (auto dest: simple_functionD simp: setsum_delta) + also have "... = f x" using x by (auto simp: indicator_def) + finally show ?thesis by auto +qed + +lemma simple_function_notspace: + "simple_function M (\x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h") +proof - + have "?h ` space M \ {0}" unfolding indicator_def by auto + hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) + have "?h -` {0} \ space M = space M" by auto + thus ?thesis unfolding simple_function_def by auto +qed + +lemma simple_function_cong: + assumes "\t. t \ space M \ f t = g t" + shows "simple_function M f \ simple_function M g" +proof - + have "f ` space M = g ` space M" + "\x. f -` {x} \ space M = g -` {x} \ space M" + using assms by (auto intro!: image_eqI) + thus ?thesis unfolding simple_function_def using assms by simp +qed + +lemma simple_function_cong_algebra: + assumes "sets N = sets M" "space N = space M" + shows "simple_function M f \ simple_function N f" + unfolding simple_function_def assms .. + +lemma simple_function_borel_measurable: + fixes f :: "'a \ 'x::{t2_space}" + assumes "f \ borel_measurable M" and "finite (f ` space M)" + shows "simple_function M f" + using assms unfolding simple_function_def + by (auto intro: borel_measurable_vimage) + +lemma simple_function_eq_measurable: + fixes f :: "'a \ ereal" + shows "simple_function M f \ finite (f`space M) \ f \ measurable M (count_space UNIV)" + using simple_function_borel_measurable[of f] measurable_simple_function[of M f] + by (fastforce simp: simple_function_def) + +lemma simple_function_const[intro, simp]: + "simple_function M (\x. c)" + by (auto intro: finite_subset simp: simple_function_def) +lemma simple_function_compose[intro, simp]: + assumes "simple_function M f" + shows "simple_function M (g \ f)" + unfolding simple_function_def +proof safe + show "finite ((g \ f) ` space M)" + using assms unfolding simple_function_def by (auto simp: image_comp [symmetric]) +next + fix x assume "x \ space M" + let ?G = "g -` {g (f x)} \ (f`space M)" + have *: "(g \ f) -` {(g \ f) x} \ space M = + (\x\?G. f -` {x} \ space M)" by auto + show "(g \ f) -` {(g \ f) x} \ space M \ sets M" + using assms unfolding simple_function_def * + by (rule_tac sets.finite_UN) auto +qed + +lemma simple_function_indicator[intro, simp]: + assumes "A \ sets M" + shows "simple_function M (indicator A)" +proof - + have "indicator A ` space M \ {0, 1}" (is "?S \ _") + by (auto simp: indicator_def) + hence "finite ?S" by (rule finite_subset) simp + moreover have "- A \ space M = space M - A" by auto + ultimately show ?thesis unfolding simple_function_def + using assms by (auto simp: indicator_def [abs_def]) +qed + +lemma simple_function_Pair[intro, simp]: + assumes "simple_function M f" + assumes "simple_function M g" + shows "simple_function M (\x. (f x, g x))" (is "simple_function M ?p") + unfolding simple_function_def +proof safe + show "finite (?p ` space M)" + using assms unfolding simple_function_def + by (rule_tac finite_subset[of _ "f`space M \ g`space M"]) auto +next + fix x assume "x \ space M" + have "(\x. (f x, g x)) -` {(f x, g x)} \ space M = + (f -` {f x} \ space M) \ (g -` {g x} \ space M)" + by auto + with `x \ space M` show "(\x. (f x, g x)) -` {(f x, g x)} \ space M \ sets M" + using assms unfolding simple_function_def by auto +qed + +lemma simple_function_compose1: + assumes "simple_function M f" + shows "simple_function M (\x. g (f x))" + using simple_function_compose[OF assms, of g] + by (simp add: comp_def) + +lemma simple_function_compose2: + assumes "simple_function M f" and "simple_function M g" + shows "simple_function M (\x. h (f x) (g x))" +proof - + have "simple_function M ((\(x, y). h x y) \ (\x. (f x, g x)))" + using assms by auto + thus ?thesis by (simp_all add: comp_def) +qed + +lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] + and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] + and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] + and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] + and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] + and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] + and simple_function_max[intro, simp] = simple_function_compose2[where h=max] + +lemma simple_function_setsum[intro, simp]: + assumes "\i. i \ P \ simple_function M (f i)" + shows "simple_function M (\x. \i\P. f i x)" +proof cases + assume "finite P" from this assms show ?thesis by induct auto +qed auto + +lemma simple_function_ereal[intro, simp]: + fixes f g :: "'a \ real" assumes sf: "simple_function M f" + shows "simple_function M (\x. ereal (f x))" + by (auto intro!: simple_function_compose1[OF sf]) + +lemma simple_function_real_of_nat[intro, simp]: + fixes f g :: "'a \ nat" assumes sf: "simple_function M f" + shows "simple_function M (\x. real (f x))" + by (auto intro!: simple_function_compose1[OF sf]) + +lemma borel_measurable_implies_simple_function_sequence: + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" + shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ + (\x. (SUP i. f i x) = max 0 (u x)) \ (\i x. 0 \ f i x)" +proof - + def f \ "\x i. if real i \ u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)" + { fix x j have "f x j \ j * 2 ^ j" unfolding f_def + proof (split split_if, intro conjI impI) + assume "\ real j \ u x" + then have "natfloor (real (u x) * 2 ^ j) \ natfloor (j * 2 ^ j)" + by (cases "u x") (auto intro!: natfloor_mono) + moreover have "real (natfloor (j * 2 ^ j)) \ j * 2^j" + by (intro real_natfloor_le) auto + ultimately show "natfloor (real (u x) * 2 ^ j) \ j * 2 ^ j" + unfolding real_of_nat_le_iff by auto + qed auto } + note f_upper = this + + have real_f: + "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))" + unfolding f_def by auto + + let ?g = "\j x. real (f x j) / 2^j :: ereal" + show ?thesis + proof (intro exI[of _ ?g] conjI allI ballI) + fix i + have "simple_function M (\x. real (f x i))" + proof (intro simple_function_borel_measurable) + show "(\x. real (f x i)) \ borel_measurable M" + using u by (auto simp: real_f) + have "(\x. real (f x i))`space M \ real`{..i*2^i}" + using f_upper[of _ i] by auto + then show "finite ((\x. real (f x i))`space M)" + by (rule finite_subset) auto + qed + then show "simple_function M (?g i)" + by (auto intro: simple_function_ereal simple_function_div) + next + show "incseq ?g" + proof (intro incseq_ereal incseq_SucI le_funI) + fix x and i :: nat + have "f x i * 2 \ f x (Suc i)" unfolding f_def + proof ((split split_if)+, intro conjI impI) + assume "ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" + then show "i * 2 ^ i * 2 \ natfloor (real (u x) * 2 ^ Suc i)" + by (cases "u x") (auto intro!: le_natfloor) + next + assume "\ ereal (real i) \ u x" "ereal (real (Suc i)) \ u x" + then show "natfloor (real (u x) * 2 ^ i) * 2 \ Suc i * 2 ^ Suc i" + by (cases "u x") auto + next + assume "\ ereal (real i) \ u x" "\ ereal (real (Suc i)) \ u x" + have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2" + by simp + also have "\ \ natfloor (real (u x) * 2 ^ i * 2)" + proof cases + assume "0 \ u x" then show ?thesis + by (intro le_mult_natfloor) + next + assume "\ 0 \ u x" then show ?thesis + by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg) + qed + also have "\ = natfloor (real (u x) * 2 ^ Suc i)" + by (simp add: ac_simps) + finally show "natfloor (real (u x) * 2 ^ i) * 2 \ natfloor (real (u x) * 2 ^ Suc i)" . + qed simp + then show "?g i x \ ?g (Suc i) x" + by (auto simp: field_simps) + qed + next + fix x show "(SUP i. ?g i x) = max 0 (u x)" + proof (rule SUP_eqI) + fix i show "?g i x \ max 0 (u x)" unfolding max_def f_def + by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg + mult_nonpos_nonneg) + next + fix y assume *: "\i. i \ UNIV \ ?g i x \ y" + have "\i. 0 \ ?g i x" by auto + from order_trans[OF this *] have "0 \ y" by simp + show "max 0 (u x) \ y" + proof (cases y) + case (real r) + with * have *: "\i. f x i \ r * 2^i" by (auto simp: divide_le_eq) + from reals_Archimedean2[of r] * have "u x \ \" by (auto simp: f_def) (metis less_le_not_le) + then have "\p. max 0 (u x) = ereal p \ 0 \ p" by (cases "u x") (auto simp: max_def) + then guess p .. note ux = this + obtain m :: nat where m: "p < real m" using reals_Archimedean2 .. + have "p \ r" + proof (rule ccontr) + assume "\ p \ r" + with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] + obtain N where "\n\N. r * 2^n < p * 2^n - 1" by (auto simp: field_simps) + then have "r * 2^max N m < p * 2^max N m - 1" by simp + moreover + have "real (natfloor (p * 2 ^ max N m)) \ r * 2 ^ max N m" + using *[of "max N m"] m unfolding real_f using ux + by (cases "0 \ u x") (simp_all add: max_def split: split_if_asm) + then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" + by (metis real_natfloor_gt_diff_one less_le_trans) + ultimately show False by auto + qed + then show "max 0 (u x) \ y" using real ux by simp + qed (insert `0 \ y`, auto) + qed + qed auto +qed + +lemma borel_measurable_implies_simple_function_sequence': + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" + obtains f where "\i. simple_function M (f i)" "incseq f" "\i. \ \ range (f i)" + "\x. (SUP i. f i x) = max 0 (u x)" "\i x. 0 \ f i x" + using borel_measurable_implies_simple_function_sequence[OF u] by auto + +lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]: + fixes u :: "'a \ ereal" + assumes u: "simple_function M u" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. P u \ P (\x. c * u x)" + assumes add: "\u v. P u \ P v \ P (\x. v x + u x)" + shows "P u" +proof (rule cong) + from AE_space show "AE x in M. (\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" + proof eventually_elim + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. + qed +next + from u have "finite (u ` space M)" + unfolding simple_function_def by auto + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" + proof induct + case empty show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + qed (auto intro!: add mult set simple_functionD u) +next + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" + apply (subst simple_function_cong) + apply (rule simple_function_indicator_representation[symmetric]) + apply (auto intro: u) + done +qed fact + +lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: + fixes u :: "'a \ ereal" + assumes u: "simple_function M u" and nn: "\x. 0 \ u x" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult: "\u c. 0 \ c \ simple_function M u \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" + assumes add: "\u v. simple_function M u \ (\x. 0 \ u x) \ P u \ simple_function M v \ (\x. 0 \ v x) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" + shows "P u" +proof - + show ?thesis + proof (rule cong) + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. + next + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" + apply (subst simple_function_cong) + apply (rule simple_function_indicator_representation[symmetric]) + apply (auto intro: u) + done + next + + from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" + unfolding simple_function_def by auto + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" + proof induct + case empty show ?case + using set[of "{}"] by (simp add: indicator_def[abs_def]) + next + case (insert x S) + { fix z have "(\y\S. y * indicator (u -` {y} \ space M) z) = 0 \ + x * indicator (u -` {x} \ space M) z = 0" + using insert by (subst setsum_ereal_0) (auto simp: indicator_def) } + note disj = this + from insert show ?case + by (auto intro!: add mult set simple_functionD u setsum_nonneg simple_function_setsum disj) + qed + qed fact +qed + +lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: + fixes u :: "'a \ ereal" + assumes u: "u \ borel_measurable M" "\x. 0 \ u x" + assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" + assumes set: "\A. A \ sets M \ P (indicator A)" + assumes mult': "\u c. 0 \ c \ c < \ \ u \ borel_measurable M \ (\x. 0 \ u x) \ (\x. x \ space M \ u x < \) \ P u \ P (\x. c * u x)" + assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ (\x. x \ space M \ u x < \) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ (\x. x \ space M \ v x < \) \ (\x. x \ space M \ u x = 0 \ v x = 0) \ P v \ P (\x. v x + u x)" + assumes seq: "\U. (\i. U i \ borel_measurable M) \ (\i x. 0 \ U i x) \ (\i x. x \ space M \ U i x < \) \ (\i. P (U i)) \ incseq U \ u = (SUP i. U i) \ P (SUP i. U i)" + shows "P u" + using u +proof (induct rule: borel_measurable_implies_simple_function_sequence') + fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and + sup: "\x. (SUP i. U i x) = max 0 (u x)" and nn: "\i x. 0 \ U i x" + have u_eq: "u = (SUP i. U i)" + using nn u sup by (auto simp: max_def) + + have not_inf: "\x i. x \ space M \ U i x < \" + using U by (auto simp: image_iff eq_commute) + + from U have "\i. U i \ borel_measurable M" + by (simp add: borel_measurable_simple_function) + + show "P u" + unfolding u_eq + proof (rule seq) + fix i show "P (U i)" + using `simple_function M (U i)` nn[of i] not_inf[of _ i] + proof (induct rule: simple_function_induct_nn) + case (mult u c) + show ?case + proof cases + assume "c = 0 \ space M = {} \ (\x\space M. u x = 0)" + with mult(2) show ?thesis + by (intro cong[of "\x. c * u x" "indicator {}"] set) + (auto dest!: borel_measurable_simple_function) + next + assume "\ (c = 0 \ space M = {} \ (\x\space M. u x = 0))" + with mult obtain x where u_fin: "\x. x \ space M \ u x < \" + and x: "x \ space M" "u x \ 0" "c \ 0" + by auto + with mult have "P u" + by auto + from x mult(5)[OF `x \ space M`] mult(1) mult(3)[of x] have "c < \" + by auto + with u_fin mult + show ?thesis + by (intro mult') (auto dest!: borel_measurable_simple_function) + qed + qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) + qed fact+ +qed + +lemma simple_function_If_set: + assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" + shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") +proof - + def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" + show ?thesis unfolding simple_function_def + proof safe + have "?IF ` space M \ f ` space M \ g ` space M" by auto + from finite_subset[OF this] assms + show "finite (?IF ` space M)" unfolding simple_function_def by auto + next + fix x assume "x \ space M" + then have *: "?IF -` {?IF x} \ space M = (if x \ A + then ((F (f x) \ (A \ space M)) \ (G (f x) - (G (f x) \ (A \ space M)))) + else ((F (g x) \ (A \ space M)) \ (G (g x) - (G (g x) \ (A \ space M)))))" + using sets.sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) + have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" + unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto + show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto + qed +qed + +lemma simple_function_If: + assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" + shows "simple_function M (\x. if P x then f x else g x)" +proof - + have "{x\space M. P x} = {x. P x} \ space M" by auto + with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp +qed + +lemma simple_function_subalgebra: + assumes "simple_function N f" + and N_subalgebra: "sets N \ sets M" "space N = space M" + shows "simple_function M f" + using assms unfolding simple_function_def by auto + +lemma simple_function_comp: + assumes T: "T \ measurable M M'" + and f: "simple_function M' f" + shows "simple_function M (\x. f (T x))" +proof (intro simple_function_def[THEN iffD2] conjI ballI) + have "(\x. f (T x)) ` space M \ f ` space M'" + using T unfolding measurable_def by auto + then show "finite ((\x. f (T x)) ` space M)" + using f unfolding simple_function_def by (auto intro: finite_subset) + fix i assume i: "i \ (\x. f (T x)) ` space M" + then have "i \ f ` space M'" + using T unfolding measurable_def by auto + then have "f -` {i} \ space M' \ sets M'" + using f unfolding simple_function_def by auto + then have "T -` (f -` {i} \ space M') \ space M \ sets M" + using T unfolding measurable_def by auto + also have "T -` (f -` {i} \ space M') \ space M = (\x. f (T x)) -` {i} \ space M" + using T unfolding measurable_def by auto + finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . +qed + +section "Simple integral" + +definition simple_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>S") where + "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" + +syntax + "_simple_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>S _. _ \_" [60,61] 110) + +translations + "\\<^sup>S x. f \M" == "CONST simple_integral M (%x. f)" + +lemma simple_integral_cong: + assumes "\t. t \ space M \ f t = g t" + shows "integral\<^sup>S M f = integral\<^sup>S M g" +proof - + have "f ` space M = g ` space M" + "\x. f -` {x} \ space M = g -` {x} \ space M" + using assms by (auto intro!: image_eqI) + thus ?thesis unfolding simple_integral_def by simp +qed + +lemma simple_integral_const[simp]: + "(\\<^sup>Sx. c \M) = c * (emeasure M) (space M)" +proof (cases "space M = {}") + case True thus ?thesis unfolding simple_integral_def by simp +next + case False hence "(\x. c) ` space M = {c}" by auto + thus ?thesis unfolding simple_integral_def by simp +qed + +lemma simple_function_partition: + assumes f: "simple_function M f" and g: "simple_function M g" + assumes sub: "\x y. x \ space M \ y \ space M \ g x = g y \ f x = f y" + assumes v: "\x. x \ space M \ f x = v (g x)" + shows "integral\<^sup>S M f = (\y\g ` space M. v y * emeasure M {x\space M. g x = y})" + (is "_ = ?r") +proof - + from f g have [simp]: "finite (f`space M)" "finite (g`space M)" + by (auto simp: simple_function_def) + from f g have [measurable]: "f \ measurable M (count_space UNIV)" "g \ measurable M (count_space UNIV)" + by (auto intro: measurable_simple_function) + + { fix y assume "y \ space M" + then have "f ` space M \ {i. \x\space M. i = f x \ g y = g x} = {v (g y)}" + by (auto cong: sub simp: v[symmetric]) } + note eq = this + + have "integral\<^sup>S M f = + (\y\f`space M. y * (\z\g`space M. + if \x\space M. y = f x \ z = g x then emeasure M {x\space M. g x = z} else 0))" + unfolding simple_integral_def + proof (safe intro!: setsum_cong ereal_left_mult_cong) + fix y assume y: "y \ space M" "f y \ 0" + have [simp]: "g ` space M \ {z. \x\space M. f y = f x \ z = g x} = + {z. \x\space M. f y = f x \ z = g x}" + by auto + have eq:"(\i\{z. \x\space M. f y = f x \ z = g x}. {x \ space M. g x = i}) = + f -` {f y} \ space M" + by (auto simp: eq_commute cong: sub rev_conj_cong) + have "finite (g`space M)" by simp + then have "finite {z. \x\space M. f y = f x \ z = g x}" + by (rule rev_finite_subset) auto + then show "emeasure M (f -` {f y} \ space M) = + (\z\g ` space M. if \x\space M. f y = f x \ z = g x then emeasure M {x \ space M. g x = z} else 0)" + apply (simp add: setsum_cases) + apply (subst setsum_emeasure) + apply (auto simp: disjoint_family_on_def eq) + done + qed + also have "\ = (\y\f`space M. (\z\g`space M. + if \x\space M. y = f x \ z = g x then y * emeasure M {x\space M. g x = z} else 0))" + by (auto intro!: setsum_cong simp: setsum_ereal_right_distrib emeasure_nonneg) + also have "\ = ?r" + by (subst setsum_commute) + (auto intro!: setsum_cong simp: setsum_cases scaleR_setsum_right[symmetric] eq) + finally show "integral\<^sup>S M f = ?r" . +qed + +lemma simple_integral_add[simp]: + assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" + shows "(\\<^sup>Sx. f x + g x \M) = integral\<^sup>S M f + integral\<^sup>S M g" +proof - + have "(\\<^sup>Sx. f x + g x \M) = + (\y\(\x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\space M. (f x, g x) = y})" + by (intro simple_function_partition) (auto intro: f g) + also have "\ = (\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) + + (\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y})" + using assms(2,4) by (auto intro!: setsum_cong ereal_left_distrib simp: setsum_addf[symmetric]) + also have "(\y\(\x. (f x, g x))`space M. fst y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. f x \M)" + by (intro simple_function_partition[symmetric]) (auto intro: f g) + also have "(\y\(\x. (f x, g x))`space M. snd y * emeasure M {x\space M. (f x, g x) = y}) = (\\<^sup>Sx. g x \M)" + by (intro simple_function_partition[symmetric]) (auto intro: f g) + finally show ?thesis . +qed + +lemma simple_integral_setsum[simp]: + assumes "\i x. i \ P \ 0 \ f i x" + assumes "\i. i \ P \ simple_function M (f i)" + shows "(\\<^sup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^sup>S M (f i))" +proof cases + assume "finite P" + from this assms show ?thesis + by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) +qed auto + +lemma simple_integral_mult[simp]: + assumes f: "simple_function M f" "\x. 0 \ f x" + shows "(\\<^sup>Sx. c * f x \M) = c * integral\<^sup>S M f" +proof - + have "(\\<^sup>Sx. c * f x \M) = (\y\f ` space M. (c * y) * emeasure M {x\space M. f x = y})" + using f by (intro simple_function_partition) auto + also have "\ = c * integral\<^sup>S M f" + using f unfolding simple_integral_def + by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult_assoc Int_def conj_commute) + finally show ?thesis . +qed + +lemma simple_integral_mono_AE: + assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" + and mono: "AE x in M. f x \ g x" + shows "integral\<^sup>S M f \ integral\<^sup>S M g" +proof - + let ?\ = "\P. emeasure M {x\space M. P x}" + have "integral\<^sup>S M f = (\y\(\x. (f x, g x))`space M. fst y * ?\ (\x. (f x, g x) = y))" + using f g by (intro simple_function_partition) auto + also have "\ \ (\y\(\x. (f x, g x))`space M. snd y * ?\ (\x. (f x, g x) = y))" + proof (clarsimp intro!: setsum_mono) + fix x assume "x \ space M" + let ?M = "?\ (\y. f y = f x \ g y = g x)" + show "f x * ?M \ g x * ?M" + proof cases + assume "?M \ 0" + then have "0 < ?M" + by (simp add: less_le emeasure_nonneg) + also have "\ \ ?\ (\y. f x \ g x)" + using mono by (intro emeasure_mono_AE) auto + finally have "\ \ f x \ g x" + by (intro notI) auto + then show ?thesis + by (intro ereal_mult_right_mono) auto + qed simp + qed + also have "\ = integral\<^sup>S M g" + using f g by (intro simple_function_partition[symmetric]) auto + finally show ?thesis . +qed + +lemma simple_integral_mono: + assumes "simple_function M f" and "simple_function M g" + and mono: "\ x. x \ space M \ f x \ g x" + shows "integral\<^sup>S M f \ integral\<^sup>S M g" + using assms by (intro simple_integral_mono_AE) auto + +lemma simple_integral_cong_AE: + assumes "simple_function M f" and "simple_function M g" + and "AE x in M. f x = g x" + shows "integral\<^sup>S M f = integral\<^sup>S M g" + using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) + +lemma simple_integral_cong': + assumes sf: "simple_function M f" "simple_function M g" + and mea: "(emeasure M) {x\space M. f x \ g x} = 0" + shows "integral\<^sup>S M f = integral\<^sup>S M g" +proof (intro simple_integral_cong_AE sf AE_I) + show "(emeasure M) {x\space M. f x \ g x} = 0" by fact + show "{x \ space M. f x \ g x} \ sets M" + using sf[THEN borel_measurable_simple_function] by auto +qed simp + +lemma simple_integral_indicator: + assumes A: "A \ sets M" + assumes f: "simple_function M f" + shows "(\\<^sup>Sx. f x * indicator A x \M) = + (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" +proof - + have eq: "(\x. (f x, indicator A x)) ` space M \ {x. snd x = 1} = (\x. (f x, 1::ereal))`A" + using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: split_if_asm) + have eq2: "\x. f x \ f ` A \ f -` {f x} \ space M \ A = {}" + by (auto simp: image_iff) + + have "(\\<^sup>Sx. f x * indicator A x \M) = + (\y\(\x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\space M. (f x, indicator A x) = y})" + using assms by (intro simple_function_partition) auto + also have "\ = (\y\(\x. (f x, indicator A x::ereal))`space M. + if snd y = 1 then fst y * emeasure M (f -` {fst y} \ space M \ A) else 0)" + by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum_cong) + also have "\ = (\y\(\x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \ space M \ A))" + using assms by (subst setsum_cases) (auto intro!: simple_functionD(1) simp: eq) + also have "\ = (\y\fst`(\x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \ space M \ A))" + by (subst setsum_reindex[where f=fst]) (auto simp: inj_on_def) + also have "\ = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M \ A))" + using A[THEN sets.sets_into_space] + by (intro setsum_mono_zero_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) + finally show ?thesis . +qed + +lemma simple_integral_indicator_only[simp]: + assumes "A \ sets M" + shows "integral\<^sup>S M (indicator A) = emeasure M A" + using simple_integral_indicator[OF assms, of "\x. 1"] sets.sets_into_space[OF assms] + by (simp_all add: image_constant_conv Int_absorb1 split: split_if_asm) + +lemma simple_integral_null_set: + assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets M" + shows "(\\<^sup>Sx. u x * indicator N x \M) = 0" +proof - + have "AE x in M. indicator N x = (0 :: ereal)" + using `N \ null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) + then have "(\\<^sup>Sx. u x * indicator N x \M) = (\\<^sup>Sx. 0 \M)" + using assms apply (intro simple_integral_cong_AE) by auto + then show ?thesis by simp +qed + +lemma simple_integral_cong_AE_mult_indicator: + assumes sf: "simple_function M f" and eq: "AE x in M. x \ S" and "S \ sets M" + shows "integral\<^sup>S M f = (\\<^sup>Sx. f x * indicator S x \M)" + using assms by (intro simple_integral_cong_AE) auto + +lemma simple_integral_cmult_indicator: + assumes A: "A \ sets M" + shows "(\\<^sup>Sx. c * indicator A x \M) = c * emeasure M A" + using simple_integral_mult[OF simple_function_indicator[OF A]] + unfolding simple_integral_indicator_only[OF A] by simp + +lemma simple_integral_positive: + assumes f: "simple_function M f" and ae: "AE x in M. 0 \ f x" + shows "0 \ integral\<^sup>S M f" +proof - + have "integral\<^sup>S M (\x. 0) \ integral\<^sup>S M f" + using simple_integral_mono_AE[OF _ f ae] by auto + then show ?thesis by simp +qed + +section "Continuous positive integration" + +definition positive_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>P") where + "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^sup>S M g)" + +syntax + "_positive_integral" :: "pttrn \ ereal \ 'a measure \ ereal" ("\\<^sup>+ _. _ \_" [60,61] 110) + +translations + "\\<^sup>+ x. f \M" == "CONST positive_integral M (%x. f)" + +lemma positive_integral_positive: + "0 \ integral\<^sup>P M f" + by (auto intro!: SUP_upper2[of "\x. 0"] simp: positive_integral_def le_fun_def) + +lemma positive_integral_not_MInfty[simp]: "integral\<^sup>P M f \ -\" + using positive_integral_positive[of M f] by auto + +lemma positive_integral_def_finite: + "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f \ range g \ {0 ..< \}}. integral\<^sup>S M g)" + (is "_ = SUPREMUM ?A ?f") + unfolding positive_integral_def +proof (safe intro!: antisym SUP_least) + fix g assume g: "simple_function M g" "g \ max 0 \ f" + let ?G = "{x \ space M. \ g x \ \}" + note gM = g(1)[THEN borel_measurable_simple_function] + have \_G_pos: "0 \ (emeasure M) ?G" using gM by auto + let ?g = "\y x. if g x = \ then y else max 0 (g x)" + from g gM have g_in_A: "\y. 0 \ y \ y \ \ \ ?g y \ ?A" + apply (safe intro!: simple_function_max simple_function_If) + apply (force simp: max_def le_fun_def split: split_if_asm)+ + done + show "integral\<^sup>S M g \ SUPREMUM ?A ?f" + proof cases + have g0: "?g 0 \ ?A" by (intro g_in_A) auto + assume "(emeasure M) ?G = 0" + with gM have "AE x in M. x \ ?G" + by (auto simp add: AE_iff_null intro!: null_setsI) + with gM g show ?thesis + by (intro SUP_upper2[OF g0] simple_integral_mono_AE) + (auto simp: max_def intro!: simple_function_If) + next + assume \_G: "(emeasure M) ?G \ 0" + have "SUPREMUM ?A (integral\<^sup>S M) = \" + proof (intro SUP_PInfty) + fix n :: nat + let ?y = "ereal (real n) / (if (emeasure M) ?G = \ then 1 else (emeasure M) ?G)" + have "0 \ ?y" "?y \ \" using \_G \_G_pos by (auto simp: ereal_divide_eq) + then have "?g ?y \ ?A" by (rule g_in_A) + have "real n \ ?y * (emeasure M) ?G" + using \_G \_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) + also have "\ = (\\<^sup>Sx. ?y * indicator ?G x \M)" + using `0 \ ?y` `?g ?y \ ?A` gM + by (subst simple_integral_cmult_indicator) auto + also have "\ \ integral\<^sup>S M (?g ?y)" using `?g ?y \ ?A` gM + by (intro simple_integral_mono) auto + finally show "\i\?A. real n \ integral\<^sup>S M i" + using `?g ?y \ ?A` by blast + qed + then show ?thesis by simp + qed +qed (auto intro: SUP_upper) + +lemma positive_integral_mono_AE: + assumes ae: "AE x in M. u x \ v x" shows "integral\<^sup>P M u \ integral\<^sup>P M v" + unfolding positive_integral_def +proof (safe intro!: SUP_mono) + fix n assume n: "simple_function M n" "n \ max 0 \ u" + from ae[THEN AE_E] guess N . note N = this + then have ae_N: "AE x in M. x \ N" by (auto intro: AE_not_in) + let ?n = "\x. n x * indicator (space M - N) x" + have "AE x in M. n x \ ?n x" "simple_function M ?n" + using n N ae_N by auto + moreover + { fix x have "?n x \ max 0 (v x)" + proof cases + assume x: "x \ space M - N" + with N have "u x \ v x" by auto + with n(2)[THEN le_funD, of x] x show ?thesis + by (auto simp: max_def split: split_if_asm) + qed simp } + then have "?n \ max 0 \ v" by (auto simp: le_funI) + moreover have "integral\<^sup>S M n \ integral\<^sup>S M ?n" + using ae_N N n by (auto intro!: simple_integral_mono_AE) + ultimately show "\m\{g. simple_function M g \ g \ max 0 \ v}. integral\<^sup>S M n \ integral\<^sup>S M m" + by force +qed + +lemma positive_integral_mono: + "(\x. x \ space M \ u x \ v x) \ integral\<^sup>P M u \ integral\<^sup>P M v" + by (auto intro: positive_integral_mono_AE) + +lemma positive_integral_cong_AE: + "AE x in M. u x = v x \ integral\<^sup>P M u = integral\<^sup>P M v" + by (auto simp: eq_iff intro!: positive_integral_mono_AE) + +lemma positive_integral_cong: + "(\x. x \ space M \ u x = v x) \ integral\<^sup>P M u = integral\<^sup>P M v" + by (auto intro: positive_integral_cong_AE) + +lemma positive_integral_cong_strong: + "M = N \ (\x. x \ space M \ u x = v x) \ integral\<^sup>P M u = integral\<^sup>P N v" + by (auto intro: positive_integral_cong) + +lemma positive_integral_eq_simple_integral: + assumes f: "simple_function M f" "\x. 0 \ f x" shows "integral\<^sup>P M f = integral\<^sup>S M f" +proof - + let ?f = "\x. f x * indicator (space M) x" + have f': "simple_function M ?f" using f by auto + with f(2) have [simp]: "max 0 \ ?f = ?f" + by (auto simp: fun_eq_iff max_def split: split_indicator) + have "integral\<^sup>P M ?f \ integral\<^sup>S M ?f" using f' + by (force intro!: SUP_least simple_integral_mono simp: le_fun_def positive_integral_def) + moreover have "integral\<^sup>S M ?f \ integral\<^sup>P M ?f" + unfolding positive_integral_def + using f' by (auto intro!: SUP_upper) + ultimately show ?thesis + by (simp cong: positive_integral_cong simple_integral_cong) +qed + +lemma positive_integral_eq_simple_integral_AE: + assumes f: "simple_function M f" "AE x in M. 0 \ f x" shows "integral\<^sup>P M f = integral\<^sup>S M f" +proof - + have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) + with f have "integral\<^sup>P M f = integral\<^sup>S M (\x. max 0 (f x))" + by (simp cong: positive_integral_cong_AE simple_integral_cong_AE + add: positive_integral_eq_simple_integral) + with assms show ?thesis + by (auto intro!: simple_integral_cong_AE split: split_max) +qed + +lemma positive_integral_SUP_approx: + assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + and u: "simple_function M u" "u \ (SUP i. f i)" "u`space M \ {0..<\}" + shows "integral\<^sup>S M u \ (SUP i. integral\<^sup>P M (f i))" (is "_ \ ?S") +proof (rule ereal_le_mult_one_interval) + have "0 \ (SUP i. integral\<^sup>P M (f i))" + using f(3) by (auto intro!: SUP_upper2 positive_integral_positive) + then show "(SUP i. integral\<^sup>P M (f i)) \ -\" by auto + have u_range: "\x. x \ space M \ 0 \ u x \ u x \ \" + using u(3) by auto + fix a :: ereal assume "0 < a" "a < 1" + hence "a \ 0" by auto + let ?B = "\i. {x \ space M. a * u x \ f i x}" + have B: "\i. ?B i \ sets M" + using f `simple_function M u`[THEN borel_measurable_simple_function] by auto + + let ?uB = "\i x. u x * indicator (?B i) x" + + { fix i have "?B i \ ?B (Suc i)" + proof safe + fix i x assume "a * u x \ f i x" + also have "\ \ f (Suc i) x" + using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto + finally show "a * u x \ f (Suc i) x" . + qed } + note B_mono = this + + note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B] + + let ?B' = "\i n. (u -` {i} \ space M) \ ?B n" + have measure_conv: "\i. (emeasure M) (u -` {i} \ space M) = (SUP n. (emeasure M) (?B' i n))" + proof - + fix i + have 1: "range (?B' i) \ sets M" using B_u by auto + have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI) + have "(\n. ?B' i n) = u -` {i} \ space M" + proof safe + fix x i assume x: "x \ space M" + show "x \ (\i. ?B' (u x) i)" + proof cases + assume "u x = 0" thus ?thesis using `x \ space M` f(3) by simp + next + assume "u x \ 0" + with `a < 1` u_range[OF `x \ space M`] + have "a * u x < 1 * u x" + by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) + also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) + finally obtain i where "a * u x < f i x" unfolding SUP_def + by (auto simp add: less_SUP_iff) + hence "a * u x \ f i x" by auto + thus ?thesis using `x \ space M` by auto + qed + qed + then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp + qed + + have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" + unfolding simple_integral_indicator[OF B `simple_function M u`] + proof (subst SUP_ereal_setsum, safe) + fix x n assume "x \ space M" + with u_range show "incseq (\i. u x * (emeasure M) (?B' (u x) i))" "\i. 0 \ u x * (emeasure M) (?B' (u x) i)" + using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) + next + show "integral\<^sup>S M u = (\i\u ` space M. SUP n. i * (emeasure M) (?B' i n))" + using measure_conv u_range B_u unfolding simple_integral_def + by (auto intro!: setsum_cong SUP_ereal_cmult [symmetric]) + qed + moreover + have "a * (SUP i. integral\<^sup>S M (?uB i)) \ ?S" + apply (subst SUP_ereal_cmult [symmetric]) + proof (safe intro!: SUP_mono bexI) + fix i + have "a * integral\<^sup>S M (?uB i) = (\\<^sup>Sx. a * ?uB i x \M)" + using B `simple_function M u` u_range + by (subst simple_integral_mult) (auto split: split_indicator) + also have "\ \ integral\<^sup>P M (f i)" + proof - + have *: "simple_function M (\x. a * ?uB i x)" using B `0 < a` u(1) by auto + show ?thesis using f(3) * u_range `0 < a` + by (subst positive_integral_eq_simple_integral[symmetric]) + (auto intro!: positive_integral_mono split: split_indicator) + qed + finally show "a * integral\<^sup>S M (?uB i) \ integral\<^sup>P M (f i)" + by auto + next + fix i show "0 \ \\<^sup>S x. ?uB i x \M" using B `0 < a` u(1) u_range + by (intro simple_integral_positive) (auto split: split_indicator) + qed (insert `0 < a`, auto) + ultimately show "a * integral\<^sup>S M u \ ?S" by simp +qed + +lemma incseq_positive_integral: + assumes "incseq f" shows "incseq (\i. integral\<^sup>P M (f i))" +proof - + have "\i x. f i x \ f (Suc i) x" + using assms by (auto dest!: incseq_SucD simp: le_fun_def) + then show ?thesis + by (auto intro!: incseq_SucI positive_integral_mono) +qed + +text {* Beppo-Levi monotone convergence theorem *} +lemma positive_integral_monotone_convergence_SUP: + assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" +proof (rule antisym) + show "(SUP j. integral\<^sup>P M (f j)) \ (\\<^sup>+ x. (SUP i. f i x) \M)" + by (auto intro!: SUP_least SUP_upper positive_integral_mono) +next + show "(\\<^sup>+ x. (SUP i. f i x) \M) \ (SUP j. integral\<^sup>P M (f j))" + unfolding positive_integral_def_finite[of _ "\x. SUP i. f i x"] + proof (safe intro!: SUP_least) + fix g assume g: "simple_function M g" + and *: "g \ max 0 \ (\x. SUP i. f i x)" "range g \ {0..<\}" + then have "\x. 0 \ (SUP i. f i x)" and g': "g`space M \ {0..<\}" + using f by (auto intro!: SUP_upper2) + with * show "integral\<^sup>S M g \ (SUP j. integral\<^sup>P M (f j))" + by (intro positive_integral_SUP_approx[OF f g _ g']) + (auto simp: le_fun_def max_def) + qed +qed + +lemma positive_integral_monotone_convergence_SUP_AE: + assumes f: "\i. AE x in M. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" + shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" +proof - + from f have "AE x in M. \i. f i x \ f (Suc i) x \ 0 \ f i x" + by (simp add: AE_all_countable) + from this[THEN AE_E] guess N . note N = this + let ?f = "\i x. if x \ space M - N then f i x else 0" + have f_eq: "AE x in M. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) + then have "(\\<^sup>+ x. (SUP i. f i x) \M) = (\\<^sup>+ x. (SUP i. ?f i x) \M)" + by (auto intro!: positive_integral_cong_AE) + also have "\ = (SUP i. (\\<^sup>+ x. ?f i x \M))" + proof (rule positive_integral_monotone_convergence_SUP) + show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) + { fix i show "(\x. if x \ space M - N then f i x else 0) \ borel_measurable M" + using f N(3) by (intro measurable_If_set) auto + fix x show "0 \ ?f i x" + using N(1) by auto } + qed + also have "\ = (SUP i. (\\<^sup>+ x. f i x \M))" + using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] positive_integral_cong_AE ext) + finally show ?thesis . +qed + +lemma positive_integral_monotone_convergence_SUP_AE_incseq: + assumes f: "incseq f" "\i. AE x in M. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" + shows "(\\<^sup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^sup>P M (f i))" + using f[unfolded incseq_Suc_iff le_fun_def] + by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel]) + auto + +lemma positive_integral_monotone_convergence_simple: + assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" + shows "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" + using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1) + f(3)[THEN borel_measurable_simple_function] f(2)] + by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPREMUM UNIV"] ext) + +lemma positive_integral_max_0: + "(\\<^sup>+x. max 0 (f x) \M) = integral\<^sup>P M f" + by (simp add: le_fun_def positive_integral_def) + +lemma positive_integral_cong_pos: + assumes "\x. x \ space M \ f x \ 0 \ g x \ 0 \ f x = g x" + shows "integral\<^sup>P M f = integral\<^sup>P M g" +proof - + have "integral\<^sup>P M (\x. max 0 (f x)) = integral\<^sup>P M (\x. max 0 (g x))" + proof (intro positive_integral_cong) + fix x assume "x \ space M" + from assms[OF this] show "max 0 (f x) = max 0 (g x)" + by (auto split: split_max) + qed + then show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma SUP_simple_integral_sequences: + assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" + and g: "incseq g" "\i x. 0 \ g i x" "\i. simple_function M (g i)" + and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" + shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" + (is "SUPREMUM _ ?F = SUPREMUM _ ?G") +proof - + have "(SUP i. integral\<^sup>S M (f i)) = (\\<^sup>+x. (SUP i. f i x) \M)" + using f by (rule positive_integral_monotone_convergence_simple) + also have "\ = (\\<^sup>+x. (SUP i. g i x) \M)" + unfolding eq[THEN positive_integral_cong_AE] .. + also have "\ = (SUP i. ?G i)" + using g by (rule positive_integral_monotone_convergence_simple[symmetric]) + finally show ?thesis by simp +qed + +lemma positive_integral_const[simp]: + "0 \ c \ (\\<^sup>+ x. c \M) = c * (emeasure M) (space M)" + by (subst positive_integral_eq_simple_integral) auto + +lemma positive_integral_linear: + assumes f: "f \ borel_measurable M" "\x. 0 \ f x" and "0 \ a" + and g: "g \ borel_measurable M" "\x. 0 \ g x" + shows "(\\<^sup>+ x. a * f x + g x \M) = a * integral\<^sup>P M f + integral\<^sup>P M g" + (is "integral\<^sup>P M ?L = _") +proof - + from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . + note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . + note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + let ?L' = "\i x. a * u i x + v i x" + + have "?L \ borel_measurable M" using assms by auto + from borel_measurable_implies_simple_function_sequence'[OF this] guess l . + note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + + have inc: "incseq (\i. a * integral\<^sup>S M (u i))" "incseq (\i. integral\<^sup>S M (v i))" + using u v `0 \ a` + by (auto simp: incseq_Suc_iff le_fun_def + intro!: add_mono ereal_mult_left_mono simple_integral_mono) + have pos: "\i. 0 \ integral\<^sup>S M (u i)" "\i. 0 \ integral\<^sup>S M (v i)" "\i. 0 \ a * integral\<^sup>S M (u i)" + using u v `0 \ a` by (auto simp: simple_integral_positive) + { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \ -\" "integral\<^sup>S M (v i) \ -\" + by (auto split: split_if_asm) } + note not_MInf = this + + have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" + proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) + show "incseq ?L'" "\i x. 0 \ ?L' i x" "\i. simple_function M (?L' i)" + using u v `0 \ a` unfolding incseq_Suc_iff le_fun_def + by (auto intro!: add_mono ereal_mult_left_mono) + { fix x + { fix i have "a * u i x \ -\" "v i x \ -\" "u i x \ -\" using `0 \ a` u(6)[of i x] v(6)[of i x] + by auto } + then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" + using `0 \ a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] + by (subst SUP_ereal_cmult [symmetric, OF u(6) `0 \ a`]) + (auto intro!: SUP_ereal_add + simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } + then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" + unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) + by (intro AE_I2) (auto split: split_max) + qed + also have "\ = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" + using u(2, 6) v(2, 6) `0 \ a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) + finally have "(\\<^sup>+ x. max 0 (a * f x + g x) \M) = a * (\\<^sup>+x. max 0 (f x) \M) + (\\<^sup>+x. max 0 (g x) \M)" + unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] + unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] + apply (subst SUP_ereal_cmult [symmetric, OF pos(1) `0 \ a`]) + apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) . + then show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma positive_integral_cmult: + assumes f: "f \ borel_measurable M" "0 \ c" + shows "(\\<^sup>+ x. c * f x \M) = c * integral\<^sup>P M f" +proof - + have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` + by (auto split: split_max simp: ereal_zero_le_0_iff) + have "(\\<^sup>+ x. c * f x \M) = (\\<^sup>+ x. c * max 0 (f x) \M)" + by (simp add: positive_integral_max_0) + then show ?thesis + using positive_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" _ "\x. 0"] f + by (auto simp: positive_integral_max_0) +qed + +lemma positive_integral_multc: + assumes "f \ borel_measurable M" "0 \ c" + shows "(\\<^sup>+ x. f x * c \M) = integral\<^sup>P M f * c" + unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp + +lemma positive_integral_indicator[simp]: + "A \ sets M \ (\\<^sup>+ x. indicator A x\M) = (emeasure M) A" + by (subst positive_integral_eq_simple_integral) + (auto simp: simple_integral_indicator) + +lemma positive_integral_cmult_indicator: + "0 \ c \ A \ sets M \ (\\<^sup>+ x. c * indicator A x \M) = c * (emeasure M) A" + by (subst positive_integral_eq_simple_integral) + (auto simp: simple_function_indicator simple_integral_indicator) + +lemma positive_integral_indicator': + assumes [measurable]: "A \ space M \ sets M" + shows "(\\<^sup>+ x. indicator A x \M) = emeasure M (A \ space M)" +proof - + have "(\\<^sup>+ x. indicator A x \M) = (\\<^sup>+ x. indicator (A \ space M) x \M)" + by (intro positive_integral_cong) (simp split: split_indicator) + also have "\ = emeasure M (A \ space M)" + by simp + finally show ?thesis . +qed + +lemma positive_integral_add: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + shows "(\\<^sup>+ x. f x + g x \M) = integral\<^sup>P M f + integral\<^sup>P M g" +proof - + have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" + using assms by (auto split: split_max) + have "(\\<^sup>+ x. f x + g x \M) = (\\<^sup>+ x. max 0 (f x + g x) \M)" + by (simp add: positive_integral_max_0) + also have "\ = (\\<^sup>+ x. max 0 (f x) + max 0 (g x) \M)" + unfolding ae[THEN positive_integral_cong_AE] .. + also have "\ = (\\<^sup>+ x. max 0 (f x) \M) + (\\<^sup>+ x. max 0 (g x) \M)" + using positive_integral_linear[of "\x. max 0 (f x)" _ 1 "\x. max 0 (g x)"] f g + by auto + finally show ?thesis + by (simp add: positive_integral_max_0) +qed + +lemma positive_integral_setsum: + assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x in M. 0 \ f i x" + shows "(\\<^sup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^sup>P M (f i))" +proof cases + assume f: "finite P" + from assms have "AE x in M. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto + from f this assms(1) show ?thesis + proof induct + case (insert i P) + then have "f i \ borel_measurable M" "AE x in M. 0 \ f i x" + "(\x. \i\P. f i x) \ borel_measurable M" "AE x in M. 0 \ (\i\P. f i x)" + by (auto intro!: setsum_nonneg) + from positive_integral_add[OF this] + show ?case using insert by auto + qed simp +qed simp + +lemma positive_integral_Markov_inequality: + assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" and "A \ sets M" and c: "0 \ c" + shows "(emeasure M) ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^sup>+ x. u x * indicator A x \M)" + (is "(emeasure M) ?A \ _ * ?PI") +proof - + have "?A \ sets M" + using `A \ sets M` u by auto + hence "(emeasure M) ?A = (\\<^sup>+ x. indicator ?A x \M)" + using positive_integral_indicator by simp + also have "\ \ (\\<^sup>+ x. c * (u x * indicator A x) \M)" using u c + by (auto intro!: positive_integral_mono_AE + simp: indicator_def ereal_zero_le_0_iff) + also have "\ = c * (\\<^sup>+ x. u x * indicator A x \M)" + using assms + by (auto intro!: positive_integral_cmult simp: ereal_zero_le_0_iff) + finally show ?thesis . +qed + +lemma positive_integral_noteq_infinite: + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + and "integral\<^sup>P M g \ \" + shows "AE x in M. g x \ \" +proof (rule ccontr) + assume c: "\ (AE x in M. g x \ \)" + have "(emeasure M) {x\space M. g x = \} \ 0" + using c g by (auto simp add: AE_iff_null) + moreover have "0 \ (emeasure M) {x\space M. g x = \}" using g by (auto intro: measurable_sets) + ultimately have "0 < (emeasure M) {x\space M. g x = \}" by auto + then have "\ = \ * (emeasure M) {x\space M. g x = \}" by auto + also have "\ \ (\\<^sup>+x. \ * indicator {x\space M. g x = \} x \M)" + using g by (subst positive_integral_cmult_indicator) auto + also have "\ \ integral\<^sup>P M g" + using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def) + finally show False using `integral\<^sup>P M g \ \` by auto +qed + +lemma positive_integral_PInf: + assumes f: "f \ borel_measurable M" + and not_Inf: "integral\<^sup>P M f \ \" + shows "(emeasure M) (f -` {\} \ space M) = 0" +proof - + have "\ * (emeasure M) (f -` {\} \ space M) = (\\<^sup>+ x. \ * indicator (f -` {\} \ space M) x \M)" + using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets) + also have "\ \ integral\<^sup>P M (\x. max 0 (f x))" + by (auto intro!: positive_integral_mono simp: indicator_def max_def) + finally have "\ * (emeasure M) (f -` {\} \ space M) \ integral\<^sup>P M f" + by (simp add: positive_integral_max_0) + moreover have "0 \ (emeasure M) (f -` {\} \ space M)" + by (rule emeasure_nonneg) + ultimately show ?thesis + using assms by (auto split: split_if_asm) +qed + +lemma positive_integral_PInf_AE: + assumes "f \ borel_measurable M" "integral\<^sup>P M f \ \" shows "AE x in M. f x \ \" +proof (rule AE_I) + show "(emeasure M) (f -` {\} \ space M) = 0" + by (rule positive_integral_PInf[OF assms]) + show "f -` {\} \ space M \ sets M" + using assms by (auto intro: borel_measurable_vimage) +qed auto + +lemma simple_integral_PInf: + assumes "simple_function M f" "\x. 0 \ f x" + and "integral\<^sup>S M f \ \" + shows "(emeasure M) (f -` {\} \ space M) = 0" +proof (rule positive_integral_PInf) + show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) + show "integral\<^sup>P M f \ \" + using assms by (simp add: positive_integral_eq_simple_integral) +qed + +lemma positive_integral_diff: + assumes f: "f \ borel_measurable M" + and g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + and fin: "integral\<^sup>P M g \ \" + and mono: "AE x in M. g x \ f x" + shows "(\\<^sup>+ x. f x - g x \M) = integral\<^sup>P M f - integral\<^sup>P M g" +proof - + have diff: "(\x. f x - g x) \ borel_measurable M" "AE x in M. 0 \ f x - g x" + using assms by (auto intro: ereal_diff_positive) + have pos_f: "AE x in M. 0 \ f x" using mono g by auto + { fix a b :: ereal assume "0 \ a" "a \ \" "0 \ b" "a \ b" then have "b - a + a = b" + by (cases rule: ereal2_cases[of a b]) auto } + note * = this + then have "AE x in M. f x = f x - g x + g x" + using mono positive_integral_noteq_infinite[OF g fin] assms by auto + then have **: "integral\<^sup>P M f = (\\<^sup>+x. f x - g x \M) + integral\<^sup>P M g" + unfolding positive_integral_add[OF diff g, symmetric] + by (rule positive_integral_cong_AE) + show ?thesis unfolding ** + using fin positive_integral_positive[of M g] + by (cases rule: ereal2_cases[of "\\<^sup>+ x. f x - g x \M" "integral\<^sup>P M g"]) auto +qed + +lemma positive_integral_suminf: + assumes f: "\i. f i \ borel_measurable M" "\i. AE x in M. 0 \ f i x" + shows "(\\<^sup>+ x. (\i. f i x) \M) = (\i. integral\<^sup>P M (f i))" +proof - + have all_pos: "AE x in M. \i. 0 \ f i x" + using assms by (auto simp: AE_all_countable) + have "(\i. integral\<^sup>P M (f i)) = (SUP n. \iP M (f i))" + using positive_integral_positive by (rule suminf_ereal_eq_SUP) + also have "\ = (SUP n. \\<^sup>+x. (\iM)" + unfolding positive_integral_setsum[OF f] .. + also have "\ = \\<^sup>+x. (SUP n. \iM" using f all_pos + by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) + (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3) + also have "\ = \\<^sup>+x. (\i. f i x) \M" using all_pos + by (intro positive_integral_cong_AE) (auto simp: suminf_ereal_eq_SUP) + finally show ?thesis by simp +qed + +lemma positive_integral_mult_bounded_inf: + assumes f: "f \ borel_measurable M" "(\\<^sup>+x. f x \M) < \" + and c: "0 \ c" "c \ \" and ae: "AE x in M. g x \ c * f x" + shows "(\\<^sup>+x. g x \M) < \" +proof - + have "(\\<^sup>+x. g x \M) \ (\\<^sup>+x. c * f x \M)" + by (intro positive_integral_mono_AE ae) + also have "(\\<^sup>+x. c * f x \M) < \" + using c f by (subst positive_integral_cmult) auto + finally show ?thesis . +qed + +text {* Fatou's lemma: convergence theorem on limes inferior *} + +lemma positive_integral_liminf: + fixes u :: "nat \ 'a \ ereal" + assumes u: "\i. u i \ borel_measurable M" "\i. AE x in M. 0 \ u i x" + shows "(\\<^sup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>P M (u n))" +proof - + have pos: "AE x in M. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) + have "(\\<^sup>+ x. liminf (\n. u n x) \M) = + (SUP n. \\<^sup>+ x. (INF i:{n..}. u i x) \M)" + unfolding liminf_SUP_INF using pos u + by (intro positive_integral_monotone_convergence_SUP_AE) + (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) + also have "\ \ liminf (\n. integral\<^sup>P M (u n))" + unfolding liminf_SUP_INF + by (auto intro!: SUP_mono exI INF_greatest positive_integral_mono INF_lower) + finally show ?thesis . +qed + +lemma le_Limsup: + "F \ bot \ eventually (\x. c \ g x) F \ c \ Limsup F g" + using Limsup_mono[of "\_. c" g F] by (simp add: Limsup_const) + +lemma Limsup_le: + "F \ bot \ eventually (\x. f x \ c) F \ Limsup F f \ c" + using Limsup_mono[of f "\_. c" F] by (simp add: Limsup_const) + +lemma ereal_mono_minus_cancel: + fixes a b c :: ereal + shows "c - a \ c - b \ 0 \ c \ c < \ \ b \ a" + by (cases a b c rule: ereal3_cases) auto + +lemma positive_integral_limsup: + fixes u :: "nat \ 'a \ ereal" + assumes [measurable]: "\i. u i \ borel_measurable M" "w \ borel_measurable M" + assumes bounds: "\i. AE x in M. 0 \ u i x" "\i. AE x in M. u i x \ w x" and w: "(\\<^sup>+x. w x \M) < \" + shows "limsup (\n. integral\<^sup>P M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" +proof - + have bnd: "AE x in M. \i. 0 \ u i x \ u i x \ w x" + using bounds by (auto simp: AE_all_countable) + + from bounds[of 0] have w_nonneg: "AE x in M. 0 \ w x" + by auto + + have "(\\<^sup>+x. w x \M) - (\\<^sup>+x. limsup (\n. u n x) \M) = (\\<^sup>+x. w x - limsup (\n. u n x) \M)" + proof (intro positive_integral_diff[symmetric]) + show "AE x in M. 0 \ limsup (\n. u n x)" + using bnd by (auto intro!: le_Limsup) + show "AE x in M. limsup (\n. u n x) \ w x" + using bnd by (auto intro!: Limsup_le) + then have "(\\<^sup>+x. limsup (\n. u n x) \M) < \" + by (intro positive_integral_mult_bounded_inf[OF _ w, of 1]) auto + then show "(\\<^sup>+x. limsup (\n. u n x) \M) \ \" + by simp + qed auto + also have "\ = (\\<^sup>+x. liminf (\n. w x - u n x) \M)" + using w_nonneg + by (intro positive_integral_cong_AE, eventually_elim) + (auto intro!: liminf_ereal_cminus[symmetric]) + also have "\ \ liminf (\n. \\<^sup>+x. w x - u n x \M)" + proof (rule positive_integral_liminf) + fix i show "AE x in M. 0 \ w x - u i x" + using bounds[of i] by eventually_elim (auto intro: ereal_diff_positive) + qed simp + also have "(\n. \\<^sup>+x. w x - u n x \M) = (\n. (\\<^sup>+x. w x \M) - (\\<^sup>+x. u n x \M))" + proof (intro ext positive_integral_diff) + fix i have "(\\<^sup>+x. u i x \M) < \" + using bounds by (intro positive_integral_mult_bounded_inf[OF _ w, of 1]) auto + then show "(\\<^sup>+x. u i x \M) \ \" by simp + qed (insert bounds, auto) + also have "liminf (\n. (\\<^sup>+x. w x \M) - (\\<^sup>+x. u n x \M)) = (\\<^sup>+x. w x \M) - limsup (\n. \\<^sup>+x. u n x \M)" + using w by (intro liminf_ereal_cminus) auto + finally show ?thesis + by (rule ereal_mono_minus_cancel) (intro w positive_integral_positive)+ +qed + +lemma positive_integral_dominated_convergence: + assumes [measurable]: + "\i. u i \ borel_measurable M" "u' \ borel_measurable M" "w \ borel_measurable M" + and bound: "\j. AE x in M. 0 \ u j x" "\j. AE x in M. u j x \ w x" + and w: "(\\<^sup>+x. w x \M) < \" + and u': "AE x in M. (\i. u i x) ----> u' x" + shows "(\i. (\\<^sup>+x. u i x \M)) ----> (\\<^sup>+x. u' x \M)" +proof - + have "limsup (\n. integral\<^sup>P M (u n)) \ (\\<^sup>+ x. limsup (\n. u n x) \M)" + by (intro positive_integral_limsup[OF _ _ bound w]) auto + moreover have "(\\<^sup>+ x. limsup (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" + using u' by (intro positive_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) + moreover have "(\\<^sup>+ x. liminf (\n. u n x) \M) = (\\<^sup>+ x. u' x \M)" + using u' by (intro positive_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) + moreover have "(\\<^sup>+x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^sup>P M (u n))" + by (intro positive_integral_liminf[OF _ bound(1)]) auto + moreover have "liminf (\n. integral\<^sup>P M (u n)) \ limsup (\n. integral\<^sup>P M (u n))" + by (intro Liminf_le_Limsup sequentially_bot) + ultimately show ?thesis + by (intro Liminf_eq_Limsup) auto +qed + +lemma positive_integral_null_set: + assumes "N \ null_sets M" shows "(\\<^sup>+ x. u x * indicator N x \M) = 0" +proof - + have "(\\<^sup>+ x. u x * indicator N x \M) = (\\<^sup>+ x. 0 \M)" + proof (intro positive_integral_cong_AE AE_I) + show "{x \ space M. u x * indicator N x \ 0} \ N" + by (auto simp: indicator_def) + show "(emeasure M) N = 0" "N \ sets M" + using assms by auto + qed + then show ?thesis by simp +qed + +lemma positive_integral_0_iff: + assumes u: "u \ borel_measurable M" and pos: "AE x in M. 0 \ u x" + shows "integral\<^sup>P M u = 0 \ emeasure M {x\space M. u x \ 0} = 0" + (is "_ \ (emeasure M) ?A = 0") +proof - + have u_eq: "(\\<^sup>+ x. u x * indicator ?A x \M) = integral\<^sup>P M u" + by (auto intro!: positive_integral_cong simp: indicator_def) + show ?thesis + proof + assume "(emeasure M) ?A = 0" + with positive_integral_null_set[of ?A M u] u + show "integral\<^sup>P M u = 0" by (simp add: u_eq null_sets_def) + next + { fix r :: ereal and n :: nat assume gt_1: "1 \ real n * r" + then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) + then have "0 \ r" by (auto simp add: ereal_zero_less_0_iff) } + note gt_1 = this + assume *: "integral\<^sup>P M u = 0" + let ?M = "\n. {x \ space M. 1 \ real (n::nat) * u x}" + have "0 = (SUP n. (emeasure M) (?M n \ ?A))" + proof - + { fix n :: nat + from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] + have "(emeasure M) (?M n \ ?A) \ 0" unfolding u_eq * using u by simp + moreover have "0 \ (emeasure M) (?M n \ ?A)" using u by auto + ultimately have "(emeasure M) (?M n \ ?A) = 0" by auto } + thus ?thesis by simp + qed + also have "\ = (emeasure M) (\n. ?M n \ ?A)" + proof (safe intro!: SUP_emeasure_incseq) + fix n show "?M n \ ?A \ sets M" + using u by (auto intro!: sets.Int) + next + show "incseq (\n. {x \ space M. 1 \ real n * u x} \ {x \ space M. u x \ 0})" + proof (safe intro!: incseq_SucI) + fix n :: nat and x + assume *: "1 \ real n * u x" + also from gt_1[OF *] have "real n * u x \ real (Suc n) * u x" + using `0 \ u x` by (auto intro!: ereal_mult_right_mono) + finally show "1 \ real (Suc n) * u x" by auto + qed + qed + also have "\ = (emeasure M) {x\space M. 0 < u x}" + proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1) + fix x assume "0 < u x" and [simp, intro]: "x \ space M" + show "x \ (\n. ?M n \ ?A)" + proof (cases "u x") + case (real r) with `0 < u x` have "0 < r" by auto + obtain j :: nat where "1 / r \ real j" using real_arch_simple .. + hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using `0 < r` by auto + hence "1 \ real j * r" using real `0 < r` by auto + thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) + qed (insert `0 < u x`, auto) + qed auto + finally have "(emeasure M) {x\space M. 0 < u x} = 0" by simp + moreover + from pos have "AE x in M. \ (u x < 0)" by auto + then have "(emeasure M) {x\space M. u x < 0} = 0" + using AE_iff_null[of M] u by auto + moreover have "(emeasure M) {x\space M. u x \ 0} = (emeasure M) {x\space M. u x < 0} + (emeasure M) {x\space M. 0 < u x}" + using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) + ultimately show "(emeasure M) ?A = 0" by simp + qed +qed + +lemma positive_integral_0_iff_AE: + assumes u: "u \ borel_measurable M" + shows "integral\<^sup>P M u = 0 \ (AE x in M. u x \ 0)" +proof - + have sets: "{x\space M. max 0 (u x) \ 0} \ sets M" + using u by auto + from positive_integral_0_iff[of "\x. max 0 (u x)"] + have "integral\<^sup>P M u = 0 \ (AE x in M. max 0 (u x) = 0)" + unfolding positive_integral_max_0 + using AE_iff_null[OF sets] u by auto + also have "\ \ (AE x in M. u x \ 0)" by (auto split: split_max) + finally show ?thesis . +qed + +lemma AE_iff_positive_integral: + "{x\space M. P x} \ sets M \ (AE x in M. P x) \ integral\<^sup>P M (indicator {x. \ P x}) = 0" + by (subst positive_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def + sets.sets_Collect_neg indicator_def[abs_def] measurable_If) + +lemma positive_integral_const_If: + "(\\<^sup>+x. a \M) = (if 0 \ a then a * (emeasure M) (space M) else 0)" + by (auto intro!: positive_integral_0_iff_AE[THEN iffD2]) + +lemma positive_integral_subalgebra: + assumes f: "f \ borel_measurable N" "\x. 0 \ f x" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" + shows "integral\<^sup>P N f = integral\<^sup>P M f" +proof - + have [simp]: "\f :: 'a \ ereal. f \ borel_measurable N \ f \ borel_measurable M" + using N by (auto simp: measurable_def) + have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" + using N by (auto simp add: eventually_ae_filter null_sets_def) + have [simp]: "\A. A \ sets N \ A \ sets M" + using N by auto + from f show ?thesis + apply induct + apply (simp_all add: positive_integral_add positive_integral_cmult positive_integral_monotone_convergence_SUP N) + apply (auto intro!: positive_integral_cong cong: positive_integral_cong simp: N(2)[symmetric]) + done +qed + +lemma positive_integral_nat_function: + fixes f :: "'a \ nat" + assumes "f \ measurable M (count_space UNIV)" + shows "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\t. emeasure M {x\space M. t < f x})" +proof - + def F \ "\i. {x\space M. i < f x}" + with assms have [measurable]: "\i. F i \ sets M" + by auto + + { fix x assume "x \ space M" + have "(\i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" + using sums_If_finite[of "\i. i < f x" "\_. 1::real"] by simp + then have "(\i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))" + unfolding sums_ereal . + moreover have "\i. ereal (if i < f x then 1 else 0) = indicator (F i) x" + using `x \ space M` by (simp add: one_ereal_def F_def) + ultimately have "ereal(of_nat(f x)) = (\i. indicator (F i) x)" + by (simp add: sums_iff) } + then have "(\\<^sup>+x. ereal (of_nat (f x)) \M) = (\\<^sup>+x. (\i. indicator (F i) x) \M)" + by (simp cong: positive_integral_cong) + also have "\ = (\i. emeasure M (F i))" + by (simp add: positive_integral_suminf) + finally show ?thesis + by (simp add: F_def) +qed + +section {* Distributions *} + +lemma positive_integral_distr': + assumes T: "T \ measurable M M'" + and f: "f \ borel_measurable (distr M M' T)" "\x. 0 \ f x" + shows "integral\<^sup>P (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" + using f +proof induct + case (cong f g) + with T show ?case + apply (subst positive_integral_cong[of _ f g]) + apply simp + apply (subst positive_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) + apply (simp add: measurable_def Pi_iff) + apply simp + done +next + case (set A) + then have eq: "\x. x \ space M \ indicator A (T x) = indicator (T -` A \ space M) x" + by (auto simp: indicator_def) + from set T show ?case + by (subst positive_integral_cong[OF eq]) + (auto simp add: emeasure_distr intro!: positive_integral_indicator[symmetric] measurable_sets) +qed (simp_all add: measurable_compose[OF T] T positive_integral_cmult positive_integral_add + positive_integral_monotone_convergence_SUP le_fun_def incseq_def) + +lemma positive_integral_distr: + "T \ measurable M M' \ f \ borel_measurable M' \ integral\<^sup>P (distr M M' T) f = (\\<^sup>+ x. f (T x) \M)" + by (subst (1 2) positive_integral_max_0[symmetric]) + (simp add: positive_integral_distr') + +section {* Lebesgue integration on @{const count_space} *} + +lemma simple_function_count_space[simp]: + "simple_function (count_space A) f \ finite (f ` A)" + unfolding simple_function_def by simp + +lemma positive_integral_count_space: + assumes A: "finite {a\A. 0 < f a}" + shows "integral\<^sup>P (count_space A) f = (\a|a\A \ 0 < f a. f a)" +proof - + have *: "(\\<^sup>+x. max 0 (f x) \count_space A) = + (\\<^sup>+ x. (\a|a\A \ 0 < f a. f a * indicator {a} x) \count_space A)" + by (auto intro!: positive_integral_cong + simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less) + also have "\ = (\a|a\A \ 0 < f a. \\<^sup>+ x. f a * indicator {a} x \count_space A)" + by (subst positive_integral_setsum) + (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) + also have "\ = (\a|a\A \ 0 < f a. f a)" + by (auto intro!: setsum_cong simp: positive_integral_cmult_indicator one_ereal_def[symmetric]) + finally show ?thesis by (simp add: positive_integral_max_0) +qed + +lemma positive_integral_count_space_finite: + "finite A \ (\\<^sup>+x. f x \count_space A) = (\a\A. max 0 (f a))" + by (subst positive_integral_max_0[symmetric]) + (auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le) + +lemma emeasure_UN_countable: + assumes sets: "\i. i \ I \ X i \ sets M" and I: "countable I" + assumes disj: "disjoint_family_on X I" + shows "emeasure M (UNION I X) = (\\<^sup>+i. emeasure M (X i) \count_space I)" +proof cases + assume "finite I" with sets disj show ?thesis + by (subst setsum_emeasure[symmetric]) + (auto intro!: setsum_cong simp add: max_def subset_eq positive_integral_count_space_finite emeasure_nonneg) +next + assume f: "\ finite I" + then have [intro]: "I \ {}" by auto + from from_nat_into_inj_infinite[OF I f] from_nat_into[OF this] disj + have disj2: "disjoint_family (\i. X (from_nat_into I i))" + unfolding disjoint_family_on_def by metis + + from f have "bij_betw (from_nat_into I) UNIV I" + using bij_betw_from_nat_into[OF I] by simp + then have "(\i\I. X i) = (\i. (X \ from_nat_into I) i)" + unfolding SUP_def image_comp [symmetric] by (simp add: bij_betw_def) + then have "emeasure M (UNION I X) = emeasure M (\i. X (from_nat_into I i))" + by simp + also have "\ = (\i. emeasure M (X (from_nat_into I i)))" + by (intro suminf_emeasure[symmetric] disj disj2) (auto intro!: sets from_nat_into[OF `I \ {}`]) + also have "\ = (\n. \\<^sup>+i. emeasure M (X i) * indicator {from_nat_into I n} i \count_space I)" + proof (intro arg_cong[where f=suminf] ext) + fix i + have eq: "{a \ I. 0 < emeasure M (X a) * indicator {from_nat_into I i} a} + = (if 0 < emeasure M (X (from_nat_into I i)) then {from_nat_into I i} else {})" + using ereal_0_less_1 + by (auto simp: ereal_zero_less_0_iff indicator_def from_nat_into `I \ {}` simp del: ereal_0_less_1) + have "(\\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I) = + (if 0 < emeasure M (X (from_nat_into I i)) then emeasure M (X (from_nat_into I i)) else 0)" + by (subst positive_integral_count_space) (simp_all add: eq) + also have "\ = emeasure M (X (from_nat_into I i))" + by (simp add: less_le emeasure_nonneg) + finally show "emeasure M (X (from_nat_into I i)) = + \\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \count_space I" .. + qed + also have "\ = (\\<^sup>+i. emeasure M (X i) \count_space I)" + apply (subst positive_integral_suminf[symmetric]) + apply (auto simp: emeasure_nonneg intro!: positive_integral_cong) + proof - + fix x assume "x \ I" + then have "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = (\i\{to_nat_on I x}. emeasure M (X x) * indicator {from_nat_into I i} x)" + by (intro suminf_finite) (auto simp: indicator_def I f) + also have "\ = emeasure M (X x)" + by (simp add: I f `x\I`) + finally show "(\i. emeasure M (X x) * indicator {from_nat_into I i} x) = emeasure M (X x)" . + qed + finally show ?thesis . +qed + +section {* Measures with Restricted Space *} + +lemma positive_integral_restrict_space: + assumes \: "\ \ sets M" and f: "f \ borel_measurable M" "\x. 0 \ f x" "\x. x \ space M - \ \ f x = 0" + shows "positive_integral (restrict_space M \) f = positive_integral M f" +using f proof (induct rule: borel_measurable_induct) + case (cong f g) then show ?case + using positive_integral_cong[of M f g] positive_integral_cong[of "restrict_space M \" f g] + sets.sets_into_space[OF `\ \ sets M`] + by (simp add: subset_eq space_restrict_space) +next + case (set A) + then have "A \ \" + unfolding indicator_eq_0_iff by (auto dest: sets.sets_into_space) + with set `\ \ sets M` sets.sets_into_space[OF `\ \ sets M`] show ?case + by (subst positive_integral_indicator') + (auto simp add: sets_restrict_space_iff space_restrict_space + emeasure_restrict_space Int_absorb2 + dest: sets.sets_into_space) +next + case (mult f c) then show ?case + by (cases "c = 0") (simp_all add: measurable_restrict_space1 \ positive_integral_cmult) +next + case (add f g) then show ?case + by (simp add: measurable_restrict_space1 \ positive_integral_add ereal_add_nonneg_eq_0_iff) +next + case (seq F) then show ?case + by (auto simp add: SUP_eq_iff measurable_restrict_space1 \ positive_integral_monotone_convergence_SUP) +qed + +section {* Measure spaces with an associated density *} + +definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where + "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" + +lemma + shows sets_density[simp]: "sets (density M f) = sets M" + and space_density[simp]: "space (density M f) = space M" + by (auto simp: density_def) + +(* FIXME: add conversion to simplify space, sets and measurable *) +lemma space_density_imp[measurable_dest]: + "\x M f. x \ space (density M f) \ x \ space M" by auto + +lemma + shows measurable_density_eq1[simp]: "g \ measurable (density Mg f) Mg' \ g \ measurable Mg Mg'" + and measurable_density_eq2[simp]: "h \ measurable Mh (density Mh' f) \ h \ measurable Mh Mh'" + and simple_function_density_eq[simp]: "simple_function (density Mu f) u \ simple_function Mu u" + unfolding measurable_def simple_function_def by simp_all + +lemma density_cong: "f \ borel_measurable M \ f' \ borel_measurable M \ + (AE x in M. f x = f' x) \ density M f = density M f'" + unfolding density_def by (auto intro!: measure_of_eq positive_integral_cong_AE sets.space_closed) + +lemma density_max_0: "density M f = density M (\x. max 0 (f x))" +proof - + have "\x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" + by (auto simp: indicator_def) + then show ?thesis + unfolding density_def by (simp add: positive_integral_max_0) +qed + +lemma density_ereal_max_0: "density M (\x. ereal (f x)) = density M (\x. ereal (max 0 (f x)))" + by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) + +lemma emeasure_density: + assumes f[measurable]: "f \ borel_measurable M" and A[measurable]: "A \ sets M" + shows "emeasure (density M f) A = (\\<^sup>+ x. f x * indicator A x \M)" + (is "_ = ?\ A") + unfolding density_def +proof (rule emeasure_measure_of_sigma) + show "sigma_algebra (space M) (sets M)" .. + show "positive (sets M) ?\" + using f by (auto simp: positive_def intro!: positive_integral_positive) + have \_eq: "?\ = (\A. \\<^sup>+ x. max 0 (f x) * indicator A x \M)" (is "?\ = ?\'") + apply (subst positive_integral_max_0[symmetric]) + apply (intro ext positive_integral_cong_AE AE_I2) + apply (auto simp: indicator_def) + done + show "countably_additive (sets M) ?\" + unfolding \_eq + proof (intro countably_additiveI) + fix A :: "nat \ 'a set" assume "range A \ sets M" + then have "\i. A i \ sets M" by auto + then have *: "\i. (\x. max 0 (f x) * indicator (A i) x) \ borel_measurable M" + by (auto simp: set_eq_iff) + assume disj: "disjoint_family A" + have "(\n. ?\' (A n)) = (\\<^sup>+ x. (\n. max 0 (f x) * indicator (A n) x) \M)" + using f * by (simp add: positive_integral_suminf) + also have "\ = (\\<^sup>+ x. max 0 (f x) * (\n. indicator (A n) x) \M)" using f + by (auto intro!: suminf_cmult_ereal positive_integral_cong_AE) + also have "\ = (\\<^sup>+ x. max 0 (f x) * indicator (\n. A n) x \M)" + unfolding suminf_indicator[OF disj] .. + finally show "(\n. ?\' (A n)) = ?\' (\x. A x)" by simp + qed +qed fact + +lemma null_sets_density_iff: + assumes f: "f \ borel_measurable M" + shows "A \ null_sets (density M f) \ A \ sets M \ (AE x in M. x \ A \ f x \ 0)" +proof - + { assume "A \ sets M" + have eq: "(\\<^sup>+x. f x * indicator A x \M) = (\\<^sup>+x. max 0 (f x) * indicator A x \M)" + apply (subst positive_integral_max_0[symmetric]) + apply (intro positive_integral_cong) + apply (auto simp: indicator_def) + done + have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ + emeasure M {x \ space M. max 0 (f x) * indicator A x \ 0} = 0" + unfolding eq + using f `A \ sets M` + by (intro positive_integral_0_iff) auto + also have "\ \ (AE x in M. max 0 (f x) * indicator A x = 0)" + using f `A \ sets M` + by (intro AE_iff_measurable[OF _ refl, symmetric]) auto + also have "(AE x in M. max 0 (f x) * indicator A x = 0) \ (AE x in M. x \ A \ f x \ 0)" + by (auto simp add: indicator_def max_def split: split_if_asm) + finally have "(\\<^sup>+x. f x * indicator A x \M) = 0 \ (AE x in M. x \ A \ f x \ 0)" . } + with f show ?thesis + by (simp add: null_sets_def emeasure_density cong: conj_cong) +qed + +lemma AE_density: + assumes f: "f \ borel_measurable M" + shows "(AE x in density M f. P x) \ (AE x in M. 0 < f x \ P x)" +proof + assume "AE x in density M f. P x" + with f obtain N where "{x \ space M. \ P x} \ N" "N \ sets M" and ae: "AE x in M. x \ N \ f x \ 0" + by (auto simp: eventually_ae_filter null_sets_density_iff) + then have "AE x in M. x \ N \ P x" by auto + with ae show "AE x in M. 0 < f x \ P x" + by (rule eventually_elim2) auto +next + fix N assume ae: "AE x in M. 0 < f x \ P x" + then obtain N where "{x \ space M. \ (0 < f x \ P x)} \ N" "N \ null_sets M" + by (auto simp: eventually_ae_filter) + then have *: "{x \ space (density M f). \ P x} \ N \ {x\space M. \ 0 < f x}" + "N \ {x\space M. \ 0 < f x} \ sets M" and ae2: "AE x in M. x \ N" + using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in) + show "AE x in density M f. P x" + using ae2 + unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] + by (intro exI[of _ "N \ {x\space M. \ 0 < f x}"] conjI *) + (auto elim: eventually_elim2) +qed + +lemma positive_integral_density': + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g: "g \ borel_measurable M" "\x. 0 \ g x" + shows "integral\<^sup>P (density M f) g = (\\<^sup>+ x. f x * g x \M)" +using g proof induct + case (cong u v) + then show ?case + apply (subst positive_integral_cong[OF cong(3)]) + apply (simp_all cong: positive_integral_cong) + done +next + case (set A) then show ?case + by (simp add: emeasure_density f) +next + case (mult u c) + moreover have "\x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) + ultimately show ?case + using f by (simp add: positive_integral_cmult) +next + case (add u v) + then have "\x. f x * (v x + u x) = f x * v x + f x * u x" + by (simp add: ereal_right_distrib) + with add f show ?case + by (auto simp add: positive_integral_add ereal_zero_le_0_iff intro!: positive_integral_add[symmetric]) +next + case (seq U) + from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" + by eventually_elim (simp add: SUP_ereal_cmult seq) + from seq f show ?case + apply (simp add: positive_integral_monotone_convergence_SUP) + apply (subst positive_integral_cong_AE[OF eq]) + apply (subst positive_integral_monotone_convergence_SUP_AE) + apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) + done +qed + +lemma positive_integral_density: + "f \ borel_measurable M \ AE x in M. 0 \ f x \ g' \ borel_measurable M \ + integral\<^sup>P (density M f) g' = (\\<^sup>+ x. f x * g' x \M)" + by (subst (1 2) positive_integral_max_0[symmetric]) + (auto intro!: positive_integral_cong_AE + simp: measurable_If max_def ereal_zero_le_0_iff positive_integral_density') + +lemma emeasure_restricted: + assumes S: "S \ sets M" and X: "X \ sets M" + shows "emeasure (density M (indicator S)) X = emeasure M (S \ X)" +proof - + have "emeasure (density M (indicator S)) X = (\\<^sup>+x. indicator S x * indicator X x \M)" + using S X by (simp add: emeasure_density) + also have "\ = (\\<^sup>+x. indicator (S \ X) x \M)" + by (auto intro!: positive_integral_cong simp: indicator_def) + also have "\ = emeasure M (S \ X)" + using S X by (simp add: sets.Int) + finally show ?thesis . +qed + +lemma measure_restricted: + "S \ sets M \ X \ sets M \ measure (density M (indicator S)) X = measure M (S \ X)" + by (simp add: emeasure_restricted measure_def) + +lemma (in finite_measure) finite_measure_restricted: + "S \ sets M \ finite_measure (density M (indicator S))" + by default (simp add: emeasure_restricted) + +lemma emeasure_density_const: + "A \ sets M \ 0 \ c \ emeasure (density M (\_. c)) A = c * emeasure M A" + by (auto simp: positive_integral_cmult_indicator emeasure_density) + +lemma measure_density_const: + "A \ sets M \ 0 < c \ c \ \ \ measure (density M (\_. c)) A = real c * measure M A" + by (auto simp: emeasure_density_const measure_def) + +lemma density_density_eq: + "f \ borel_measurable M \ g \ borel_measurable M \ AE x in M. 0 \ f x \ + density (density M f) g = density M (\x. f x * g x)" + by (auto intro!: measure_eqI simp: emeasure_density positive_integral_density ac_simps) + +lemma distr_density_distr: + assumes T: "T \ measurable M M'" and T': "T' \ measurable M' M" + and inv: "\x\space M. T' (T x) = x" + assumes f: "f \ borel_measurable M'" + shows "distr (density (distr M M' T) f) M T' = density M (f \ T)" (is "?R = ?L") +proof (rule measure_eqI) + fix A assume A: "A \ sets ?R" + { fix x assume "x \ space M" + with sets.sets_into_space[OF A] + have "indicator (T' -` A \ space M') (T x) = (indicator A x :: ereal)" + using T inv by (auto simp: indicator_def measurable_space) } + with A T T' f show "emeasure ?R A = emeasure ?L A" + by (simp add: measurable_comp emeasure_density emeasure_distr + positive_integral_distr measurable_sets cong: positive_integral_cong) +qed simp + +lemma density_density_divide: + fixes f g :: "'a \ real" + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + assumes ac: "AE x in M. f x = 0 \ g x = 0" + shows "density (density M f) (\x. g x / f x) = density M g" +proof - + have "density M g = density M (\x. f x * (g x / f x))" + using f g ac by (auto intro!: density_cong measurable_If) + then show ?thesis + using f g by (subst density_density_eq) auto +qed + +section {* Point measure *} + +definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where + "point_measure A f = density (count_space A) f" + +lemma + shows space_point_measure: "space (point_measure A f) = A" + and sets_point_measure: "sets (point_measure A f) = Pow A" + by (auto simp: point_measure_def) + +lemma measurable_point_measure_eq1[simp]: + "g \ measurable (point_measure A f) M \ g \ A \ space M" + unfolding point_measure_def by simp + +lemma measurable_point_measure_eq2_finite[simp]: + "finite A \ + g \ measurable M (point_measure A f) \ + (g \ space M \ A \ (\a\A. g -` {a} \ space M \ sets M))" + unfolding point_measure_def by (simp add: measurable_count_space_eq2) + +lemma simple_function_point_measure[simp]: + "simple_function (point_measure A f) g \ finite (g ` A)" + by (simp add: point_measure_def) + +lemma emeasure_point_measure: + assumes A: "finite {a\X. 0 < f a}" "X \ A" + shows "emeasure (point_measure A f) X = (\a|a\X \ 0 < f a. f a)" +proof - + have "{a. (a \ X \ a \ A \ 0 < f a) \ a \ X} = {a\X. 0 < f a}" + using `X \ A` by auto + with A show ?thesis + by (simp add: emeasure_density positive_integral_count_space ereal_zero_le_0_iff + point_measure_def indicator_def) +qed + +lemma emeasure_point_measure_finite: + "finite A \ (\i. i \ A \ 0 \ f i) \ X \ A \ emeasure (point_measure A f) X = (\a\X. f a)" + by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) + +lemma emeasure_point_measure_finite2: + "X \ A \ finite X \ (\i. i \ X \ 0 \ f i) \ emeasure (point_measure A f) X = (\a\X. f a)" + by (subst emeasure_point_measure) + (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le) + +lemma null_sets_point_measure_iff: + "X \ null_sets (point_measure A f) \ X \ A \ (\x\X. f x \ 0)" + by (auto simp: AE_count_space null_sets_density_iff point_measure_def) + +lemma AE_point_measure: + "(AE x in point_measure A f. P x) \ (\x\A. 0 < f x \ P x)" + unfolding point_measure_def + by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) + +lemma positive_integral_point_measure: + "finite {a\A. 0 < f a \ 0 < g a} \ + integral\<^sup>P (point_measure A f) g = (\a|a\A \ 0 < f a \ 0 < g a. f a * g a)" + unfolding point_measure_def + apply (subst density_max_0) + apply (subst positive_integral_density) + apply (simp_all add: AE_count_space positive_integral_density) + apply (subst positive_integral_count_space ) + apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff) + apply (rule finite_subset) + prefer 2 + apply assumption + apply auto + done + +lemma positive_integral_point_measure_finite: + "finite A \ (\a. a \ A \ 0 \ f a) \ (\a. a \ A \ 0 \ g a) \ + integral\<^sup>P (point_measure A f) g = (\a\A. f a * g a)" + by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le) + +section {* Uniform measure *} + +definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" + +lemma + shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M" + and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" + by (auto simp: uniform_measure_def) + +lemma emeasure_uniform_measure[simp]: + assumes A: "A \ sets M" and B: "B \ sets M" + shows "emeasure (uniform_measure M A) B = emeasure M (A \ B) / emeasure M A" +proof - + from A B have "emeasure (uniform_measure M A) B = (\\<^sup>+x. (1 / emeasure M A) * indicator (A \ B) x \M)" + by (auto simp add: uniform_measure_def emeasure_density split: split_indicator + intro!: positive_integral_cong) + also have "\ = emeasure M (A \ B) / emeasure M A" + using A B + by (subst positive_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg) + finally show ?thesis . +qed + +lemma measure_uniform_measure[simp]: + assumes A: "emeasure M A \ 0" "emeasure M A \ \" and B: "B \ sets M" + shows "measure (uniform_measure M A) B = measure M (A \ B) / measure M A" + using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A + by (cases "emeasure M A" "emeasure M (A \ B)" rule: ereal2_cases) (simp_all add: measure_def) + +section {* Uniform count measure *} + +definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" + +lemma + shows space_uniform_count_measure: "space (uniform_count_measure A) = A" + and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" + unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) + +lemma emeasure_uniform_count_measure: + "finite A \ X \ A \ emeasure (uniform_count_measure A) X = card X / card A" + by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) + +lemma measure_uniform_count_measure: + "finite A \ X \ A \ measure (uniform_count_measure A) X = card X / card A" + by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) + +end diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Probability_Measure.thy Mon May 19 12:04:45 2014 +0200 @@ -112,6 +112,7 @@ qed lemma (in prob_space) expectation_less: + fixes X :: "_ \ real" assumes [simp]: "integrable M X" assumes gt: "AE x in M. X x < b" shows "expectation X < b" @@ -123,6 +124,7 @@ qed lemma (in prob_space) expectation_greater: + fixes X :: "_ \ real" assumes [simp]: "integrable M X" assumes gt: "AE x in M. a < X x" shows "a < expectation X" @@ -134,7 +136,7 @@ qed lemma (in prob_space) jensens_inequality: - fixes a b :: real + fixes q :: "real \ real" assumes X: "integrable M X" "AE x in M. X x \ I" assumes I: "I = {a <..< b} \ I = {a <..} \ I = {..< b} \ I = UNIV" assumes q: "integrable M (\x. q (X x))" "convex_on I q" @@ -173,8 +175,8 @@ using prob_space by (simp add: X) also have "\ \ expectation (\w. q (X w))" using `x \ I` `open I` X(2) - apply (intro integral_mono_AE integral_add integral_cmult integral_diff - lebesgue_integral_const X q) + apply (intro integral_mono_AE integrable_add integrable_mult_right integrable_diff + integrable_const X q) apply (elim eventually_elim1) apply (intro convex_le_Inf_differential) apply (auto simp: interior_open q) @@ -452,7 +454,7 @@ lemma distributed_integral: "distributed M N X f \ g \ borel_measurable N \ (\x. f x * g x \N) = (\x. g (X x) \M)" by (auto simp: distributed_real_AE - distributed_distr_eq_density[symmetric] integral_density[symmetric] integral_distr) + distributed_distr_eq_density[symmetric] integral_real_density[symmetric] integral_distr) lemma distributed_transform_integral: assumes Px: "distributed M N X Px" @@ -520,7 +522,7 @@ also have "\ = (\\<^sup>+x. (\\<^sup>+y. (f (x, y) * indicator B y) * indicator A x \T) \S)" using f by (auto simp add: eq positive_integral_multc intro!: positive_integral_cong) also have "\ = emeasure ?R E" - by (auto simp add: emeasure_density T.positive_integral_fst_measurable(2)[symmetric] + by (auto simp add: emeasure_density T.positive_integral_fst[symmetric] intro!: positive_integral_cong split: split_indicator) finally show "emeasure ?L E = emeasure ?R E" . qed @@ -588,7 +590,7 @@ show X: "X \ measurable M S" by simp show borel: "Px \ borel_measurable S" - by (auto intro!: T.positive_integral_fst_measurable simp: Px_def) + by (auto intro!: T.positive_integral_fst simp: Px_def) interpret Pxy: prob_space "distr M (S \\<^sub>M T) (\x. (X x, Y x))" by (intro prob_space_distr) simp @@ -606,7 +608,7 @@ using Pxy by (simp add: distributed_def) also have "\ = \\<^sup>+ x. \\<^sup>+ y. Pxy (x, y) * indicator (A \ space T) (x, y) \T \S" using A borel Pxy - by (simp add: emeasure_density T.positive_integral_fst_measurable(2)[symmetric]) + by (simp add: emeasure_density T.positive_integral_fst[symmetric]) also have "\ = \\<^sup>+ x. Px x * indicator A x \S" apply (rule positive_integral_cong_AE) using Pxy[THEN distributed_AE, THEN ST.AE_pair] AE_space diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Mon May 19 12:04:45 2014 +0200 @@ -5,7 +5,7 @@ header {*Radon-Nikod{\'y}m derivative*} theory Radon_Nikodym -imports Lebesgue_Integration +imports Bochner_Integration begin definition "diff_measure M N = @@ -1062,35 +1062,53 @@ section "Radon-Nikodym derivative" -definition - "RN_deriv M N \ SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N" +definition RN_deriv :: "'a measure \ 'a measure \ 'a \ ereal" where + "RN_deriv M N = + (if \f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N + then SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N + else (\_. 0))" -lemma - assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" - shows borel_measurable_RN_deriv_density: "RN_deriv M (density M f) \ borel_measurable M" (is ?borel) - and density_RN_deriv_density: "density M (RN_deriv M (density M f)) = density M f" (is ?density) - and RN_deriv_density_nonneg: "0 \ RN_deriv M (density M f) x" (is ?pos) +lemma RN_derivI: + assumes "f \ borel_measurable M" "\x. 0 \ f x" "density M f = N" + shows "density M (RN_deriv M N) = N" proof - - let ?f = "\x. max 0 (f x)" - let ?P = "\g. g \ borel_measurable M \ (\x. 0 \ g x) \ density M g = density M f" - from f have "?P ?f" using f by (auto intro!: density_cong simp: split: split_max) - then have "?P (RN_deriv M (density M f))" - unfolding RN_deriv_def by (rule someI[where P="?P"]) - then show ?borel ?density ?pos by auto + have "\f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N" + using assms by auto + moreover then have "density M (SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N) = N" + by (rule someI2_ex) auto + ultimately show ?thesis + by (auto simp: RN_deriv_def) qed -lemma (in sigma_finite_measure) RN_deriv: - assumes "absolutely_continuous M N" "sets N = sets M" - shows borel_measurable_RN_deriv[measurable]: "RN_deriv M N \ borel_measurable M" (is ?borel) - and density_RN_deriv: "density M (RN_deriv M N) = N" (is ?density) - and RN_deriv_nonneg: "0 \ RN_deriv M N x" (is ?pos) +lemma + shows borel_measurable_RN_deriv[measurable]: "RN_deriv M N \ borel_measurable M" (is ?m) + and RN_deriv_nonneg: "0 \ RN_deriv M N x" (is ?nn) proof - - note Ex = Radon_Nikodym[OF assms, unfolded Bex_def] - from Ex show ?borel unfolding RN_deriv_def by (rule someI2_ex) simp - from Ex show ?density unfolding RN_deriv_def by (rule someI2_ex) simp - from Ex show ?pos unfolding RN_deriv_def by (rule someI2_ex) simp + { assume ex: "\f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N" + have 1: "(SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N) \ borel_measurable M" + using ex by (rule someI2_ex) auto + moreover + have 2: "0 \ (SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ density M f = N) x" + using ex by (rule someI2_ex) auto + note 1 2 } + from this show ?m ?nn + by (auto simp: RN_deriv_def) qed +lemma density_RN_deriv_density: + assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" + shows "density M (RN_deriv M (density M f)) = density M f" +proof (rule RN_derivI) + show "(\x. max 0 (f x)) \ borel_measurable M" "\x. 0 \ max 0 (f x)" + using f by auto + show "density M (\x. max 0 (f x)) = density M f" + using f by (intro density_cong) (auto simp: max_def) +qed + +lemma (in sigma_finite_measure) density_RN_deriv: + "absolutely_continuous M N \ sets N = sets M \ density M (RN_deriv M N) = N" + by (metis RN_derivI Radon_Nikodym) + lemma (in sigma_finite_measure) RN_deriv_positive_integral: assumes N: "absolutely_continuous M N" "sets N = sets M" and f: "f \ borel_measurable M" @@ -1099,7 +1117,7 @@ have "integral\<^sup>P N f = integral\<^sup>P (density M (RN_deriv M N)) f" using N by (simp add: density_RN_deriv) also have "\ = (\\<^sup>+x. RN_deriv M N x * f x \M)" - using RN_deriv(1,3)[OF N] f by (simp add: positive_integral_density) + using f by (simp add: positive_integral_density RN_deriv_nonneg) finally show ?thesis by simp qed @@ -1111,16 +1129,16 @@ and eq: "density M f = N" shows "AE x in M. f x = RN_deriv M N x" unfolding eq[symmetric] - by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv_density - RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric]) + by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv + RN_deriv_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric]) lemma RN_deriv_unique_sigma_finite: assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" and eq: "density M f = N" and fin: "sigma_finite_measure N" shows "AE x in M. f x = RN_deriv M N x" using fin unfolding eq[symmetric] - by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv_density - RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric]) + by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv + RN_deriv_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric]) lemma (in sigma_finite_measure) RN_deriv_distr: fixes T :: "'a \ 'b" @@ -1165,7 +1183,7 @@ using T ac by measurable then show "(\x. RN_deriv ?M' ?N' (T x)) \ borel_measurable M" by (simp add: comp_def) - show "AE x in M. 0 \ RN_deriv ?M' ?N' (T x)" using M'.RN_deriv_nonneg[OF ac] by auto + show "AE x in M. 0 \ RN_deriv ?M' ?N' (T x)" by (auto intro: RN_deriv_nonneg) have "N = distr N M (T' \ T)" by (subst measure_of_of_measure[of N, symmetric]) @@ -1175,7 +1193,7 @@ also have "\ = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'" using ac by (simp add: M'.density_RN_deriv) also have "\ = density M (RN_deriv (distr M M' T) (distr N M' T) \ T)" - using M'.borel_measurable_RN_deriv[OF ac] by (simp add: distr_density_distr[OF T T', OF inv]) + by (simp add: distr_density_distr[OF T T', OF inv]) finally show "density M (\x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N" by (simp add: comp_def) qed @@ -1186,7 +1204,9 @@ proof - interpret N: sigma_finite_measure N by fact from N show ?thesis - using sigma_finite_iff_density_finite[OF RN_deriv(1)[OF ac]] RN_deriv(2,3)[OF ac] by simp + using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N] + density_RN_deriv[OF ac] + by (simp add: RN_deriv_nonneg) qed lemma (in sigma_finite_measure) @@ -1194,29 +1214,31 @@ and f: "f \ borel_measurable M" shows RN_deriv_integrable: "integrable N f \ integrable M (\x. real (RN_deriv M N x) * f x)" (is ?integrable) - and RN_deriv_integral: "integral\<^sup>L N f = - (\x. real (RN_deriv M N x) * f x \M)" (is ?integral) + and RN_deriv_integral: "integral\<^sup>L N f = (\x. real (RN_deriv M N x) * f x \M)" (is ?integral) proof - note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp] interpret N: sigma_finite_measure N by fact - have minus_cong: "\A B A' B'::ereal. A = A' \ B = B' \ real A - real B = real A' - real B'" by simp - have f': "(\x. - f x) \ borel_measurable M" using f by auto - have Nf: "f \ borel_measurable N" using f by (simp add: measurable_def) - { fix f :: "'a \ real" - { fix x assume *: "RN_deriv M N x \ \" - have "ereal (real (RN_deriv M N x)) * ereal (f x) = ereal (real (RN_deriv M N x) * f x)" - by (simp add: mult_le_0_iff) - then have "RN_deriv M N x * ereal (f x) = ereal (real (RN_deriv M N x) * f x)" - using RN_deriv(3)[OF ac] * by (auto simp add: ereal_real split: split_if_asm) } - then have "(\\<^sup>+x. ereal (real (RN_deriv M N x) * f x) \M) = (\\<^sup>+x. RN_deriv M N x * ereal (f x) \M)" - "(\\<^sup>+x. ereal (- (real (RN_deriv M N x) * f x)) \M) = (\\<^sup>+x. RN_deriv M N x * ereal (- f x) \M)" - using RN_deriv_finite[OF N ac] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric] - by (auto intro!: positive_integral_cong_AE) } - note * = this - show ?integral ?integrable - unfolding lebesgue_integral_def integrable_def * - using Nf f RN_deriv(1)[OF ac] - by (auto simp: RN_deriv_positive_integral[OF ac]) + + have eq: "density M (RN_deriv M N) = density M (\x. real (RN_deriv M N x))" + proof (rule density_cong) + from RN_deriv_finite[OF assms(1,2,3)] + show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))" + by eventually_elim (insert RN_deriv_nonneg[of M N], auto simp: ereal_real) + qed (insert ac, auto) + + show ?integrable + apply (subst density_RN_deriv[OF ac, symmetric]) + unfolding eq + apply (intro integrable_real_density f AE_I2 real_of_ereal_pos RN_deriv_nonneg) + apply (insert ac, auto) + done + + show ?integral + apply (subst density_RN_deriv[OF ac, symmetric]) + unfolding eq + apply (intro integral_real_density f AE_I2 real_of_ereal_pos RN_deriv_nonneg) + apply (insert ac, auto) + done qed lemma (in sigma_finite_measure) real_RN_deriv: @@ -1229,7 +1251,7 @@ proof interpret N: finite_measure N by fact - note RN = RN_deriv[OF ac] + note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac] RN_deriv_nonneg[of M N] let ?RN = "\t. {x \ space M. RN_deriv M N x = t}" @@ -1277,11 +1299,10 @@ and x: "{x} \ sets M" shows "N {x} = RN_deriv M N x * emeasure M {x}" proof - - note deriv = RN_deriv[OF ac] - from deriv(1,3) `{x} \ sets M` + from `{x} \ sets M` have "density M (RN_deriv M N) {x} = (\\<^sup>+w. RN_deriv M N x * indicator {x} w \M)" by (auto simp: indicator_def emeasure_density intro!: positive_integral_cong) - with x deriv show ?thesis + with x density_RN_deriv[OF ac] RN_deriv_nonneg[of M N] show ?thesis by (auto simp: positive_integral_cmult_indicator) qed diff -r d0e5225601d3 -r e5366291d6aa src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon May 19 11:27:02 2014 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon May 19 12:04:45 2014 +0200 @@ -1391,6 +1391,11 @@ unfolding measurable_def using assms by (simp cong: vimage_inter_cong Pi_cong) +lemma measurable_cong_strong: + "M = N \ M' = N' \ (\w. w \ space M \ f w = g w) \ + f \ measurable M M' \ g \ measurable N N'" + by (metis measurable_cong) + lemma measurable_eqI: "\ space m1 = space m1' ; space m2 = space m2' ; sets m1 = sets m1' ; sets m2 = sets m2' \ @@ -1540,6 +1545,21 @@ unfolding measurable_def by auto qed +lemma measurable_count_space_eq2_countable: + fixes f :: "'a => 'c::countable" + shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" +proof - + { fix X assume "X \ A" "f \ space M \ A" + assume *: "\a. a\A \ f -` {a} \ space M \ sets M" + have "f -` X \ space M = (\a\X. f -` {a} \ space M)" + by auto + also have "\ \ sets M" + using * `X \ A` by (intro sets.countable_UN) auto + finally have "f -` X \ space M \ sets M" . } + then show ?thesis + unfolding measurable_def by auto +qed + lemma measurable_compose_countable: assumes f: "\i::'i::countable. (\x. f i x) \ measurable M N" and g: "g \ measurable M (count_space UNIV)" shows "(\x. f (g x) x) \ measurable M N"