--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Binary_Product_Measure.thy Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,1110 @@
+(* Title: HOL/Probability/Binary_Product_Measure.thy
+ Author: Johannes Hölzl, TU München
+*)
+
+section \<open>Binary product measures\<close>
+
+theory Binary_Product_Measure
+imports Nonnegative_Lebesgue_Integration
+begin
+
+lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
+ by auto
+
+lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
+ by auto
+
+subsection "Binary products"
+
+definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
+ "A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
+ {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
+ (\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
+
+lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
+ using sets.space_closed[of A] sets.space_closed[of B] by auto
+
+lemma space_pair_measure:
+ "space (A \<Otimes>\<^sub>M B) = space A \<times> space B"
+ unfolding pair_measure_def using pair_measure_closed[of A B]
+ by (rule space_measure_of)
+
+lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}"
+ by (auto simp: space_pair_measure)
+
+lemma sets_pair_measure:
+ "sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
+ unfolding pair_measure_def using pair_measure_closed[of A B]
+ by (rule sets_measure_of)
+
+lemma sets_pair_measure_cong[measurable_cong, cong]:
+ "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')"
+ unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
+
+lemma pair_measureI[intro, simp, measurable]:
+ "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)"
+ by (auto simp: sets_pair_measure)
+
+lemma sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ using pair_measureI[of "{x}" M1 "{y}" M2] by simp
+
+lemma measurable_pair_measureI:
+ assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"
+ assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M"
+ shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
+ unfolding pair_measure_def using 1 2
+ by (intro measurable_measure_of) (auto dest: sets.sets_into_space)
+
+lemma measurable_split_replace[measurable (raw)]:
+ "(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. case_prod (f x) (g x)) \<in> measurable M N"
+ unfolding split_beta' .
+
+lemma measurable_Pair[measurable (raw)]:
+ assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2"
+ shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
+proof (rule measurable_pair_measureI)
+ show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2"
+ using f g by (auto simp: measurable_def)
+ fix A B assume *: "A \<in> sets M1" "B \<in> sets M2"
+ have "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
+ by auto
+ also have "\<dots> \<in> sets M"
+ by (rule sets.Int) (auto intro!: measurable_sets * f g)
+ finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" .
+qed
+
+lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"
+ by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
+ measurable_def)
+
+lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"
+ by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times
+ measurable_def)
+
+lemma measurable_Pair_compose_split[measurable_dest]:
+ assumes f: "case_prod f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N"
+ assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2"
+ shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N"
+ using measurable_compose[OF measurable_Pair f, OF g h] by simp
+
+lemma measurable_Pair1_compose[measurable_dest]:
+ assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
+ assumes [measurable]: "h \<in> measurable N M"
+ shows "(\<lambda>x. f (h x)) \<in> measurable N M1"
+ using measurable_compose[OF f measurable_fst] by simp
+
+lemma measurable_Pair2_compose[measurable_dest]:
+ assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
+ assumes [measurable]: "h \<in> measurable N M"
+ shows "(\<lambda>x. g (h x)) \<in> measurable N M2"
+ using measurable_compose[OF f measurable_snd] by simp
+
+lemma measurable_pair:
+ assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
+ shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"
+ using measurable_Pair[OF assms] by simp
+
+lemma
+ assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)"
+ shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"
+ and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"
+ by simp_all
+
+lemma
+ assumes f[measurable]: "f \<in> measurable M N"
+ shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N"
+ and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N"
+ by simp_all
+
+lemma sets_pair_in_sets:
+ assumes "\<And>a b. a \<in> sets A \<Longrightarrow> b \<in> sets B \<Longrightarrow> a \<times> b \<in> sets N"
+ shows "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets N"
+ unfolding sets_pair_measure
+ by (intro sets.sigma_sets_subset') (auto intro!: assms)
+
+lemma sets_pair_eq_sets_fst_snd:
+ "sets (A \<Otimes>\<^sub>M B) = sets (Sup {vimage_algebra (space A \<times> space B) fst A, vimage_algebra (space A \<times> space B) snd B})"
+ (is "?P = sets (Sup {?fst, ?snd})")
+proof -
+ { fix a b assume ab: "a \<in> sets A" "b \<in> sets B"
+ then have "a \<times> b = (fst -` a \<inter> (space A \<times> space B)) \<inter> (snd -` b \<inter> (space A \<times> space B))"
+ by (auto dest: sets.sets_into_space)
+ also have "\<dots> \<in> sets (Sup {?fst, ?snd})"
+ apply (rule sets.Int)
+ apply (rule in_sets_Sup)
+ apply auto []
+ apply (rule insertI1)
+ apply (auto intro: ab in_vimage_algebra) []
+ apply (rule in_sets_Sup)
+ apply auto []
+ apply (rule insertI2)
+ apply (auto intro: ab in_vimage_algebra)
+ done
+ finally have "a \<times> b \<in> sets (Sup {?fst, ?snd})" . }
+ moreover have "sets ?fst \<subseteq> sets (A \<Otimes>\<^sub>M B)"
+ by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric])
+ moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)"
+ by (rule sets_image_in_sets) (auto simp: space_pair_measure)
+ ultimately show ?thesis
+ apply (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets)
+ apply simp
+ apply simp
+ apply simp
+ apply (elim disjE)
+ apply (simp add: space_pair_measure)
+ apply (simp add: space_pair_measure)
+ apply (auto simp add: space_pair_measure)
+ done
+qed
+
+lemma measurable_pair_iff:
+ "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
+ by (auto intro: measurable_pair[of f M M1 M2])
+
+lemma measurable_split_conv:
+ "(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"
+ by (intro arg_cong2[where f="op \<in>"]) auto
+
+lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"
+ by (auto intro!: measurable_Pair simp: measurable_split_conv)
+
+lemma measurable_pair_swap:
+ assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M"
+ using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def)
+
+lemma measurable_pair_swap_iff:
+ "f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M"
+ by (auto dest: measurable_pair_swap)
+
+lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"
+ by simp
+
+lemma sets_Pair1[measurable (raw)]:
+ assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2"
+proof -
+ have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
+ using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
+ also have "\<dots> \<in> sets M2"
+ using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: if_split_asm)
+ finally show ?thesis .
+qed
+
+lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"
+ by (auto intro!: measurable_Pair)
+
+lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
+proof -
+ have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
+ using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
+ also have "\<dots> \<in> sets M1"
+ using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: if_split_asm)
+ finally show ?thesis .
+qed
+
+lemma measurable_Pair2:
+ assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1"
+ shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
+ using measurable_comp[OF measurable_Pair1' f, OF x]
+ by (simp add: comp_def)
+
+lemma measurable_Pair1:
+ assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"
+ shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
+ using measurable_comp[OF measurable_Pair2' f, OF y]
+ by (simp add: comp_def)
+
+lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
+ unfolding Int_stable_def
+ by safe (auto simp add: times_Int_times)
+
+lemma (in finite_measure) finite_measure_cut_measurable:
+ assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)"
+ shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N"
+ (is "?s Q \<in> _")
+ using Int_stable_pair_measure_generator pair_measure_closed assms
+ unfolding sets_pair_measure
+proof (induct rule: sigma_sets_induct_disjoint)
+ case (compl A)
+ with sets.sets_into_space have "\<And>x. emeasure M (Pair x -` ((space N \<times> space M) - A)) =
+ (if x \<in> space N then emeasure M (space M) - ?s A x else 0)"
+ unfolding sets_pair_measure[symmetric]
+ by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1)
+ with compl sets.top show ?case
+ by (auto intro!: measurable_If simp: space_pair_measure)
+next
+ case (union F)
+ then have "\<And>x. emeasure M (Pair x -` (\<Union>i. F i)) = (\<Sum>i. ?s (F i) x)"
+ by (simp add: suminf_emeasure disjoint_family_on_vimageI subset_eq vimage_UN sets_pair_measure[symmetric])
+ with union show ?case
+ unfolding sets_pair_measure[symmetric] by simp
+qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If)
+
+lemma (in sigma_finite_measure) measurable_emeasure_Pair:
+ assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")
+proof -
+ from sigma_finite_disjoint guess F . note F = this
+ then have F_sets: "\<And>i. F i \<in> sets M" by auto
+ let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"
+ { fix i
+ have [simp]: "space N \<times> F i \<inter> space N \<times> space M = space N \<times> F i"
+ using F sets.sets_into_space by auto
+ let ?R = "density M (indicator (F i))"
+ have "finite_measure ?R"
+ using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)
+ then have "(\<lambda>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))) \<in> borel_measurable N"
+ by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q)
+ moreover have "\<And>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))
+ = emeasure M (F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q))"
+ using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1)
+ moreover have "\<And>x. F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q) = ?C x i"
+ using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure)
+ ultimately have "(\<lambda>x. emeasure M (?C x i)) \<in> borel_measurable N"
+ by simp }
+ moreover
+ { fix x
+ have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)"
+ proof (intro suminf_emeasure)
+ show "range (?C x) \<subseteq> sets M"
+ using F \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> by (auto intro!: sets_Pair1)
+ have "disjoint_family F" using F by auto
+ show "disjoint_family (?C x)"
+ by (rule disjoint_family_on_bisimulation[OF \<open>disjoint_family F\<close>]) auto
+ qed
+ also have "(\<Union>i. ?C x i) = Pair x -` Q"
+ using F sets.sets_into_space[OF \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close>]
+ by (auto simp: space_pair_measure)
+ finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))"
+ by simp }
+ ultimately show ?thesis using \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> F_sets
+ by auto
+qed
+
+lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:
+ assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M"
+ assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
+ shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N"
+proof -
+ from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} = A x"
+ by (auto simp: space_pair_measure)
+ with measurable_emeasure_Pair[OF A] show ?thesis
+ by (auto cong: measurable_cong)
+qed
+
+lemma (in sigma_finite_measure) emeasure_pair_measure:
+ assumes "X \<in> sets (N \<Otimes>\<^sub>M M)"
+ shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")
+proof (rule emeasure_measure_of[OF pair_measure_def])
+ show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
+ by (auto simp: positive_def)
+ have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"
+ by (auto simp: indicator_def)
+ show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"
+ proof (rule countably_additiveI)
+ fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F"
+ from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" by auto
+ moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
+ by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
+ moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M"
+ using F by (auto simp: sets_Pair1)
+ ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"
+ by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure
+ intro!: nn_integral_cong nn_integral_indicator[symmetric])
+ qed
+ show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)"
+ using sets.space_closed[of N] sets.space_closed[of M] by auto
+qed fact
+
+lemma (in sigma_finite_measure) emeasure_pair_measure_alt:
+ assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)"
+ shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)"
+proof -
+ have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"
+ by (auto simp: indicator_def)
+ show ?thesis
+ using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1)
+qed
+
+lemma (in sigma_finite_measure) emeasure_pair_measure_Times:
+ assumes A: "A \<in> sets N" and B: "B \<in> sets M"
+ shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B"
+proof -
+ have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"
+ using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt)
+ also have "\<dots> = emeasure M B * emeasure N A"
+ using A by (simp add: nn_integral_cmult_indicator)
+ finally show ?thesis
+ by (simp add: ac_simps)
+qed
+
+subsection \<open>Binary products of $\sigma$-finite emeasure spaces\<close>
+
+locale pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2
+ for M1 :: "'a measure" and M2 :: "'b measure"
+
+lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
+ "Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
+ using M2.measurable_emeasure_Pair .
+
+lemma (in pair_sigma_finite) measurable_emeasure_Pair2:
+ assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
+proof -
+ have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
+ using Q measurable_pair_swap' by (auto intro: measurable_sets)
+ note M1.measurable_emeasure_Pair[OF this]
+ moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1)) = (\<lambda>x. (x, y)) -` Q"
+ using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure)
+ ultimately show ?thesis by simp
+qed
+
+lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
+ defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
+ shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and>
+ (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"
+proof -
+ from M1.sigma_finite_incseq guess F1 . note F1 = this
+ from M2.sigma_finite_incseq guess F2 . note F2 = this
+ from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
+ let ?F = "\<lambda>i. F1 i \<times> F2 i"
+ show ?thesis
+ proof (intro exI[of _ ?F] conjI allI)
+ show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD)
+ next
+ have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
+ proof (intro subsetI)
+ fix x assume "x \<in> space M1 \<times> space M2"
+ then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
+ by (auto simp: space)
+ then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
+ using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_def
+ by (force split: split_max)+
+ then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
+ by (intro SigmaI) (auto simp add: max.commute)
+ then show "x \<in> (\<Union>i. ?F i)" by auto
+ qed
+ then show "(\<Union>i. ?F i) = space M1 \<times> space M2"
+ using space by (auto simp: space)
+ next
+ fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
+ using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_Suc_iff by auto
+ next
+ fix i
+ from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
+ with F1 F2 show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
+ by (auto simp add: emeasure_pair_measure_Times ennreal_mult_eq_top_iff)
+ qed
+qed
+
+sublocale pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"
+proof
+ from M1.sigma_finite_countable guess F1 ..
+ moreover from M2.sigma_finite_countable guess F2 ..
+ ultimately show
+ "\<exists>A. countable A \<and> A \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> \<Union>A = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>a\<in>A. emeasure (M1 \<Otimes>\<^sub>M M2) a \<noteq> \<infinity>)"
+ by (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (F1 \<times> F2)"] conjI)
+ (auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq ennreal_mult_eq_top_iff)
+qed
+
+lemma sigma_finite_pair_measure:
+ assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"
+ shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)"
+proof -
+ interpret A: sigma_finite_measure A by fact
+ interpret B: sigma_finite_measure B by fact
+ interpret AB: pair_sigma_finite A B ..
+ show ?thesis ..
+qed
+
+lemma sets_pair_swap:
+ assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"
+ using measurable_pair_swap' assms by (rule measurable_sets)
+
+lemma (in pair_sigma_finite) distr_pair_swap:
+ "M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
+proof -
+ from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+ let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
+ show ?thesis
+ proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
+ show "?E \<subseteq> Pow (space ?P)"
+ using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)
+ show "sets ?P = sigma_sets (space ?P) ?E"
+ by (simp add: sets_pair_measure space_pair_measure)
+ then show "sets ?D = sigma_sets (space ?P) ?E"
+ by simp
+ next
+ show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
+ using F by (auto simp: space_pair_measure)
+ next
+ fix X assume "X \<in> ?E"
+ then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto
+ have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^sub>M M1) = B \<times> A"
+ using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure)
+ with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X"
+ by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr
+ measurable_pair_swap' ac_simps)
+ qed
+qed
+
+lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:
+ assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
+ (is "_ = ?\<nu> A")
+proof -
+ have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A"
+ using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
+ show ?thesis using A
+ by (subst distr_pair_swap)
+ (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap']
+ M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
+qed
+
+lemma (in pair_sigma_finite) AE_pair:
+ assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x"
+ shows "AE x in M1. (AE y in M2. Q (x, y))"
+proof -
+ obtain N where N: "N \<in> sets (M1 \<Otimes>\<^sub>M M2)" "emeasure (M1 \<Otimes>\<^sub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> Q x} \<subseteq> N"
+ using assms unfolding eventually_ae_filter by auto
+ show ?thesis
+ proof (rule AE_I)
+ from N measurable_emeasure_Pair1[OF \<open>N \<in> sets (M1 \<Otimes>\<^sub>M M2)\<close>]
+ show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
+ by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff)
+ show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
+ by (intro borel_measurable_eq measurable_emeasure_Pair1 N sets.sets_Collect_neg N) simp
+ { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
+ have "AE y in M2. Q (x, y)"
+ proof (rule AE_I)
+ show "emeasure M2 (Pair x -` N) = 0" by fact
+ show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)
+ show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
+ using N \<open>x \<in> space M1\<close> unfolding space_pair_measure by auto
+ qed }
+ then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}"
+ by auto
+ qed
+qed
+
+lemma (in pair_sigma_finite) AE_pair_measure:
+ assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ assumes ae: "AE x in M1. AE y in M2. P (x, y)"
+ shows "AE x in M1 \<Otimes>\<^sub>M M2. P x"
+proof (subst AE_iff_measurable[OF _ refl])
+ show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ by (rule sets.sets_Collect) fact
+ then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} =
+ (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
+ by (simp add: M2.emeasure_pair_measure)
+ also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)"
+ using ae
+ apply (safe intro!: nn_integral_cong_AE)
+ apply (intro AE_I2)
+ apply (safe intro!: nn_integral_cong_AE)
+ apply auto
+ done
+ finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
+qed
+
+lemma (in pair_sigma_finite) AE_pair_iff:
+ "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow>
+ (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))"
+ using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto
+
+lemma (in pair_sigma_finite) AE_commute:
+ assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
+ shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"
+proof -
+ interpret Q: pair_sigma_finite M2 M1 ..
+ have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x"
+ by auto
+ have "{x \<in> space (M2 \<Otimes>\<^sub>M M1). P (snd x) (fst x)} =
+ (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^sub>M M1)"
+ by (auto simp: space_pair_measure)
+ also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)"
+ by (intro sets_pair_swap P)
+ finally show ?thesis
+ apply (subst AE_pair_iff[OF P])
+ apply (subst distr_pair_swap)
+ apply (subst AE_distr_iff[OF measurable_pair_swap' P])
+ apply (subst Q.AE_pair_iff)
+ apply simp_all
+ done
+qed
+
+subsection "Fubinis theorem"
+
+lemma measurable_compose_Pair1:
+ "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
+ by simp
+
+lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst:
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
+ shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"
+using f proof induct
+ case (cong u v)
+ then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)"
+ by (auto simp: space_pair_measure)
+ show ?case
+ apply (subst measurable_cong)
+ apply (rule nn_integral_cong)
+ apply fact+
+ done
+next
+ case (set Q)
+ have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y"
+ by (auto simp: indicator_def)
+ have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M"
+ by (simp add: sets_Pair1[OF set])
+ from this measurable_emeasure_Pair[OF set] show ?case
+ by (rule measurable_cong[THEN iffD1])
+qed (simp_all add: nn_integral_add nn_integral_cmult measurable_compose_Pair1
+ nn_integral_monotone_convergence_SUP incseq_def le_fun_def
+ cong: measurable_cong)
+
+lemma (in sigma_finite_measure) nn_integral_fst:
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"
+ shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")
+using f proof induct
+ case (cong u v)
+ then have "?I u = ?I v"
+ by (intro nn_integral_cong) (auto simp: space_pair_measure)
+ with cong show ?case
+ by (simp cong: nn_integral_cong)
+qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add
+ nn_integral_monotone_convergence_SUP measurable_compose_Pair1
+ borel_measurable_nn_integral_fst nn_integral_mono incseq_def le_fun_def
+ cong: nn_integral_cong)
+
+lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:
+ "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"
+ using borel_measurable_nn_integral_fst[of "case_prod f" N] by simp
+
+lemma (in pair_sigma_finite) nn_integral_snd:
+ assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
+proof -
+ note measurable_pair_swap[OF f]
+ from M1.nn_integral_fst[OF this]
+ have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))"
+ by simp
+ also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"
+ by (subst distr_pair_swap) (auto simp add: nn_integral_distr intro!: nn_integral_cong)
+ finally show ?thesis .
+qed
+
+lemma (in pair_sigma_finite) Fubini:
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
+ unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] ..
+
+lemma (in pair_sigma_finite) Fubini':
+ assumes f: "case_prod f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)"
+ using Fubini[OF f] by simp
+
+subsection \<open>Products on counting spaces, densities and distributions\<close>
+
+lemma sigma_prod:
+ assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X"
+ assumes Y_cover: "\<exists>E\<subseteq>B. countable E \<and> Y = \<Union>E" and B: "B \<subseteq> Pow Y"
+ shows "sigma X A \<Otimes>\<^sub>M sigma Y B = sigma (X \<times> Y) {a \<times> b | a b. a \<in> A \<and> b \<in> B}"
+ (is "?P = ?S")
+proof (rule measure_eqI)
+ have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X"
+ by auto
+ let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}"
+ have "sets ?P = sets (SUP xy:?XY. sigma (X \<times> Y) xy)"
+ by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B)
+ also have "\<dots> = sets (sigma (X \<times> Y) (\<Union>?XY))"
+ by (intro Sup_sigma arg_cong[where f=sets]) auto
+ also have "\<dots> = sets ?S"
+ proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)
+ show "\<Union>?XY \<subseteq> Pow (X \<times> Y)" "{a \<times> b |a b. a \<in> A \<and> b \<in> B} \<subseteq> Pow (X \<times> Y)"
+ using A B by auto
+ next
+ interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
+ using A B by (intro sigma_algebra_sigma_sets) auto
+ fix Z assume "Z \<in> \<Union>?XY"
+ then show "Z \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
+ proof safe
+ fix a assume "a \<in> A"
+ from Y_cover obtain E where E: "E \<subseteq> B" "countable E" and "Y = \<Union>E"
+ by auto
+ with \<open>a \<in> A\<close> A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)"
+ by auto
+ show "fst -` a \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
+ using \<open>a \<in> A\<close> E unfolding eq by (auto intro!: XY.countable_UN')
+ next
+ fix b assume "b \<in> B"
+ from X_cover obtain E where E: "E \<subseteq> A" "countable E" and "X = \<Union>E"
+ by auto
+ with \<open>b \<in> B\<close> B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)"
+ by auto
+ show "snd -` b \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
+ using \<open>b \<in> B\<close> E unfolding eq by (auto intro!: XY.countable_UN')
+ qed
+ next
+ fix Z assume "Z \<in> {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
+ then obtain a b where "Z = a \<times> b" and ab: "a \<in> A" "b \<in> B"
+ by auto
+ then have Z: "Z = (fst -` a \<inter> X \<times> Y) \<inter> (snd -` b \<inter> X \<times> Y)"
+ using A B by auto
+ interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) (\<Union>?XY)"
+ by (intro sigma_algebra_sigma_sets) auto
+ show "Z \<in> sigma_sets (X \<times> Y) (\<Union>?XY)"
+ unfolding Z by (rule XY.Int) (blast intro: ab)+
+ qed
+ finally show "sets ?P = sets ?S" .
+next
+ interpret finite_measure "sigma X A" for X A
+ proof qed (simp add: emeasure_sigma)
+ fix A assume "A \<in> sets ?P" then show "emeasure ?P A = emeasure ?S A"
+ by (simp add: emeasure_pair_measure_alt emeasure_sigma)
+qed
+
+lemma sigma_sets_pair_measure_generator_finite:
+ assumes "finite A" and "finite B"
+ shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
+ (is "sigma_sets ?prod ?sets = _")
+proof safe
+ have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
+ fix x assume subset: "x \<subseteq> A \<times> B"
+ hence "finite x" using fin by (rule finite_subset)
+ from this subset show "x \<in> sigma_sets ?prod ?sets"
+ proof (induct x)
+ case empty show ?case by (rule sigma_sets.Empty)
+ next
+ case (insert a x)
+ hence "{a} \<in> sigma_sets ?prod ?sets" by auto
+ moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
+ ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
+ qed
+next
+ fix x a b
+ assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
+ from sigma_sets_into_sp[OF _ this(1)] this(2)
+ show "a \<in> A" and "b \<in> B" by auto
+qed
+
+lemma borel_prod:
+ "(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)"
+ (is "?P = ?B")
+proof -
+ have "?B = sigma UNIV {A \<times> B | A B. open A \<and> open B}"
+ by (rule second_countable_borel_measurable[OF open_prod_generated])
+ also have "\<dots> = ?P"
+ unfolding borel_def
+ by (subst sigma_prod) (auto intro!: exI[of _ "{UNIV}"])
+ finally show ?thesis ..
+qed
+
+lemma pair_measure_count_space:
+ assumes A: "finite A" and B: "finite B"
+ shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
+proof (rule measure_eqI)
+ interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact
+ interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact
+ interpret P: pair_sigma_finite "count_space A" "count_space B" ..
+ show eq: "sets ?P = sets ?C"
+ by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B)
+ fix X assume X: "X \<in> sets ?P"
+ with eq have X_subset: "X \<subseteq> A \<times> B" by simp
+ with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"
+ by (intro finite_subset[OF _ B]) auto
+ have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)
+ have pos_card: "(0::ennreal) < of_nat (card (Pair x -` X)) \<longleftrightarrow> Pair x -` X \<noteq> {}" for x
+ by (auto simp: card_eq_0_iff fin_Pair) blast
+
+ show "emeasure ?P X = emeasure ?C X"
+ using X_subset A fin_Pair fin_X
+ apply (subst B.emeasure_pair_measure_alt[OF X])
+ apply (subst emeasure_count_space)
+ apply (auto simp add: emeasure_count_space nn_integral_count_space
+ pos_card of_nat_setsum[symmetric] card_SigmaI[symmetric]
+ simp del: of_nat_setsum card_SigmaI
+ intro!: arg_cong[where f=card])
+ done
+qed
+
+
+lemma emeasure_prod_count_space:
+ assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M M)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")
+ shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator A (x, y) \<partial>?B \<partial>?A)"
+ by (rule emeasure_measure_of[OF pair_measure_def])
+ (auto simp: countably_additive_def positive_def suminf_indicator A
+ nn_integral_suminf[symmetric] dest: sets.sets_into_space)
+
+lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
+proof -
+ have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ennreal)"
+ by (auto split: split_indicator)
+ show ?thesis
+ by (cases x) (auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair)
+qed
+
+lemma emeasure_count_space_prod_eq:
+ fixes A :: "('a \<times> 'b) set"
+ assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M count_space UNIV)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")
+ shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A"
+proof -
+ { fix A :: "('a \<times> 'b) set" assume "countable A"
+ then have "emeasure (?A \<Otimes>\<^sub>M ?B) (\<Union>a\<in>A. {a}) = (\<integral>\<^sup>+a. emeasure (?A \<Otimes>\<^sub>M ?B) {a} \<partial>count_space A)"
+ by (intro emeasure_UN_countable) (auto simp: sets_Pair disjoint_family_on_def)
+ also have "\<dots> = (\<integral>\<^sup>+a. indicator A a \<partial>count_space UNIV)"
+ by (subst nn_integral_count_space_indicator) auto
+ finally have "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A"
+ by simp }
+ note * = this
+
+ show ?thesis
+ proof cases
+ assume "finite A" then show ?thesis
+ by (intro * countable_finite)
+ next
+ assume "infinite A"
+ then obtain C where "countable C" and "infinite C" and "C \<subseteq> A"
+ by (auto dest: infinite_countable_subset')
+ with A have "emeasure (?A \<Otimes>\<^sub>M ?B) C \<le> emeasure (?A \<Otimes>\<^sub>M ?B) A"
+ by (intro emeasure_mono) auto
+ also have "emeasure (?A \<Otimes>\<^sub>M ?B) C = emeasure (count_space UNIV) C"
+ using \<open>countable C\<close> by (rule *)
+ finally show ?thesis
+ using \<open>infinite C\<close> \<open>infinite A\<close> by (simp add: top_unique)
+ qed
+qed
+
+lemma nn_integral_count_space_prod_eq:
+ "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"
+ (is "nn_integral ?P f = _")
+proof cases
+ assume cntbl: "countable {x. f x \<noteq> 0}"
+ have [simp]: "\<And>x. card ({x} \<inter> {x. f x \<noteq> 0}) = (indicator {x. f x \<noteq> 0} x::ennreal)"
+ by (auto split: split_indicator)
+ have [measurable]: "\<And>y. (\<lambda>x. indicator {y} x) \<in> borel_measurable ?P"
+ by (rule measurable_discrete_difference[of "\<lambda>x. 0" _ borel "{y}" "\<lambda>x. indicator {y} x" for y])
+ (auto intro: sets_Pair)
+
+ have "(\<integral>\<^sup>+x. f x \<partial>?P) = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x * indicator {x} x' \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)"
+ by (auto simp add: nn_integral_cmult nn_integral_indicator' intro!: nn_integral_cong split: split_indicator)
+ also have "\<dots> = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x' * indicator {x'} x \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)"
+ by (auto intro!: nn_integral_cong split: split_indicator)
+ also have "\<dots> = (\<integral>\<^sup>+x'. \<integral>\<^sup>+x. f x' * indicator {x'} x \<partial>?P \<partial>count_space {x. f x \<noteq> 0})"
+ by (intro nn_integral_count_space_nn_integral cntbl) auto
+ also have "\<dots> = (\<integral>\<^sup>+x'. f x' \<partial>count_space {x. f x \<noteq> 0})"
+ by (intro nn_integral_cong) (auto simp: nn_integral_cmult sets_Pair)
+ finally show ?thesis
+ by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
+next
+ { fix x assume "f x \<noteq> 0"
+ then have "(\<exists>r\<ge>0. 0 < r \<and> f x = ennreal r) \<or> f x = \<infinity>"
+ by (cases "f x" rule: ennreal_cases) (auto simp: less_le)
+ then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f x"
+ by (auto elim!: nat_approx_posE intro!: less_imp_le) }
+ note * = this
+
+ assume cntbl: "uncountable {x. f x \<noteq> 0}"
+ also have "{x. f x \<noteq> 0} = (\<Union>n. {x. 1/Suc n \<le> f x})"
+ using * by auto
+ finally obtain n where "infinite {x. 1/Suc n \<le> f x}"
+ by (meson countableI_type countable_UN uncountable_infinite)
+ then obtain C where C: "C \<subseteq> {x. 1/Suc n \<le> f x}" and "countable C" "infinite C"
+ by (metis infinite_countable_subset')
+
+ have [measurable]: "C \<in> sets ?P"
+ using sets.countable[OF _ \<open>countable C\<close>, of ?P] by (auto simp: sets_Pair)
+
+ have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f"
+ using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])
+ moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>"
+ using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq ennreal_mult_top)
+ moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f"
+ using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])
+ moreover have "(\<integral>\<^sup>+x. ennreal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>"
+ using \<open>infinite C\<close> by (simp add: nn_integral_cmult ennreal_mult_top)
+ ultimately show ?thesis
+ by (simp add: top_unique)
+qed
+
+lemma pair_measure_density:
+ assumes f: "f \<in> borel_measurable M1"
+ assumes g: "g \<in> borel_measurable M2"
+ assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)"
+ shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
+proof (rule measure_eqI)
+ interpret M2: sigma_finite_measure M2 by fact
+ interpret D2: sigma_finite_measure "density M2 g" by fact
+
+ fix A assume A: "A \<in> sets ?L"
+ with f g have "(\<integral>\<^sup>+ x. f x * \<integral>\<^sup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) =
+ (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)"
+ by (intro nn_integral_cong_AE)
+ (auto simp add: nn_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 nn_integral_density
+ M2.nn_integral_fst[symmetric]
+ cong: nn_integral_cong)
+qed simp
+
+lemma sigma_finite_measure_distr:
+ assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"
+ shows "sigma_finite_measure M"
+proof -
+ interpret sigma_finite_measure "distr M N f" by fact
+ from sigma_finite_countable guess A .. note A = this
+ show ?thesis
+ proof
+ show "\<exists>A. countable A \<and> A \<subseteq> sets M \<and> \<Union>A = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
+ using A f
+ by (intro exI[of _ "(\<lambda>a. f -` a \<inter> space M) ` A"])
+ (auto simp: emeasure_distr set_eq_iff subset_eq intro: measurable_space)
+ qed
+qed
+
+lemma pair_measure_distr:
+ assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"
+ assumes "sigma_finite_measure (distr N T g)"
+ shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
+proof (rule measure_eqI)
+ interpret T: sigma_finite_measure "distr N T g" by fact
+ interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+
+
+ fix A assume A: "A \<in> sets ?P"
+ with f g show "emeasure ?P A = emeasure ?D A"
+ by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr
+ T.emeasure_pair_measure_alt nn_integral_distr
+ intro!: nn_integral_cong arg_cong[where f="emeasure N"])
+qed simp
+
+lemma pair_measure_eqI:
+ assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
+ assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M"
+ assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)"
+ shows "M1 \<Otimes>\<^sub>M M2 = M"
+proof -
+ interpret M1: sigma_finite_measure M1 by fact
+ interpret M2: sigma_finite_measure M2 by fact
+ interpret pair_sigma_finite M1 M2 ..
+ from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+ let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
+ let ?P = "M1 \<Otimes>\<^sub>M M2"
+ show ?thesis
+ proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
+ show "?E \<subseteq> Pow (space ?P)"
+ using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)
+ show "sets ?P = sigma_sets (space ?P) ?E"
+ by (simp add: sets_pair_measure space_pair_measure)
+ then show "sets M = sigma_sets (space ?P) ?E"
+ using sets[symmetric] by simp
+ next
+ show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
+ using F by (auto simp: space_pair_measure)
+ next
+ fix X assume "X \<in> ?E"
+ then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto
+ then have "emeasure ?P X = emeasure M1 A * emeasure M2 B"
+ by (simp add: M2.emeasure_pair_measure_Times)
+ also have "\<dots> = emeasure M (A \<times> B)"
+ using A B emeasure by auto
+ finally show "emeasure ?P X = emeasure M X"
+ by simp
+ qed
+qed
+
+lemma sets_pair_countable:
+ assumes "countable S1" "countable S2"
+ assumes M: "sets M = Pow S1" and N: "sets N = Pow S2"
+ shows "sets (M \<Otimes>\<^sub>M N) = Pow (S1 \<times> S2)"
+proof auto
+ fix x a b assume x: "x \<in> sets (M \<Otimes>\<^sub>M N)" "(a, b) \<in> x"
+ from sets.sets_into_space[OF x(1)] x(2)
+ sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N
+ show "a \<in> S1" "b \<in> S2"
+ by (auto simp: space_pair_measure)
+next
+ fix X assume X: "X \<subseteq> S1 \<times> S2"
+ then have "countable X"
+ by (metis countable_subset \<open>countable S1\<close> \<open>countable S2\<close> countable_SIGMA)
+ have "X = (\<Union>(a, b)\<in>X. {a} \<times> {b})" by auto
+ also have "\<dots> \<in> sets (M \<Otimes>\<^sub>M N)"
+ using X
+ by (safe intro!: sets.countable_UN' \<open>countable X\<close> subsetI pair_measureI) (auto simp: M N)
+ finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" .
+qed
+
+lemma pair_measure_countable:
+ assumes "countable S1" "countable S2"
+ shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)"
+proof (rule pair_measure_eqI)
+ show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)"
+ using assms by (auto intro!: sigma_finite_measure_count_space_countable)
+ show "sets (count_space S1 \<Otimes>\<^sub>M count_space S2) = sets (count_space (S1 \<times> S2))"
+ by (subst sets_pair_countable[OF assms]) auto
+next
+ fix A B assume "A \<in> sets (count_space S1)" "B \<in> sets (count_space S2)"
+ then show "emeasure (count_space S1) A * emeasure (count_space S2) B =
+ emeasure (count_space (S1 \<times> S2)) (A \<times> B)"
+ by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff ennreal_mult_top ennreal_top_mult)
+qed
+
+lemma nn_integral_fst_count_space:
+ "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
+ (is "?lhs = ?rhs")
+proof(cases)
+ assume *: "countable {xy. f xy \<noteq> 0}"
+ let ?A = "fst ` {xy. f xy \<noteq> 0}"
+ let ?B = "snd ` {xy. f xy \<noteq> 0}"
+ from * have [simp]: "countable ?A" "countable ?B" by(rule countable_image)+
+ have "?lhs = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space ?A)"
+ by(rule nn_integral_count_space_eq)
+ (auto simp add: nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI)
+ also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space ?B \<partial>count_space ?A)"
+ by(intro nn_integral_count_space_eq nn_integral_cong)(auto intro: rev_image_eqI)
+ also have "\<dots> = (\<integral>\<^sup>+ xy. f xy \<partial>count_space (?A \<times> ?B))"
+ by(subst sigma_finite_measure.nn_integral_fst)
+ (simp_all add: sigma_finite_measure_count_space_countable pair_measure_countable)
+ also have "\<dots> = ?rhs"
+ by(rule nn_integral_count_space_eq)(auto intro: rev_image_eqI)
+ finally show ?thesis .
+next
+ { fix xy assume "f xy \<noteq> 0"
+ then have "(\<exists>r\<ge>0. 0 < r \<and> f xy = ennreal r) \<or> f xy = \<infinity>"
+ by (cases "f xy" rule: ennreal_cases) (auto simp: less_le)
+ then have "\<exists>n. ennreal (1 / real (Suc n)) \<le> f xy"
+ by (auto elim!: nat_approx_posE intro!: less_imp_le) }
+ note * = this
+
+ assume cntbl: "uncountable {xy. f xy \<noteq> 0}"
+ also have "{xy. f xy \<noteq> 0} = (\<Union>n. {xy. 1/Suc n \<le> f xy})"
+ using * by auto
+ finally obtain n where "infinite {xy. 1/Suc n \<le> f xy}"
+ by (meson countableI_type countable_UN uncountable_infinite)
+ then obtain C where C: "C \<subseteq> {xy. 1/Suc n \<le> f xy}" and "countable C" "infinite C"
+ by (metis infinite_countable_subset')
+
+ have "\<infinity> = (\<integral>\<^sup>+ xy. ennreal (1 / Suc n) * indicator C xy \<partial>count_space UNIV)"
+ using \<open>infinite C\<close> by(simp add: nn_integral_cmult ennreal_mult_top)
+ also have "\<dots> \<le> ?rhs" using C
+ by(intro nn_integral_mono)(auto split: split_indicator)
+ finally have "?rhs = \<infinity>" by (simp add: top_unique)
+ moreover have "?lhs = \<infinity>"
+ proof(cases "finite (fst ` C)")
+ case True
+ then obtain x C' where x: "x \<in> fst ` C"
+ and C': "C' = fst -` {x} \<inter> C"
+ and "infinite C'"
+ using \<open>infinite C\<close> by(auto elim!: inf_img_fin_domE')
+ from x C C' have **: "C' \<subseteq> {xy. 1 / Suc n \<le> f xy}" by auto
+
+ from C' \<open>infinite C'\<close> have "infinite (snd ` C')"
+ by(auto dest!: finite_imageD simp add: inj_on_def)
+ then have "\<infinity> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator (snd ` C') y \<partial>count_space UNIV)"
+ by(simp add: nn_integral_cmult ennreal_mult_top)
+ also have "\<dots> = (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV)"
+ by(rule nn_integral_cong)(force split: split_indicator intro: rev_image_eqI simp add: C')
+ also have "\<dots> = (\<integral>\<^sup>+ x'. (\<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV) * indicator {x} x' \<partial>count_space UNIV)"
+ by(simp add: one_ereal_def[symmetric])
+ also have "\<dots> \<le> (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV \<partial>count_space UNIV)"
+ by(rule nn_integral_mono)(simp split: split_indicator)
+ also have "\<dots> \<le> ?lhs" using **
+ by(intro nn_integral_mono)(auto split: split_indicator)
+ finally show ?thesis by (simp add: top_unique)
+ next
+ case False
+ define C' where "C' = fst ` C"
+ have "\<infinity> = \<integral>\<^sup>+ x. ennreal (1 / Suc n) * indicator C' x \<partial>count_space UNIV"
+ using C'_def False by(simp add: nn_integral_cmult ennreal_mult_top)
+ also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV"
+ by(auto simp add: one_ereal_def[symmetric] max_def intro: nn_integral_cong)
+ also have "\<dots> \<le> \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ennreal (1 / Suc n) * indicator C (x, y) \<partial>count_space UNIV \<partial>count_space UNIV"
+ by(intro nn_integral_mono)(auto simp add: C'_def split: split_indicator intro: someI)
+ also have "\<dots> \<le> ?lhs" using C
+ by(intro nn_integral_mono)(auto split: split_indicator)
+ finally show ?thesis by (simp add: top_unique)
+ qed
+ ultimately show ?thesis by simp
+qed
+
+lemma nn_integral_snd_count_space:
+ "(\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"
+ (is "?lhs = ?rhs")
+proof -
+ have "?lhs = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. (\<lambda>(y, x). f (x, y)) (y, x) \<partial>count_space UNIV \<partial>count_space UNIV)"
+ by(simp)
+ also have "\<dots> = \<integral>\<^sup>+ yx. (\<lambda>(y, x). f (x, y)) yx \<partial>count_space UNIV"
+ by(rule nn_integral_fst_count_space)
+ also have "\<dots> = \<integral>\<^sup>+ xy. f xy \<partial>count_space ((\<lambda>(x, y). (y, x)) ` UNIV)"
+ by(subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric])
+ (simp_all add: inj_on_def split_def)
+ also have "\<dots> = ?rhs" by(rule nn_integral_count_space_eq) auto
+ finally show ?thesis .
+qed
+
+lemma measurable_pair_measure_countable1:
+ assumes "countable A"
+ and [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N K"
+ shows "f \<in> measurable (count_space A \<Otimes>\<^sub>M N) K"
+using _ _ assms(1)
+by(rule measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all
+
+subsection \<open>Product of Borel spaces\<close>
+
+lemma borel_Times:
+ fixes A :: "'a::topological_space set" and B :: "'b::topological_space set"
+ assumes A: "A \<in> sets borel" and B: "B \<in> sets borel"
+ shows "A \<times> B \<in> sets borel"
+proof -
+ have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)"
+ by auto
+ moreover
+ { have "A \<in> sigma_sets UNIV {S. open S}" using A by (simp add: sets_borel)
+ then have "A\<times>UNIV \<in> sets borel"
+ proof (induct A)
+ case (Basic S) then show ?case
+ by (auto intro!: borel_open open_Times)
+ next
+ case (Compl A)
+ moreover have *: "(UNIV - A) \<times> UNIV = UNIV - (A \<times> UNIV)"
+ by auto
+ ultimately show ?case
+ unfolding * by auto
+ next
+ case (Union A)
+ moreover have *: "(UNION UNIV A) \<times> UNIV = UNION UNIV (\<lambda>i. A i \<times> UNIV)"
+ by auto
+ ultimately show ?case
+ unfolding * by auto
+ qed simp }
+ moreover
+ { have "B \<in> sigma_sets UNIV {S. open S}" using B by (simp add: sets_borel)
+ then have "UNIV\<times>B \<in> sets borel"
+ proof (induct B)
+ case (Basic S) then show ?case
+ by (auto intro!: borel_open open_Times)
+ next
+ case (Compl B)
+ moreover have *: "UNIV \<times> (UNIV - B) = UNIV - (UNIV \<times> B)"
+ by auto
+ ultimately show ?case
+ unfolding * by auto
+ next
+ case (Union B)
+ moreover have *: "UNIV \<times> (UNION UNIV B) = UNION UNIV (\<lambda>i. UNIV \<times> B i)"
+ by auto
+ ultimately show ?case
+ unfolding * by auto
+ qed simp }
+ ultimately show ?thesis
+ by auto
+qed
+
+lemma finite_measure_pair_measure:
+ assumes "finite_measure M" "finite_measure N"
+ shows "finite_measure (N \<Otimes>\<^sub>M M)"
+proof (rule finite_measureI)
+ interpret M: finite_measure M by fact
+ interpret N: finite_measure N by fact
+ show "emeasure (N \<Otimes>\<^sub>M M) (space (N \<Otimes>\<^sub>M M)) \<noteq> \<infinity>"
+ by (auto simp: space_pair_measure M.emeasure_pair_measure_Times ennreal_mult_eq_top_iff)
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Bochner_Integration.thy Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,3066 @@
+(* Title: HOL/Probability/Bochner_Integration.thy
+ Author: Johannes Hölzl, TU München
+*)
+
+section \<open>Bochner Integration for Vector-Valued Functions\<close>
+
+theory Bochner_Integration
+ imports Finite_Product_Measure
+begin
+
+text \<open>
+
+In the following development of the Bochner integral we use second countable topologies instead
+of separable spaces. A second countable topology is also separable.
+
+\<close>
+
+lemma borel_measurable_implies_sequence_metric:
+ fixes f :: "'a \<Rightarrow> 'b :: {metric_space, second_countable_topology}"
+ assumes [measurable]: "f \<in> borel_measurable M"
+ shows "\<exists>F. (\<forall>i. simple_function M (F i)) \<and> (\<forall>x\<in>space M. (\<lambda>i. F i x) \<longlonglongrightarrow> f x) \<and>
+ (\<forall>i. \<forall>x\<in>space M. dist (F i x) z \<le> 2 * dist (f x) z)"
+proof -
+ obtain D :: "'b set" where "countable D" and D: "\<And>X. open X \<Longrightarrow> X \<noteq> {} \<Longrightarrow> \<exists>d\<in>D. d \<in> X"
+ by (erule countable_dense_setE)
+
+ define e where "e = from_nat_into D"
+ { fix n x
+ obtain d where "d \<in> D" and d: "d \<in> ball x (1 / Suc n)"
+ using D[of "ball x (1 / Suc n)"] by auto
+ from \<open>d \<in> D\<close> D[of UNIV] \<open>countable D\<close> obtain i where "d = e i"
+ unfolding e_def by (auto dest: from_nat_into_surj)
+ with d have "\<exists>i. dist x (e i) < 1 / Suc n"
+ by auto }
+ note e = this
+
+ define A where [abs_def]: "A m n =
+ {x\<in>space M. dist (f x) (e n) < 1 / (Suc m) \<and> 1 / (Suc m) \<le> dist (f x) z}" for m n
+ define B where [abs_def]: "B m = disjointed (A m)" for m
+
+ define m where [abs_def]: "m N x = Max {m. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}" for N x
+ define F where [abs_def]: "F N x =
+ (if (\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)) \<and> (\<exists>n\<le>N. x \<in> B (m N x) n)
+ then e (LEAST n. x \<in> B (m N x) n) else z)" for N x
+
+ have B_imp_A[intro, simp]: "\<And>x m n. x \<in> B m n \<Longrightarrow> x \<in> A m n"
+ using disjointed_subset[of "A m" for m] unfolding B_def by auto
+
+ { fix m
+ have "\<And>n. A m n \<in> sets M"
+ by (auto simp: A_def)
+ then have "\<And>n. B m n \<in> sets M"
+ using sets.range_disjointed_sets[of "A m" M] by (auto simp: B_def) }
+ note this[measurable]
+
+ { fix N i x assume "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)"
+ then have "m N x \<in> {m::nat. m \<le> N \<and> x \<in> (\<Union>n\<le>N. B m n)}"
+ unfolding m_def by (intro Max_in) auto
+ then have "m N x \<le> N" "\<exists>n\<le>N. x \<in> B (m N x) n"
+ by auto }
+ note m = this
+
+ { fix j N i x assume "j \<le> N" "i \<le> N" "x \<in> B j i"
+ then have "j \<le> 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]: "\<And>i. F i \<in> borel_measurable M"
+ unfolding F_def m_def by measurable
+ show "\<And>x i. F i -` {x} \<inter> space M \<in> sets M"
+ by measurable
+
+ { fix i
+ { fix n x assume "x \<in> B (m i x) n"
+ then have "(LEAST n. x \<in> B (m i x) n) \<le> n"
+ by (intro Least_le)
+ also assume "n \<le> i"
+ finally have "(LEAST n. x \<in> B (m i x) n) \<le> i" . }
+ then have "F i ` space M \<subseteq> {z} \<union> 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 \<le> N" "n \<le> N" "x \<in> B i n"
+ then have 1: "\<exists>m\<le>N. x \<in> (\<Union>n\<le>N. B m n)" by auto
+ from m[OF this] obtain n where n: "m N x \<le> N" "n \<le> N" "x \<in> B (m N x) n" by auto
+ moreover
+ define L where "L = (LEAST n. x \<in> B (m N x) n)"
+ have "dist (f x) (e L) < 1 / Suc (m N x)"
+ proof -
+ have "x \<in> B (m N x) L"
+ using n(3) unfolding L_def by (rule LeastI)
+ then have "x \<in> 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 \<in> space M"
+ show "(\<lambda>i. F i x) \<longlonglongrightarrow> f x"
+ proof cases
+ assume "f x = z"
+ then have "\<And>i n. x \<notin> A i n"
+ unfolding A_def by auto
+ then have "\<And>i. F i x = z"
+ by (auto simp: F_def)
+ then show ?thesis
+ using \<open>f x = z\<close> by auto
+ next
+ assume "f x \<noteq> z"
+
+ show ?thesis
+ proof (rule tendstoI)
+ fix e :: real assume "0 < e"
+ with \<open>f x \<noteq> z\<close> 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 \<open>x\<in>space M\<close> \<open>f x \<noteq> z\<close> have "x \<in> (\<Union>i. B n i)"
+ unfolding A_def B_def UN_disjointed_eq using e by auto
+ then obtain i where i: "x \<in> B n i" by auto
+
+ show "eventually (\<lambda>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 \<le> j"
+ with i have "dist (f x) (F j x) < 1 / Suc (m j x)"
+ by (intro *[OF _ _ i]) auto
+ also have "\<dots> \<le> 1 / Suc n"
+ using j m_upper[OF _ _ i]
+ by (auto simp: field_simps)
+ also note \<open>1 / Suc n < e\<close>
+ 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 \<in> A n m"
+ then have "dist (e m) (f x) + dist (f x) z \<le> 2 * dist (f x) z"
+ unfolding A_def by (auto simp: dist_commute)
+ also have "dist (e m) z \<le> dist (e m) (f x) + dist (f x) z"
+ by (rule dist_triangle)
+ finally (xtrans) have "dist (e m) z \<le> 2 * dist (f x) z" . }
+ then show "dist (F i x) z \<le> 2 * dist (f x) z"
+ unfolding F_def
+ apply auto
+ apply (rule LeastI2)
+ apply auto
+ done
+ qed
+qed
+
+lemma
+ fixes f :: "'a \<Rightarrow> 'b::semiring_1" assumes "finite A"
+ shows setsum_mult_indicator[simp]: "(\<Sum>x \<in> A. f x * indicator (B x) (g x)) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
+ and setsum_indicator_mult[simp]: "(\<Sum>x \<in> A. indicator (B x) (g x) * f x) = (\<Sum>x\<in>{x\<in>A. g x \<in> B x}. f x)"
+ unfolding indicator_def
+ using assms by (auto intro!: setsum.mono_neutral_cong_right split: if_split_asm)
+
+lemma borel_measurable_induct_real[consumes 2, case_names set mult add seq]:
+ fixes P :: "('a \<Rightarrow> real) \<Rightarrow> bool"
+ assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
+ assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
+ assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+ assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
+ assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. U i x) \<longlonglongrightarrow> u x) \<Longrightarrow> P u"
+ shows "P u"
+proof -
+ have "(\<lambda>x. ennreal (u x)) \<in> borel_measurable M" using u by auto
+ from borel_measurable_implies_simple_function_sequence'[OF this]
+ obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i x. U i x < top" and
+ sup: "\<And>x. (SUP i. U i x) = ennreal (u x)"
+ by blast
+
+ define U' where [abs_def]: "U' i x = indicator (space M) x * enn2real (U i x)" for i x
+ then have U'_sf[measurable]: "\<And>i. simple_function M (U' i)"
+ using U by (auto intro!: simple_function_compose1[where g=enn2real])
+
+ show "P u"
+ proof (rule seq)
+ show U': "U' i \<in> borel_measurable M" "\<And>x. 0 \<le> U' i x" for i
+ using U by (auto
+ intro: borel_measurable_simple_function
+ intro!: borel_measurable_enn2real borel_measurable_times
+ simp: U'_def zero_le_mult_iff enn2real_nonneg)
+ show "incseq U'"
+ using U(2,3)
+ by (auto simp: incseq_def le_fun_def image_iff eq_commute U'_def indicator_def enn2real_mono)
+
+ fix x assume x: "x \<in> space M"
+ have "(\<lambda>i. U i x) \<longlonglongrightarrow> (SUP i. U i x)"
+ using U(2) by (intro LIMSEQ_SUP) (auto simp: incseq_def le_fun_def)
+ moreover have "(\<lambda>i. U i x) = (\<lambda>i. ennreal (U' i x))"
+ using x U(3) by (auto simp: fun_eq_iff U'_def image_iff eq_commute)
+ moreover have "(SUP i. U i x) = ennreal (u x)"
+ using sup u(2) by (simp add: max_def)
+ ultimately show "(\<lambda>i. U' i x) \<longlonglongrightarrow> u x"
+ using u U' by simp
+ next
+ fix i
+ have "U' i ` space M \<subseteq> enn2real ` (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 "\<And>z. {y. U' i z = y \<and> y \<in> U' i ` space M \<and> z \<in> space M} = (if z \<in> space M then {U' i z} else {})"
+ by auto
+ ultimately have U': "(\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>space M. U' i x = y} z) = U' i"
+ by (simp add: U'_def fun_eq_iff)
+ have "\<And>x. x \<in> U' i ` space M \<Longrightarrow> 0 \<le> x"
+ by (auto simp: U'_def enn2real_nonneg)
+ with fin have "P (\<lambda>z. \<Sum>y\<in>U' i`space M. y * indicator {x\<in>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 \<noteq> 0 \<Longrightarrow> r = p) \<Longrightarrow> r *\<^sub>R x = p *\<^sub>R x"
+ by (cases "x = 0") auto
+
+inductive simple_bochner_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> bool" for M f where
+ "simple_function M f \<Longrightarrow> emeasure M {y\<in>space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow>
+ simple_bochner_integrable M f"
+
+lemma simple_bochner_integrable_compose2:
+ assumes p_0: "p 0 0 = 0"
+ shows "simple_bochner_integrable M f \<Longrightarrow> simple_bochner_integrable M g \<Longrightarrow>
+ simple_bochner_integrable M (\<lambda>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 (\<lambda>x. p (f x) (g x))"
+ by (rule simple_function_compose2)
+
+ from sf have [measurable]:
+ "f \<in> measurable M (count_space UNIV)"
+ "g \<in> measurable M (count_space UNIV)"
+ by (auto intro: measurable_simple_function)
+
+ assume fin: "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" "emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity>"
+
+ have "emeasure M {x\<in>space M. p (f x) (g x) \<noteq> 0} \<le>
+ emeasure M ({x\<in>space M. f x \<noteq> 0} \<union> {x\<in>space M. g x \<noteq> 0})"
+ by (intro emeasure_mono) (auto simp: p_0)
+ also have "\<dots> \<le> emeasure M {x\<in>space M. f x \<noteq> 0} + emeasure M {x\<in>space M. g x \<noteq> 0}"
+ by (intro emeasure_subadditive) auto
+ finally show "emeasure M {y \<in> space M. p (f y) (g y) \<noteq> 0} \<noteq> \<infinity>"
+ using fin by (auto simp: top_unique)
+qed
+
+lemma simple_function_finite_support:
+ assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and nn: "\<And>x. 0 \<le> f x"
+ shows "emeasure M {x\<in>space M. f x \<noteq> 0} \<noteq> \<infinity>"
+proof cases
+ from f have meas[measurable]: "f \<in> borel_measurable M"
+ by (rule borel_measurable_simple_function)
+
+ assume non_empty: "\<exists>x\<in>space M. f x \<noteq> 0"
+
+ define m where "m = Min (f`space M - {0})"
+ have "m \<in> 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\<in>space M. 0 \<noteq> f x} =
+ (\<integral>\<^sup>+x. m * indicator {x\<in>space M. 0 \<noteq> f x} x \<partial>M)"
+ using f by (intro nn_integral_cmult_indicator[symmetric]) auto
+ also have "\<dots> \<le> (\<integral>\<^sup>+x. f x \<partial>M)"
+ using AE_space
+ proof (intro nn_integral_mono_AE, eventually_elim)
+ fix x assume "x \<in> space M"
+ with nn show "m * indicator {x \<in> space M. 0 \<noteq> f x} x \<le> f x"
+ using f by (auto split: split_indicator simp: simple_function_def m_def)
+ qed
+ also note \<open>\<dots> < \<infinity>\<close>
+ finally show ?thesis
+ using m by (auto simp: ennreal_mult_less_top)
+next
+ assume "\<not> (\<exists>x\<in>space M. f x \<noteq> 0)"
+ with nn have *: "{x\<in>space M. f x \<noteq> 0} = {}"
+ by auto
+ show ?thesis unfolding * by simp
+qed
+
+lemma simple_bochner_integrableI_bounded:
+ assumes f: "simple_function M f" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
+ shows "simple_bochner_integrable M f"
+proof
+ have "emeasure M {y \<in> space M. ennreal (norm (f y)) \<noteq> 0} \<noteq> \<infinity>"
+ proof (rule simple_function_finite_support)
+ show "simple_function M (\<lambda>x. ennreal (norm (f x)))"
+ using f by (rule simple_function_compose1)
+ show "(\<integral>\<^sup>+ y. ennreal (norm (f y)) \<partial>M) < \<infinity>" by fact
+ qed simp
+ then show "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" by simp
+qed fact
+
+definition simple_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b::real_vector) \<Rightarrow> 'b" where
+ "simple_bochner_integral M f = (\<Sum>y\<in>f`space M. measure M {x\<in>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: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y"
+ assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)"
+ shows "simple_bochner_integral M f = (\<Sum>y\<in>g ` space M. measure M {x\<in>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 \<in> measurable M (count_space UNIV)"
+ by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ from g have [measurable]: "g \<in> measurable M (count_space UNIV)"
+ by (auto intro: measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ { fix y assume "y \<in> space M"
+ then have "f ` space M \<inter> {i. \<exists>x\<in>space M. i = f x \<and> g y = g x} = {v (g y)}"
+ by (auto cong: sub simp: v[symmetric]) }
+ note eq = this
+
+ have "simple_bochner_integral M f =
+ (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
+ if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>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 \<in> space M" "f y \<noteq> 0"
+ have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
+ {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
+ by auto
+ have eq:"{x \<in> space M. f x = f y} =
+ (\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> 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. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
+ by (rule rev_finite_subset) auto
+ moreover
+ { fix x assume "x \<in> space M" "f x = f y"
+ then have "x \<in> space M" "f x \<noteq> 0"
+ using y by auto
+ then have "emeasure M {y \<in> space M. g y = g x} \<le> emeasure M {y \<in> space M. f y \<noteq> 0}"
+ by (auto intro!: emeasure_mono cong: sub)
+ then have "emeasure M {xa \<in> space M. g xa = g x} < \<infinity>"
+ using f by (auto simp: simple_bochner_integrable.simps less_top) }
+ ultimately
+ show "measure M {x \<in> space M. f x = f y} =
+ (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then measure M {x \<in> space M. g x = z} else 0)"
+ apply (simp add: setsum.If_cases eq)
+ apply (subst measure_finite_Union[symmetric])
+ apply (auto simp: disjoint_family_on_def less_top)
+ done
+ qed
+ also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M.
+ if \<exists>x\<in>space M. y = f x \<and> z = g x then measure M {x\<in>space M. g x = z} *\<^sub>R y else 0))"
+ by (auto intro!: setsum.cong simp: scaleR_setsum_left)
+ also have "\<dots> = ?r"
+ by (subst setsum.commute)
+ (auto intro!: setsum.cong simp: setsum.If_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 (\<lambda>x. f x + g x) =
+ simple_bochner_integral M f + simple_bochner_integral M g"
+proof -
+ from f g have "simple_bochner_integral M (\<lambda>x. f x + g x) =
+ (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> 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 =
+ (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> 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 =
+ (\<Sum>y\<in>(\<lambda>x. (f x, g x)) ` space M. measure M {x \<in> 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.distrib[symmetric] scaleR_add_right)
+qed
+
+lemma (in linear) simple_bochner_integral_linear:
+ assumes g: "simple_bochner_integrable M g"
+ shows "simple_bochner_integral M (\<lambda>x. f (g x)) = f (simple_bochner_integral M g)"
+proof -
+ from g have "simple_bochner_integral M (\<lambda>x. f (g x)) =
+ (\<Sum>y\<in>g ` space M. measure M {x \<in> space M. g x = y} *\<^sub>R f y)"
+ by (intro simple_bochner_integral_partition)
+ (auto simp: simple_bochner_integrable_compose2[where p="\<lambda>x y. f x"] zero
+ elim: simple_bochner_integrable.cases)
+ also have "\<dots> = 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 (\<lambda>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 (\<lambda>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="\<lambda>x y. - y"])
+
+lemma simple_bochner_integral_norm_bound:
+ assumes f: "simple_bochner_integrable M f"
+ shows "norm (simple_bochner_integral M f) \<le> simple_bochner_integral M (\<lambda>x. norm (f x))"
+proof -
+ have "norm (simple_bochner_integral M f) \<le>
+ (\<Sum>y\<in>f ` space M. norm (measure M {x \<in> space M. f x = y} *\<^sub>R y))"
+ unfolding simple_bochner_integral_def by (rule norm_setsum)
+ also have "\<dots> = (\<Sum>y\<in>f ` space M. measure M {x \<in> space M. f x = y} *\<^sub>R norm y)"
+ by simp
+ also have "\<dots> = simple_bochner_integral M (\<lambda>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_nonneg[simp]:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "(\<And>x. 0 \<le> f x) \<Longrightarrow> 0 \<le> simple_bochner_integral M f"
+ by (simp add: setsum_nonneg simple_bochner_integral_def)
+
+lemma simple_bochner_integral_eq_nn_integral:
+ assumes f: "simple_bochner_integrable M f" "\<And>x. 0 \<le> f x"
+ shows "simple_bochner_integral M f = (\<integral>\<^sup>+x. f x \<partial>M)"
+proof -
+ { fix x y z have "(x \<noteq> 0 \<Longrightarrow> y = z) \<Longrightarrow> ennreal x * y = ennreal x * z"
+ by (cases "x = 0") (auto simp: zero_ennreal_def[symmetric]) }
+ note ennreal_cong_mult = this
+
+ have [measurable]: "f \<in> borel_measurable M"
+ using f(1) by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ { fix y assume y: "y \<in> space M" "f y \<noteq> 0"
+ have "ennreal (measure M {x \<in> space M. f x = f y}) = emeasure M {x \<in> space M. f x = f y}"
+ proof (rule emeasure_eq_ennreal_measure[symmetric])
+ have "emeasure M {x \<in> space M. f x = f y} \<le> emeasure M {x \<in> space M. f x \<noteq> 0}"
+ using y by (intro emeasure_mono) auto
+ with f show "emeasure M {x \<in> space M. f x = f y} \<noteq> top"
+ by (auto simp: simple_bochner_integrable.simps top_unique)
+ qed
+ moreover have "{x \<in> space M. f x = f y} = (\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M"
+ using f by auto
+ ultimately have "ennreal (measure M {x \<in> space M. f x = f y}) =
+ emeasure M ((\<lambda>x. ennreal (f x)) -` {ennreal (f y)} \<inter> space M)" by simp }
+ with f have "simple_bochner_integral M f = (\<integral>\<^sup>Sx. f x \<partial>M)"
+ unfolding simple_integral_def
+ by (subst simple_bochner_integral_partition[OF f(1), where g="\<lambda>x. ennreal (f x)" and v=enn2real])
+ (auto intro: f simple_function_compose1 elim: simple_bochner_integrable.cases
+ intro!: setsum.cong ennreal_cong_mult
+ simp: setsum_ennreal[symmetric] ac_simps ennreal_mult
+ simp del: setsum_ennreal)
+ also have "\<dots> = (\<integral>\<^sup>+x. f x \<partial>M)"
+ using f
+ by (intro nn_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 \<Rightarrow> 'b::{real_normed_vector, second_countable_topology}"
+ assumes f[measurable]: "f \<in> borel_measurable M"
+ assumes s: "simple_bochner_integrable M s" and t: "simple_bochner_integrable M t"
+ shows "ennreal (norm (simple_bochner_integral M s - simple_bochner_integral M t)) \<le>
+ (\<integral>\<^sup>+ x. norm (f x - s x) \<partial>M) + (\<integral>\<^sup>+ x. norm (f x - t x) \<partial>M)"
+ (is "ennreal (norm (?s - ?t)) \<le> ?S + ?T")
+proof -
+ have [measurable]: "s \<in> borel_measurable M" "t \<in> borel_measurable M"
+ using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ have "ennreal (norm (?s - ?t)) = norm (simple_bochner_integral M (\<lambda>x. s x - t x))"
+ using s t by (subst simple_bochner_integral_diff) auto
+ also have "\<dots> \<le> simple_bochner_integral M (\<lambda>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 "\<dots> = (\<integral>\<^sup>+x. norm (s x - t x) \<partial>M)"
+ using simple_bochner_integrable_compose2[of "\<lambda>x y. norm (x - y)" M "s" "t"] s t
+ by (auto intro!: simple_bochner_integral_eq_nn_integral)
+ also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s x)) + ennreal (norm (f x - t x)) \<partial>M)"
+ by (auto intro!: nn_integral_mono simp: ennreal_plus[symmetric] simp del: ennreal_plus)
+ (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 "\<dots> = ?S + ?T"
+ by (rule nn_integral_add) auto
+ finally show ?thesis .
+qed
+
+inductive has_bochner_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b::{real_normed_vector, second_countable_topology} \<Rightarrow> bool"
+ for M f x where
+ "f \<in> borel_measurable M \<Longrightarrow>
+ (\<And>i. simple_bochner_integrable M (s i)) \<Longrightarrow>
+ (\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0 \<Longrightarrow>
+ (\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x \<Longrightarrow>
+ has_bochner_integral M f x"
+
+lemma has_bochner_integral_cong:
+ assumes "M = N" "\<And>x. x \<in> space N \<Longrightarrow> f x = g x" "x = y"
+ shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral N g y"
+ unfolding has_bochner_integral.simps assms(1,3)
+ using assms(2) by (simp cong: measurable_cong_strong nn_integral_cong_strong)
+
+lemma has_bochner_integral_cong_AE:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow>
+ has_bochner_integral M f x \<longleftrightarrow> 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="\<lambda>x. x \<longlonglongrightarrow> 0"]
+ nn_integral_cong_AE)
+ auto
+
+lemma borel_measurable_has_bochner_integral:
+ "has_bochner_integral M f x \<Longrightarrow> f \<in> borel_measurable M"
+ by (rule has_bochner_integral.cases)
+
+lemma borel_measurable_has_bochner_integral'[measurable_dest]:
+ "has_bochner_integral M f x \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
+ using borel_measurable_has_bochner_integral[measurable] by measurable
+
+lemma has_bochner_integral_simple_bochner_integrable:
+ "simple_bochner_integrable M f \<Longrightarrow> has_bochner_integral M f (simple_bochner_integral M f)"
+ by (rule has_bochner_integral.intros[where s="\<lambda>_. f"])
+ (auto intro: borel_measurable_simple_function
+ elim: simple_bochner_integrable.cases
+ simp: zero_ennreal_def[symmetric])
+
+lemma has_bochner_integral_real_indicator:
+ assumes [measurable]: "A \<in> sets M" and A: "emeasure M A < \<infinity>"
+ shows "has_bochner_integral M (indicator A) (measure M A)"
+proof -
+ have sbi: "simple_bochner_integrable M (indicator A::'a \<Rightarrow> real)"
+ proof
+ have "{y \<in> space M. (indicator A y::real) \<noteq> 0} = A"
+ using sets.sets_into_space[OF \<open>A\<in>sets M\<close>] by (auto split: split_indicator)
+ then show "emeasure M {y \<in> space M. (indicator A y::real) \<noteq> 0} \<noteq> \<infinity>"
+ 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_nn_integral[OF sbi] A
+ by (simp add: ennreal_indicator emeasure_eq_ennreal_measure)
+ ultimately show ?thesis
+ by (metis has_bochner_integral_simple_bochner_integrable)
+qed
+
+lemma has_bochner_integral_add[intro]:
+ "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow>
+ has_bochner_integral M (\<lambda>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: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - sf i x) \<partial>M) \<longlonglongrightarrow> 0"
+ assume g_sg: "(\<lambda>i. \<integral>\<^sup>+ x. norm (g x - sg i x) \<partial>M) \<longlonglongrightarrow> 0"
+
+ assume sf: "\<forall>i. simple_bochner_integrable M (sf i)"
+ and sg: "\<forall>i. simple_bochner_integrable M (sg i)"
+ then have [measurable]: "\<And>i. sf i \<in> borel_measurable M" "\<And>i. sg i \<in> borel_measurable M"
+ by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+ assume [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+
+ show "\<And>i. simple_bochner_integrable M (\<lambda>x. sf i x + sg i x)"
+ using sf sg by (simp add: simple_bochner_integrable_compose2)
+
+ show "(\<lambda>i. \<integral>\<^sup>+ x. (norm (f x + g x - (sf i x + sg i x))) \<partial>M) \<longlonglongrightarrow> 0"
+ (is "?f \<longlonglongrightarrow> 0")
+ proof (rule tendsto_sandwich)
+ show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0"
+ by auto
+ show "eventually (\<lambda>i. ?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) \<partial>M) + \<integral>\<^sup>+ x. (norm (g x - sg i x)) \<partial>M) sequentially"
+ (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
+ proof (intro always_eventually allI)
+ fix i have "?f i \<le> (\<integral>\<^sup>+ x. (norm (f x - sf i x)) + ennreal (norm (g x - sg i x)) \<partial>M)"
+ by (auto intro!: nn_integral_mono norm_diff_triangle_ineq
+ simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+ also have "\<dots> = ?g i"
+ by (intro nn_integral_add) auto
+ finally show "?f i \<le> ?g i" .
+ qed
+ show "?g \<longlonglongrightarrow> 0"
+ using tendsto_add[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 \<Longrightarrow> has_bochner_integral M (\<lambda>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 \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 T.continuous_on continuous_on_id)
+ assume [measurable]: "f \<in> borel_measurable M"
+ then show "(\<lambda>x. T (f x)) \<in> borel_measurable M"
+ by auto
+
+ fix s assume f_s: "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0"
+ assume s: "\<forall>i. simple_bochner_integrable M (s i)"
+ then show "\<And>i. simple_bochner_integrable M (\<lambda>x. T (s i x))"
+ by (auto intro: simple_bochner_integrable_compose2 T.zero)
+
+ have [measurable]: "\<And>i. s i \<in> borel_measurable M"
+ using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ obtain K where K: "K > 0" "\<And>x i. norm (T (f x) - T (s i x)) \<le> norm (f x - s i x) * K"
+ using T.pos_bounded by (auto simp: T.diff[symmetric])
+
+ show "(\<lambda>i. \<integral>\<^sup>+ x. norm (T (f x) - T (s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+ (is "?f \<longlonglongrightarrow> 0")
+ proof (rule tendsto_sandwich)
+ show "eventually (\<lambda>n. 0 \<le> ?f n) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> 0"
+ by auto
+
+ show "eventually (\<lambda>i. ?f i \<le> K * (\<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M)) sequentially"
+ (is "eventually (\<lambda>i. ?f i \<le> ?g i) sequentially")
+ proof (intro always_eventually allI)
+ fix i have "?f i \<le> (\<integral>\<^sup>+ x. ennreal K * norm (f x - s i x) \<partial>M)"
+ using K by (intro nn_integral_mono) (auto simp: ac_simps ennreal_mult[symmetric])
+ also have "\<dots> = ?g i"
+ using K by (intro nn_integral_cmult) auto
+ finally show "?f i \<le> ?g i" .
+ qed
+ show "?g \<longlonglongrightarrow> 0"
+ using ennreal_tendsto_cmult[OF _ f_s] by simp
+ qed
+
+ assume "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x"
+ with s show "(\<lambda>i. simple_bochner_integral M (\<lambda>x. T (s i x))) \<longlonglongrightarrow> T x"
+ by (auto intro!: T.tendsto simp: T.simple_bochner_integral_linear)
+qed
+
+lemma has_bochner_integral_zero[intro]: "has_bochner_integral M (\<lambda>x. 0) 0"
+ by (auto intro!: has_bochner_integral.intros[where s="\<lambda>_ _. 0"]
+ simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps
+ simple_bochner_integral_def image_constant_conv)
+
+lemma has_bochner_integral_scaleR_left[intro]:
+ "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>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 \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>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 \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>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 \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>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, second_countable_topology}"
+ shows "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x / c) (x / c)"
+ using has_bochner_integral_divide by (cases "c = 0") auto
+
+lemma has_bochner_integral_inner_left[intro]:
+ "(c \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. f x \<bullet> c) (x \<bullet> 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 \<noteq> 0 \<Longrightarrow> has_bochner_integral M f x) \<Longrightarrow> has_bochner_integral M (\<lambda>x. c \<bullet> f x) (c \<bullet> 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 \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+ has_bochner_integral M (\<lambda>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 \<Longrightarrow> has_bochner_integral M g y \<Longrightarrow>
+ has_bochner_integral M (\<lambda>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:
+ "(\<And>i. i \<in> I \<Longrightarrow> has_bochner_integral M (f i) (x i)) \<Longrightarrow>
+ has_bochner_integral M (\<lambda>x. \<Sum>i\<in>I. f i x) (\<Sum>i\<in>I. x i)"
+ by (induct I rule: infinite_finite_induct) auto
+
+lemma has_bochner_integral_implies_finite_norm:
+ "has_bochner_integral M f x \<Longrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
+proof (elim has_bochner_integral.cases)
+ fix s v
+ assume [measurable]: "f \<in> borel_measurable M" and s: "\<And>i. simple_bochner_integrable M (s i)" and
+ lim_0: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+ from order_tendstoD[OF lim_0, of "\<infinity>"]
+ obtain i where f_s_fin: "(\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) < \<infinity>"
+ by (auto simp: eventually_sequentially)
+
+ have [measurable]: "\<And>i. s i \<in> borel_measurable M"
+ using s by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ define m where "m = (if space M = {} then 0 else Max ((\<lambda>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 "\<And>x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> m" "0 \<le> m"
+ by (auto simp: m_def image_comp comp_def Max_ge_iff)
+ then have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal m * indicator {x\<in>space M. s i x \<noteq> 0} x \<partial>M)"
+ by (auto split: split_indicator intro!: Max_ge nn_integral_mono simp:)
+ also have "\<dots> < \<infinity>"
+ using s by (subst nn_integral_cmult_indicator) (auto simp: \<open>0 \<le> m\<close> simple_bochner_integrable.simps ennreal_mult_less_top less_top)
+ finally have s_fin: "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>" .
+
+ have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) + ennreal (norm (s i x)) \<partial>M)"
+ by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+ (metis add.commute norm_triangle_sub)
+ also have "\<dots> = (\<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) + (\<integral>\<^sup>+x. norm (s i x) \<partial>M)"
+ by (rule nn_integral_add) auto
+ also have "\<dots> < \<infinity>"
+ using s_fin f_s_fin by auto
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
+qed
+
+lemma has_bochner_integral_norm_bound:
+ assumes i: "has_bochner_integral M f x"
+ shows "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
+using assms proof
+ fix s assume
+ x: "(\<lambda>i. simple_bochner_integral M (s i)) \<longlonglongrightarrow> x" (is "?s \<longlonglongrightarrow> x") and
+ s[simp]: "\<And>i. simple_bochner_integrable M (s i)" and
+ lim: "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0" and
+ f[measurable]: "f \<in> borel_measurable M"
+
+ have [measurable]: "\<And>i. s i \<in> borel_measurable M"
+ using s by (auto simp: simple_bochner_integrable.simps intro: borel_measurable_simple_function)
+
+ show "norm x \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
+ proof (rule LIMSEQ_le)
+ show "(\<lambda>i. ennreal (norm (?s i))) \<longlonglongrightarrow> norm x"
+ using x by (auto simp: tendsto_ennreal_iff intro: tendsto_intros)
+ show "\<exists>N. \<forall>n\<ge>N. norm (?s n) \<le> (\<integral>\<^sup>+x. norm (f x - s n x) \<partial>M) + (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
+ (is "\<exists>N. \<forall>n\<ge>N. _ \<le> ?t n")
+ proof (intro exI allI impI)
+ fix n
+ have "ennreal (norm (?s n)) \<le> simple_bochner_integral M (\<lambda>x. norm (s n x))"
+ by (auto intro!: simple_bochner_integral_norm_bound)
+ also have "\<dots> = (\<integral>\<^sup>+x. norm (s n x) \<partial>M)"
+ by (intro simple_bochner_integral_eq_nn_integral)
+ (auto intro: s simple_bochner_integrable_compose2)
+ also have "\<dots> \<le> (\<integral>\<^sup>+x. ennreal (norm (f x - s n x)) + norm (f x) \<partial>M)"
+ by (auto intro!: nn_integral_mono simp del: ennreal_plus simp add: ennreal_plus[symmetric])
+ (metis add.commute norm_minus_commute norm_triangle_sub)
+ also have "\<dots> = ?t n"
+ by (rule nn_integral_add) auto
+ finally show "norm (?s n) \<le> ?t n" .
+ qed
+ have "?t \<longlonglongrightarrow> 0 + (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
+ using has_bochner_integral_implies_finite_norm[OF i]
+ by (intro tendsto_add tendsto_const lim)
+ then show "?t \<longlonglongrightarrow> \<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M"
+ by simp
+ qed
+qed
+
+lemma has_bochner_integral_eq:
+ "has_bochner_integral M f x \<Longrightarrow> has_bochner_integral M f y \<Longrightarrow> x = y"
+proof (elim has_bochner_integral.cases)
+ assume f[measurable]: "f \<in> borel_measurable M"
+
+ fix s t
+ assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0")
+ assume "(\<lambda>i. \<integral>\<^sup>+ x. norm (f x - t i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?T \<longlonglongrightarrow> 0")
+ assume s: "\<And>i. simple_bochner_integrable M (s i)"
+ assume t: "\<And>i. simple_bochner_integrable M (t i)"
+
+ have [measurable]: "\<And>i. s i \<in> borel_measurable M" "\<And>i. t i \<in> borel_measurable M"
+ using s t by (auto intro: borel_measurable_simple_function elim: simple_bochner_integrable.cases)
+
+ let ?s = "\<lambda>i. simple_bochner_integral M (s i)"
+ let ?t = "\<lambda>i. simple_bochner_integral M (t i)"
+ assume "?s \<longlonglongrightarrow> x" "?t \<longlonglongrightarrow> y"
+ then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> norm (x - y)"
+ by (intro tendsto_intros)
+ moreover
+ have "(\<lambda>i. ennreal (norm (?s i - ?t i))) \<longlonglongrightarrow> ennreal 0"
+ proof (rule tendsto_sandwich)
+ show "eventually (\<lambda>i. 0 \<le> ennreal (norm (?s i - ?t i))) sequentially" "(\<lambda>_. 0) \<longlonglongrightarrow> ennreal 0"
+ by auto
+
+ show "eventually (\<lambda>i. norm (?s i - ?t i) \<le> ?S i + ?T i) sequentially"
+ by (intro always_eventually allI simple_bochner_integral_bounded s t f)
+ show "(\<lambda>i. ?S i + ?T i) \<longlonglongrightarrow> ennreal 0"
+ using tendsto_add[OF \<open>?S \<longlonglongrightarrow> 0\<close> \<open>?T \<longlonglongrightarrow> 0\<close>] by simp
+ qed
+ then have "(\<lambda>i. norm (?s i - ?t i)) \<longlonglongrightarrow> 0"
+ by (simp add: ennreal_0[symmetric] del: ennreal_0)
+ 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 \<in> 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 "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+ also have "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) = (\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M)"
+ using ae
+ by (intro ext nn_integral_cong_AE, eventually_elim) simp
+ finally show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (g x - s i x)) \<partial>M) \<longlonglongrightarrow> 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
+
+lemma simple_bochner_integrable_restrict_space:
+ fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
+ assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
+ shows "simple_bochner_integrable (restrict_space M \<Omega>) f \<longleftrightarrow>
+ simple_bochner_integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
+ by (simp add: simple_bochner_integrable.simps space_restrict_space
+ simple_function_restrict_space[OF \<Omega>] emeasure_restrict_space[OF \<Omega>] Collect_restrict
+ indicator_eq_0_iff conj_ac)
+
+lemma simple_bochner_integral_restrict_space:
+ fixes f :: "_ \<Rightarrow> 'b::real_normed_vector"
+ assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M"
+ assumes f: "simple_bochner_integrable (restrict_space M \<Omega>) f"
+ shows "simple_bochner_integral (restrict_space M \<Omega>) f =
+ simple_bochner_integral M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
+proof -
+ have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x)`space M)"
+ using f simple_bochner_integrable_restrict_space[OF \<Omega>, of f]
+ by (simp add: simple_bochner_integrable.simps simple_function_def)
+ then show ?thesis
+ by (auto simp: space_restrict_space measure_restrict_space[OF \<Omega>(1)] le_infI2
+ simple_bochner_integral_def Collect_restrict
+ split: split_indicator split_indicator_asm
+ intro!: setsum.mono_neutral_cong_left arg_cong2[where f=measure])
+qed
+
+context
+ notes [[inductive_internals]]
+begin
+
+inductive integrable for M f where
+ "has_bochner_integral M f x \<Longrightarrow> integrable M f"
+
+end
+
+definition lebesgue_integral ("integral\<^sup>L") where
+ "integral\<^sup>L M f = (if \<exists>x. has_bochner_integral M f x then THE x. has_bochner_integral M f x else 0)"
+
+syntax
+ "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral>((2 _./ _)/ \<partial>_)" [60,61] 110)
+
+translations
+ "\<integral> x. f \<partial>M" == "CONST lebesgue_integral M (\<lambda>x. f)"
+
+syntax
+ "_ascii_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" ("(3LINT (1_)/|(_)./ _)" [0,110,60] 60)
+
+translations
+ "LINT x|M. f" == "CONST lebesgue_integral M (\<lambda>x. f)"
+
+lemma has_bochner_integral_integral_eq: "has_bochner_integral M f x \<Longrightarrow> 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 \<Longrightarrow> 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 \<longleftrightarrow> integrable M f \<and> 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 \<Longrightarrow> 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: "\<not> integrable M f \<Longrightarrow> integral\<^sup>L M f = 0"
+ unfolding integrable.simps lebesgue_integral_def by (auto intro!: arg_cong[where f=The])
+
+lemma integral_eq_cases:
+ "integrable M f \<longleftrightarrow> integrable N g \<Longrightarrow>
+ (integrable M f \<Longrightarrow> integrable N g \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g) \<Longrightarrow>
+ 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 \<Longrightarrow> f \<in> borel_measurable M"
+ by (auto elim: integrable.cases has_bochner_integral.cases)
+
+lemma borel_measurable_integrable'[measurable_dest]:
+ "integrable M f \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
+ using borel_measurable_integrable[measurable] by measurable
+
+lemma integrable_cong:
+ "M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable N g"
+ by (simp cong: has_bochner_integral_cong add: integrable.simps)
+
+lemma integrable_cong_AE:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow>
+ integrable M f \<longleftrightarrow> integrable M g"
+ unfolding integrable.simps
+ by (intro has_bochner_integral_cong_AE arg_cong[where f=Ex] ext)
+
+lemma integral_cong:
+ "M = N \<Longrightarrow> (\<And>x. x \<in> space N \<Longrightarrow> f x = g x) \<Longrightarrow> integral\<^sup>L M f = integral\<^sup>L N g"
+ by (simp cong: has_bochner_integral_cong cong del: if_weak_cong add: lebesgue_integral_def)
+
+lemma integral_cong_AE:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. f x = g x \<Longrightarrow>
+ integral\<^sup>L M f = integral\<^sup>L M g"
+ unfolding lebesgue_integral_def
+ by (rule arg_cong[where x="has_bochner_integral M f"]) (intro has_bochner_integral_cong_AE ext)
+
+lemma integrable_add[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x + g x)"
+ by (auto simp: integrable.simps)
+
+lemma integrable_zero[simp, intro]: "integrable M (\<lambda>x. 0)"
+ by (metis has_bochner_integral_zero integrable.simps)
+
+lemma integrable_setsum[simp, intro]: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow> integrable M (\<lambda>x. \<Sum>i\<in>I. f i x)"
+ by (metis has_bochner_integral_setsum integrable.simps)
+
+lemma integrable_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+ integrable M (\<lambda>x. indicator A x *\<^sub>R c)"
+ by (metis has_bochner_integral_indicator integrable.simps)
+
+lemma integrable_real_indicator[simp, intro]: "A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow>
+ integrable M (indicator A :: 'a \<Rightarrow> real)"
+ by (metis has_bochner_integral_real_indicator integrable.simps)
+
+lemma integrable_diff[simp, intro]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> integrable M (\<lambda>x. f x - g x)"
+ by (auto simp: integrable.simps intro: has_bochner_integral_diff)
+
+lemma integrable_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. T (f x))"
+ by (auto simp: integrable.simps intro: has_bochner_integral_bounded_linear)
+
+lemma integrable_scaleR_left[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x *\<^sub>R c)"
+ unfolding integrable.simps by fastforce
+
+lemma integrable_scaleR_right[simp, intro]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>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 \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>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 \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c * f x)"
+ unfolding integrable.simps by fastforce
+
+lemma integrable_divide_zero[simp, intro]:
+ fixes c :: "_::{real_normed_field, field, second_countable_topology}"
+ shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x / c)"
+ unfolding integrable.simps by fastforce
+
+lemma integrable_inner_left[simp, intro]:
+ "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. f x \<bullet> c)"
+ unfolding integrable.simps by fastforce
+
+lemma integrable_inner_right[simp, intro]:
+ "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> integrable M (\<lambda>x. c \<bullet> 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 (\<lambda>x. 0) = 0"
+ by (intro has_bochner_integral_integral_eq has_bochner_integral_zero)
+
+lemma integral_add[simp]: "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow>
+ integral\<^sup>L M (\<lambda>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 \<Longrightarrow> integrable M g \<Longrightarrow>
+ integral\<^sup>L M (\<lambda>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: "(\<And>i. i \<in> I \<Longrightarrow> integrable M (f i)) \<Longrightarrow>
+ integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
+ by (intro has_bochner_integral_integral_eq has_bochner_integral_setsum has_bochner_integral_integrable)
+
+lemma integral_setsum'[simp]: "(\<And>i. i \<in> I =simp=> integrable M (f i)) \<Longrightarrow>
+ integral\<^sup>L M (\<lambda>x. \<Sum>i\<in>I. f i x) = (\<Sum>i\<in>I. integral\<^sup>L M (f i))"
+ unfolding simp_implies_def by (rule integral_setsum)
+
+lemma integral_bounded_linear: "bounded_linear T \<Longrightarrow> integrable M f \<Longrightarrow>
+ integral\<^sup>L M (\<lambda>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_bounded_linear':
+ assumes T: "bounded_linear T" and T': "bounded_linear T'"
+ assumes *: "\<not> (\<forall>x. T x = 0) \<Longrightarrow> (\<forall>x. T' (T x) = x)"
+ shows "integral\<^sup>L M (\<lambda>x. T (f x)) = T (integral\<^sup>L M f)"
+proof cases
+ assume "(\<forall>x. T x = 0)" then show ?thesis
+ by simp
+next
+ assume **: "\<not> (\<forall>x. T x = 0)"
+ show ?thesis
+ proof cases
+ assume "integrable M f" with T show ?thesis
+ by (rule integral_bounded_linear)
+ next
+ assume not: "\<not> integrable M f"
+ moreover have "\<not> integrable M (\<lambda>x. T (f x))"
+ proof
+ assume "integrable M (\<lambda>x. T (f x))"
+ from integrable_bounded_linear[OF T' this] not *[OF **]
+ show False
+ by auto
+ qed
+ ultimately show ?thesis
+ using T by (simp add: not_integrable_integral_eq linear_simps)
+ qed
+qed
+
+lemma integral_scaleR_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x *\<^sub>R c \<partial>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]: "(\<integral> x. c *\<^sub>R f x \<partial>M) = c *\<^sub>R integral\<^sup>L M f"
+ by (rule integral_bounded_linear'[OF bounded_linear_scaleR_right bounded_linear_scaleR_right[of "1 / c"]]) simp
+
+lemma integral_mult_left[simp]:
+ fixes c :: "_::{real_normed_algebra,second_countable_topology}"
+ shows "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x * c \<partial>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 \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c * f x \<partial>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_mult_left_zero[simp]:
+ fixes c :: "_::{real_normed_field,second_countable_topology}"
+ shows "(\<integral> x. f x * c \<partial>M) = integral\<^sup>L M f * c"
+ by (rule integral_bounded_linear'[OF bounded_linear_mult_left bounded_linear_mult_left[of "1 / c"]]) simp
+
+lemma integral_mult_right_zero[simp]:
+ fixes c :: "_::{real_normed_field,second_countable_topology}"
+ shows "(\<integral> x. c * f x \<partial>M) = c * integral\<^sup>L M f"
+ by (rule integral_bounded_linear'[OF bounded_linear_mult_right bounded_linear_mult_right[of "1 / c"]]) simp
+
+lemma integral_inner_left[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. f x \<bullet> c \<partial>M) = integral\<^sup>L M f \<bullet> c"
+ by (intro has_bochner_integral_integral_eq has_bochner_integral_integrable has_bochner_integral_inner_left)
+
+lemma integral_inner_right[simp]: "(c \<noteq> 0 \<Longrightarrow> integrable M f) \<Longrightarrow> (\<integral> x. c \<bullet> f x \<partial>M) = c \<bullet> 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, second_countable_topology}"
+ shows "integral\<^sup>L M (\<lambda>x. f x / c) = integral\<^sup>L M f / c"
+ by (rule integral_bounded_linear'[OF bounded_linear_divide bounded_linear_mult_left[of c]]) simp
+
+lemma integral_minus[simp]: "integral\<^sup>L M (\<lambda>x. - f x) = - integral\<^sup>L M f"
+ by (rule integral_bounded_linear'[OF bounded_linear_minus[OF bounded_linear_ident] bounded_linear_minus[OF bounded_linear_ident]]) simp
+
+lemma integral_complex_of_real[simp]: "integral\<^sup>L M (\<lambda>x. complex_of_real (f x)) = of_real (integral\<^sup>L M f)"
+ by (rule integral_bounded_linear'[OF bounded_linear_of_real bounded_linear_Re]) simp
+
+lemma integral_cnj[simp]: "integral\<^sup>L M (\<lambda>x. cnj (f x)) = cnj (integral\<^sup>L M f)"
+ by (rule integral_bounded_linear'[OF bounded_linear_cnj bounded_linear_cnj]) simp
+
+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_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_ennreal:
+ "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>\<^sup>+x. norm (f x) \<partial>M)"
+ by (metis has_bochner_integral_integrable has_bochner_integral_norm_bound)
+
+lemma integrableI_sequence:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes f[measurable]: "f \<in> borel_measurable M"
+ assumes s: "\<And>i. simple_bochner_integrable M (s i)"
+ assumes lim: "(\<lambda>i. \<integral>\<^sup>+x. norm (f x - s i x) \<partial>M) \<longlonglongrightarrow> 0" (is "?S \<longlonglongrightarrow> 0")
+ shows "integrable M f"
+proof -
+ let ?s = "\<lambda>n. simple_bochner_integral M (s n)"
+
+ have "\<exists>x. ?s \<longlonglongrightarrow> x"
+ unfolding convergent_eq_cauchy
+ proof (rule metric_CauchyI)
+ fix e :: real assume "0 < e"
+ then have "0 < ennreal (e / 2)" by auto
+ from order_tendstoD(2)[OF lim this]
+ obtain M where M: "\<And>n. M \<le> n \<Longrightarrow> ?S n < e / 2"
+ by (auto simp: eventually_sequentially)
+ show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (?s m) (?s n) < e"
+ proof (intro exI allI impI)
+ fix m n assume m: "M \<le> m" and n: "M \<le> n"
+ have "?S n \<noteq> \<infinity>"
+ using M[OF n] by auto
+ have "norm (?s n - ?s m) \<le> ?S n + ?S m"
+ by (intro simple_bochner_integral_bounded s f)
+ also have "\<dots> < ennreal (e / 2) + e / 2"
+ by (intro add_strict_mono M n m)
+ also have "\<dots> = e" using \<open>0<e\<close> by (simp del: ennreal_plus add: ennreal_plus[symmetric])
+ finally show "dist (?s n) (?s m) < e"
+ using \<open>0<e\<close> by (simp add: dist_norm ennreal_less_iff)
+ qed
+ qed
+ then obtain x where "?s \<longlonglongrightarrow> x" ..
+ show ?thesis
+ by (rule, rule) fact+
+qed
+
+lemma nn_integral_dominated_convergence_norm:
+ fixes u' :: "_ \<Rightarrow> _::{real_normed_vector, second_countable_topology}"
+ assumes [measurable]:
+ "\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M"
+ and bound: "\<And>j. AE x in M. norm (u j x) \<le> w x"
+ and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
+ and u': "AE x in M. (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
+ shows "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> 0"
+proof -
+ have "AE x in M. \<forall>j. norm (u j x) \<le> w x"
+ unfolding AE_all_countable by rule fact
+ with u' have bnd: "AE x in M. \<forall>j. norm (u' x - u j x) \<le> 2 * w x"
+ proof (eventually_elim, intro allI)
+ fix i x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x" "\<forall>j. norm (u j x) \<le> w x" "\<forall>j. norm (u j x) \<le> w x"
+ then have "norm (u' x) \<le> w x" "norm (u i x) \<le> w x"
+ by (auto intro: LIMSEQ_le_const2 tendsto_norm)
+ then have "norm (u' x) + norm (u i x) \<le> 2 * w x"
+ by simp
+ also have "norm (u' x - u i x) \<le> norm (u' x) + norm (u i x)"
+ by (rule norm_triangle_ineq4)
+ finally (xtrans) show "norm (u' x - u i x) \<le> 2 * w x" .
+ qed
+ have w_nonneg: "AE x in M. 0 \<le> w x"
+ using bound[of 0] by (auto intro: order_trans[OF norm_ge_zero])
+
+ have "(\<lambda>i. (\<integral>\<^sup>+x. norm (u' x - u i x) \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. 0 \<partial>M)"
+ proof (rule nn_integral_dominated_convergence)
+ show "(\<integral>\<^sup>+x. 2 * w x \<partial>M) < \<infinity>"
+ by (rule nn_integral_mult_bounded_inf[OF _ w, of 2]) (insert w_nonneg, auto simp: ennreal_mult )
+ show "AE x in M. (\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
+ using u'
+ proof eventually_elim
+ fix x assume "(\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
+ from tendsto_diff[OF tendsto_const[of "u' x"] this]
+ show "(\<lambda>i. ennreal (norm (u' x - u i x))) \<longlonglongrightarrow> 0"
+ by (simp add: tendsto_norm_zero_iff ennreal_0[symmetric] del: ennreal_0)
+ qed
+ qed (insert bnd w_nonneg, auto)
+ then show ?thesis by simp
+qed
+
+lemma integrableI_bounded:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes f[measurable]: "f \<in> borel_measurable M" and fin: "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
+ shows "integrable M f"
+proof -
+ from borel_measurable_implies_sequence_metric[OF f, of 0] obtain s where
+ s: "\<And>i. simple_function M (s i)" and
+ pointwise: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" and
+ bound: "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)"
+ by simp metis
+
+ show ?thesis
+ proof (rule integrableI_sequence)
+ { fix i
+ have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)"
+ by (intro nn_integral_mono) (simp add: bound)
+ also have "\<dots> = 2 * (\<integral>\<^sup>+x. ennreal (norm (f x)) \<partial>M)"
+ by (simp add: ennreal_mult nn_integral_cmult)
+ also have "\<dots> < top"
+ using fin by (simp add: ennreal_mult_less_top)
+ finally have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) < \<infinity>"
+ by simp }
+ note fin_s = this
+
+ show "\<And>i. simple_bochner_integrable M (s i)"
+ by (rule simple_bochner_integrableI_bounded) fact+
+
+ show "(\<lambda>i. \<integral>\<^sup>+ x. ennreal (norm (f x - s i x)) \<partial>M) \<longlonglongrightarrow> 0"
+ proof (rule nn_integral_dominated_convergence_norm)
+ show "\<And>j. AE x in M. norm (s j x) \<le> 2 * norm (f x)"
+ using bound by auto
+ show "\<And>i. s i \<in> borel_measurable M" "(\<lambda>x. 2 * norm (f x)) \<in> borel_measurable M"
+ using s by (auto intro: borel_measurable_simple_function)
+ show "(\<integral>\<^sup>+ x. ennreal (2 * norm (f x)) \<partial>M) < \<infinity>"
+ using fin by (simp add: nn_integral_cmult ennreal_mult ennreal_mult_less_top)
+ show "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
+ using pointwise by auto
+ qed fact
+ qed fact
+qed
+
+lemma integrableI_bounded_set:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M"
+ assumes finite: "emeasure M A < \<infinity>"
+ and bnd: "AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B"
+ and null: "AE x in M. x \<notin> A \<longrightarrow> f x = 0"
+ shows "integrable M f"
+proof (rule integrableI_bounded)
+ { fix x :: 'b have "norm x \<le> B \<Longrightarrow> 0 \<le> B"
+ using norm_ge_zero[of x] by arith }
+ with bnd null have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (max 0 B) * indicator A x \<partial>M)"
+ by (intro nn_integral_mono_AE) (auto split: split_indicator split_max)
+ also have "\<dots> < \<infinity>"
+ using finite by (subst nn_integral_cmult_indicator) (auto simp: ennreal_mult_less_top)
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
+qed simp
+
+lemma integrableI_bounded_set_indicator:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "A \<in> sets M \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow>
+ emeasure M A < \<infinity> \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> norm (f x) \<le> B) \<Longrightarrow>
+ integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
+ by (rule integrableI_bounded_set[where A=A]) auto
+
+lemma integrableI_nonneg:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>"
+ shows "integrable M f"
+proof -
+ have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
+ using assms by (intro nn_integral_cong_AE) auto
+ then show ?thesis
+ using assms by (intro integrableI_bounded) auto
+qed
+
+lemma integrable_iff_bounded:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and> (\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
+ 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 \<Rightarrow> 'b::{banach, second_countable_topology}"
+ and g :: "'a \<Rightarrow> 'c::{banach, second_countable_topology}"
+ shows "integrable M f \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. norm (g x) \<le> norm (f x)) \<Longrightarrow>
+ integrable M g"
+ unfolding integrable_iff_bounded
+proof safe
+ assume "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ assume "AE x in M. norm (g x) \<le> norm (f x)"
+ then have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
+ by (intro nn_integral_mono_AE) auto
+ also assume "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>" .
+qed
+
+lemma integrable_mult_indicator:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
+ by (rule integrable_bound[of M f]) (auto split: split_indicator)
+
+lemma integrable_real_mult_indicator:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "A \<in> sets M \<Longrightarrow> integrable M f \<Longrightarrow> integrable M (\<lambda>x. f x * indicator A x)"
+ using integrable_mult_indicator[of A M f] by (simp add: mult_ac)
+
+lemma integrable_abs[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. \<bar>f x\<bar>)"
+ using assms by (rule integrable_bound) auto
+
+lemma integrable_norm[simp, intro]:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes [measurable]: "integrable M f" shows "integrable M (\<lambda>x. norm (f x))"
+ using assms by (rule integrable_bound) auto
+
+lemma integrable_norm_cancel:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes [measurable]: "integrable M (\<lambda>x. norm (f x))" "f \<in> borel_measurable M" shows "integrable M f"
+ using assms by (rule integrable_bound) auto
+
+lemma integrable_norm_iff:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. norm (f x)) \<longleftrightarrow> integrable M f"
+ by (auto intro: integrable_norm_cancel)
+
+lemma integrable_abs_cancel:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes [measurable]: "integrable M (\<lambda>x. \<bar>f x\<bar>)" "f \<in> borel_measurable M" shows "integrable M f"
+ using assms by (rule integrable_bound) auto
+
+lemma integrable_abs_iff:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda>x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f"
+ by (auto intro: integrable_abs_cancel)
+
+lemma integrable_max[simp, intro]:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes fg[measurable]: "integrable M f" "integrable M g"
+ shows "integrable M (\<lambda>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 \<Rightarrow> real"
+ assumes fg[measurable]: "integrable M f" "integrable M g"
+ shows "integrable M (\<lambda>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 (\<lambda>x. - f x ::'a::{banach, second_countable_topology}) \<longleftrightarrow> integrable M f"
+ unfolding integrable_iff_bounded
+ by (auto intro: borel_measurable_uminus[of "\<lambda>x. - f x" M, simplified])
+
+lemma integrable_indicator_iff:
+ "integrable M (indicator A::_ \<Rightarrow> real) \<longleftrightarrow> A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
+ by (simp add: integrable_iff_bounded borel_measurable_indicator_iff ennreal_indicator nn_integral_indicator'
+ cong: conj_cong)
+
+lemma integral_indicator[simp]: "integral\<^sup>L M (indicator A) = measure M (A \<inter> space M)"
+proof cases
+ assume *: "A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>"
+ have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M))"
+ by (intro integral_cong) (auto split: split_indicator)
+ also have "\<dots> = measure M (A \<inter> space M)"
+ using * by (intro has_bochner_integral_integral_eq has_bochner_integral_real_indicator) auto
+ finally show ?thesis .
+next
+ assume *: "\<not> (A \<inter> space M \<in> sets M \<and> emeasure M (A \<inter> space M) < \<infinity>)"
+ have "integral\<^sup>L M (indicator A) = integral\<^sup>L M (indicator (A \<inter> space M) :: _ \<Rightarrow> real)"
+ by (intro integral_cong) (auto split: split_indicator)
+ also have "\<dots> = 0"
+ using * by (subst not_integrable_integral_eq) (auto simp: integrable_indicator_iff)
+ also have "\<dots> = measure M (A \<inter> space M)"
+ using * by (auto simp: measure_def emeasure_notin_sets not_less top_unique)
+ finally show ?thesis .
+qed
+
+lemma integrable_discrete_difference:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes X: "countable X"
+ assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
+ assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+ assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+ shows "integrable M f \<longleftrightarrow> integrable M g"
+ unfolding integrable_iff_bounded
+proof (rule conj_cong)
+ { assume "f \<in> borel_measurable M" then have "g \<in> borel_measurable M"
+ by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) }
+ moreover
+ { assume "g \<in> borel_measurable M" then have "f \<in> borel_measurable M"
+ by (rule measurable_discrete_difference[where X=X]) (auto simp: assms) }
+ ultimately show "f \<in> borel_measurable M \<longleftrightarrow> g \<in> borel_measurable M" ..
+next
+ have "AE x in M. x \<notin> X"
+ by (rule AE_discrete_difference) fact+
+ then have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. norm (g x) \<partial>M)"
+ by (intro nn_integral_cong_AE) (auto simp: eq)
+ then show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity> \<longleftrightarrow> (\<integral>\<^sup>+ x. norm (g x) \<partial>M) < \<infinity>"
+ by simp
+qed
+
+lemma integral_discrete_difference:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes X: "countable X"
+ assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
+ assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+ assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+ shows "integral\<^sup>L M f = integral\<^sup>L M g"
+proof (rule integral_eq_cases)
+ show eq: "integrable M f \<longleftrightarrow> integrable M g"
+ by (rule integrable_discrete_difference[where X=X]) fact+
+
+ assume f: "integrable M f"
+ show "integral\<^sup>L M f = integral\<^sup>L M g"
+ proof (rule integral_cong_AE)
+ show "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ using f eq by (auto intro: borel_measurable_integrable)
+
+ have "AE x in M. x \<notin> X"
+ by (rule AE_discrete_difference) fact+
+ with AE_space show "AE x in M. f x = g x"
+ by eventually_elim fact
+ qed
+qed
+
+lemma has_bochner_integral_discrete_difference:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes X: "countable X"
+ assumes null: "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0"
+ assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+ assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+ shows "has_bochner_integral M f x \<longleftrightarrow> has_bochner_integral M g x"
+ using integrable_discrete_difference[of X M f g, OF assms]
+ using integral_discrete_difference[of X M f g, OF assms]
+ by (metis has_bochner_integral_iff)
+
+lemma
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real"
+ assumes "f \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M" "integrable M w"
+ assumes lim: "AE x in M. (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
+ assumes bound: "\<And>i. AE x in M. norm (s i x) \<le> w x"
+ shows integrable_dominated_convergence: "integrable M f"
+ and integrable_dominated_convergence2: "\<And>i. integrable M (s i)"
+ and integral_dominated_convergence: "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
+proof -
+ have w_nonneg: "AE x in M. 0 \<le> w x"
+ using bound[of 0] by eventually_elim (auto intro: norm_ge_zero order_trans)
+ then have "(\<integral>\<^sup>+x. w x \<partial>M) = (\<integral>\<^sup>+x. norm (w x) \<partial>M)"
+ by (intro nn_integral_cong_AE) auto
+ with \<open>integrable M w\<close> have w: "w \<in> borel_measurable M" "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
+ unfolding integrable_iff_bounded by auto
+
+ show int_s: "\<And>i. integrable M (s i)"
+ unfolding integrable_iff_bounded
+ proof
+ fix i
+ have "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
+ using bound[of i] w_nonneg by (intro nn_integral_mono_AE) auto
+ with w show "(\<integral>\<^sup>+ x. ennreal (norm (s i x)) \<partial>M) < \<infinity>" by auto
+ qed fact
+
+ have all_bound: "AE x in M. \<forall>i. norm (s i x) \<le> w x"
+ using bound unfolding AE_all_countable by auto
+
+ show int_f: "integrable M f"
+ unfolding integrable_iff_bounded
+ proof
+ have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) \<le> (\<integral>\<^sup>+x. w x \<partial>M)"
+ using all_bound lim w_nonneg
+ proof (intro nn_integral_mono_AE, eventually_elim)
+ fix x assume "\<forall>i. norm (s i x) \<le> w x" "(\<lambda>i. s i x) \<longlonglongrightarrow> f x" "0 \<le> w x"
+ then show "ennreal (norm (f x)) \<le> ennreal (w x)"
+ by (intro LIMSEQ_le_const2[where X="\<lambda>i. ennreal (norm (s i x))"]) (auto intro: tendsto_intros)
+ qed
+ with w show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" by auto
+ qed fact
+
+ have "(\<lambda>n. ennreal (norm (integral\<^sup>L M (s n) - integral\<^sup>L M f))) \<longlonglongrightarrow> ennreal 0" (is "?d \<longlonglongrightarrow> ennreal 0")
+ proof (rule tendsto_sandwich)
+ show "eventually (\<lambda>n. ennreal 0 \<le> ?d n) sequentially" "(\<lambda>_. ennreal 0) \<longlonglongrightarrow> ennreal 0" by auto
+ show "eventually (\<lambda>n. ?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)) sequentially"
+ proof (intro always_eventually allI)
+ fix n
+ have "?d n = norm (integral\<^sup>L M (\<lambda>x. s n x - f x))"
+ using int_f int_s by simp
+ also have "\<dots> \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)"
+ by (intro int_f int_s integrable_diff integral_norm_bound_ennreal)
+ finally show "?d n \<le> (\<integral>\<^sup>+x. norm (s n x - f x) \<partial>M)" .
+ qed
+ show "(\<lambda>n. \<integral>\<^sup>+x. norm (s n x - f x) \<partial>M) \<longlonglongrightarrow> ennreal 0"
+ unfolding ennreal_0
+ apply (subst norm_minus_commute)
+ proof (rule nn_integral_dominated_convergence_norm[where w=w])
+ show "\<And>n. s n \<in> borel_measurable M"
+ using int_s unfolding integrable_iff_bounded by auto
+ qed fact+
+ qed
+ then have "(\<lambda>n. integral\<^sup>L M (s n) - integral\<^sup>L M f) \<longlonglongrightarrow> 0"
+ by (simp add: tendsto_norm_zero_iff del: ennreal_0)
+ from tendsto_add[OF this tendsto_const[of "integral\<^sup>L M f"]]
+ show "(\<lambda>i. integral\<^sup>L M (s i)) \<longlonglongrightarrow> integral\<^sup>L M f" by simp
+qed
+
+context
+ fixes s :: "real \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}" and w :: "'a \<Rightarrow> real"
+ and f :: "'a \<Rightarrow> 'b" and M
+ assumes "f \<in> borel_measurable M" "\<And>t. s t \<in> borel_measurable M" "integrable M w"
+ assumes lim: "AE x in M. ((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top"
+ assumes bound: "\<forall>\<^sub>F i in at_top. AE x in M. norm (s i x) \<le> w x"
+begin
+
+lemma integral_dominated_convergence_at_top: "((\<lambda>t. integral\<^sup>L M (s t)) \<longlongrightarrow> integral\<^sup>L M f) at_top"
+proof (rule tendsto_at_topI_sequentially)
+ fix X :: "nat \<Rightarrow> real" assume X: "filterlim X at_top sequentially"
+ from filterlim_iff[THEN iffD1, OF this, rule_format, OF bound]
+ obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s (X n) x) \<le> w x"
+ by (auto simp: eventually_sequentially)
+
+ show "(\<lambda>n. integral\<^sup>L M (s (X n))) \<longlonglongrightarrow> integral\<^sup>L M f"
+ proof (rule LIMSEQ_offset, rule integral_dominated_convergence)
+ show "AE x in M. norm (s (X (n + N)) x) \<le> w x" for n
+ by (rule w) auto
+ show "AE x in M. (\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x"
+ using lim
+ proof eventually_elim
+ fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top"
+ then show "(\<lambda>n. s (X (n + N)) x) \<longlonglongrightarrow> f x"
+ by (intro LIMSEQ_ignore_initial_segment filterlim_compose[OF _ X])
+ qed
+ qed fact+
+qed
+
+lemma integrable_dominated_convergence_at_top: "integrable M f"
+proof -
+ from bound obtain N where w: "\<And>n. N \<le> n \<Longrightarrow> AE x in M. norm (s n x) \<le> w x"
+ by (auto simp: eventually_at_top_linorder)
+ show ?thesis
+ proof (rule integrable_dominated_convergence)
+ show "AE x in M. norm (s (N + i) x) \<le> w x" for i :: nat
+ by (intro w) auto
+ show "AE x in M. (\<lambda>i. s (N + real i) x) \<longlonglongrightarrow> f x"
+ using lim
+ proof eventually_elim
+ fix x assume "((\<lambda>i. s i x) \<longlongrightarrow> f x) at_top"
+ then show "(\<lambda>n. s (N + n) x) \<longlonglongrightarrow> f x"
+ by (rule filterlim_compose)
+ (auto intro!: filterlim_tendsto_add_at_top filterlim_real_sequentially)
+ qed
+ qed fact+
+qed
+
+end
+
+lemma integrable_mult_left_iff:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "integrable M (\<lambda>x. c * f x) \<longleftrightarrow> c = 0 \<or> integrable M f"
+ using integrable_mult_left[of c M f] integrable_mult_left[of "1 / c" M "\<lambda>x. c * f x"]
+ by (cases "c = 0") auto
+
+lemma integrableI_nn_integral_finite:
+ assumes [measurable]: "f \<in> borel_measurable M"
+ and nonneg: "AE x in M. 0 \<le> f x"
+ and finite: "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
+ shows "integrable M f"
+proof (rule integrableI_bounded)
+ have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
+ using nonneg by (intro nn_integral_cong_AE) auto
+ with finite show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+ by auto
+qed simp
+
+lemma integral_nonneg_AE:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes nonneg: "AE x in M. 0 \<le> f x"
+ shows "0 \<le> integral\<^sup>L M f"
+proof cases
+ assume f: "integrable M f"
+ then have [measurable]: "f \<in> M \<rightarrow>\<^sub>M borel"
+ by auto
+ have "(\<lambda>x. max 0 (f x)) \<in> M \<rightarrow>\<^sub>M borel" "\<And>x. 0 \<le> max 0 (f x)" "integrable M (\<lambda>x. max 0 (f x))"
+ using f by auto
+ from this have "0 \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))"
+ proof (induction rule: borel_measurable_induct_real)
+ case (add f g)
+ then have "integrable M f" "integrable M g"
+ by (auto intro!: integrable_bound[OF add.prems])
+ with add show ?case
+ by (simp add: nn_integral_add)
+ next
+ case (seq U)
+ show ?case
+ proof (rule LIMSEQ_le_const)
+ have U_le: "x \<in> space M \<Longrightarrow> U i x \<le> max 0 (f x)" for x i
+ using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def)
+ with seq nonneg show "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> LINT x|M. max 0 (f x)"
+ by (intro integral_dominated_convergence) auto
+ have "integrable M (U i)" for i
+ using seq.prems by (rule integrable_bound) (insert U_le seq, auto)
+ with seq show "\<exists>N. \<forall>n\<ge>N. 0 \<le> integral\<^sup>L M (U n)"
+ by auto
+ qed
+ qed (auto simp: measure_nonneg integrable_mult_left_iff)
+ also have "\<dots> = integral\<^sup>L M f"
+ using nonneg by (auto intro!: integral_cong_AE)
+ finally show ?thesis .
+qed (simp add: not_integrable_integral_eq)
+
+lemma integral_nonneg[simp]:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "(\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x) \<Longrightarrow> 0 \<le> integral\<^sup>L M f"
+ by (intro integral_nonneg_AE) auto
+
+lemma nn_integral_eq_integral:
+ assumes f: "integrable M f"
+ assumes nonneg: "AE x in M. 0 \<le> f x"
+ shows "(\<integral>\<^sup>+ x. f x \<partial>M) = integral\<^sup>L M f"
+proof -
+ { fix f :: "'a \<Rightarrow> real" assume f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "integrable M f"
+ then have "(\<integral>\<^sup>+ x. f x \<partial>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 ennreal_indicator emeasure_eq_ennreal_measure)
+ next
+ case (mult f c) then show ?case
+ by (auto simp add: integrable_mult_left_iff nn_integral_cmult ennreal_mult integral_nonneg_AE)
+ next
+ case (add g f)
+ then have "integrable M f" "integrable M g"
+ by (auto intro!: integrable_bound[OF add.prems])
+ with add show ?case
+ by (simp add: nn_integral_add integral_nonneg_AE)
+ next
+ case (seq U)
+ show ?case
+ proof (rule LIMSEQ_unique)
+ have U_le_f: "x \<in> space M \<Longrightarrow> U i x \<le> f x" for x i
+ using seq by (intro incseq_le) (auto simp: incseq_def le_fun_def)
+ have int_U: "\<And>i. integrable M (U i)"
+ using seq f U_le_f by (intro integrable_bound[OF f(3)]) auto
+ from U_le_f seq have "(\<lambda>i. integral\<^sup>L M (U i)) \<longlonglongrightarrow> integral\<^sup>L M f"
+ by (intro integral_dominated_convergence) auto
+ then show "(\<lambda>i. ennreal (integral\<^sup>L M (U i))) \<longlonglongrightarrow> ennreal (integral\<^sup>L M f)"
+ using seq f int_U by (simp add: f integral_nonneg_AE)
+ have "(\<lambda>i. \<integral>\<^sup>+ x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+ x. f x \<partial>M"
+ using seq U_le_f f
+ by (intro nn_integral_dominated_convergence[where w=f]) (auto simp: integrable_iff_bounded)
+ then show "(\<lambda>i. \<integral>x. U i x \<partial>M) \<longlonglongrightarrow> \<integral>\<^sup>+x. f x \<partial>M"
+ using seq int_U by simp
+ qed
+ qed }
+ from this[of "\<lambda>x. max 0 (f x)"] assms have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = integral\<^sup>L M (\<lambda>x. max 0 (f x))"
+ by simp
+ also have "\<dots> = integral\<^sup>L M f"
+ using assms by (auto intro!: integral_cong_AE simp: integral_nonneg_AE)
+ also have "(\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
+ using assms by (auto intro!: nn_integral_cong_AE simp: max_def)
+ finally show ?thesis .
+qed
+
+lemma
+ fixes f :: "_ \<Rightarrow> _ \<Rightarrow> 'a :: {banach, second_countable_topology}"
+ assumes integrable[measurable]: "\<And>i. integrable M (f i)"
+ and summable: "AE x in M. summable (\<lambda>i. norm (f i x))"
+ and sums: "summable (\<lambda>i. (\<integral>x. norm (f i x) \<partial>M))"
+ shows integrable_suminf: "integrable M (\<lambda>x. (\<Sum>i. f i x))" (is "integrable M ?S")
+ and sums_integral: "(\<lambda>i. integral\<^sup>L M (f i)) sums (\<integral>x. (\<Sum>i. f i x) \<partial>M)" (is "?f sums ?x")
+ and integral_suminf: "(\<integral>x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>L M (f i))"
+ and summable_integral: "summable (\<lambda>i. integral\<^sup>L M (f i))"
+proof -
+ have 1: "integrable M (\<lambda>x. \<Sum>i. norm (f i x))"
+ proof (rule integrableI_bounded)
+ have "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i. ennreal (norm (f i x))) \<partial>M)"
+ apply (intro nn_integral_cong_AE)
+ using summable
+ apply eventually_elim
+ apply (simp add: suminf_nonneg ennreal_suminf_neq_top)
+ done
+ also have "\<dots> = (\<Sum>i. \<integral>\<^sup>+ x. norm (f i x) \<partial>M)"
+ by (intro nn_integral_suminf) auto
+ also have "\<dots> = (\<Sum>i. ennreal (\<integral>x. norm (f i x) \<partial>M))"
+ by (intro arg_cong[where f=suminf] ext nn_integral_eq_integral integrable_norm integrable) auto
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Sum>i. norm (f i x))) \<partial>M) < \<infinity>"
+ by (simp add: sums ennreal_suminf_neq_top less_top[symmetric] integral_nonneg_AE)
+ qed simp
+
+ have 2: "AE x in M. (\<lambda>n. \<Sum>i<n. f i x) \<longlonglongrightarrow> (\<Sum>i. f i x)"
+ using summable by eventually_elim (auto intro: summable_LIMSEQ summable_norm_cancel)
+
+ have 3: "\<And>j. AE x in M. norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))"
+ using summable
+ proof eventually_elim
+ fix j x assume [simp]: "summable (\<lambda>i. norm (f i x))"
+ have "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i<j. norm (f i x))" by (rule norm_setsum)
+ also have "\<dots> \<le> (\<Sum>i. norm (f i x))"
+ using setsum_le_suminf[of "\<lambda>i. norm (f i x)"] unfolding sums_iff by auto
+ finally show "norm (\<Sum>i<j. f i x) \<le> (\<Sum>i. norm (f i x))" by simp
+ qed
+
+ note ibl = integrable_dominated_convergence[OF _ _ 1 2 3]
+ note int = integral_dominated_convergence[OF _ _ 1 2 3]
+
+ show "integrable M ?S"
+ by (rule ibl) measurable
+
+ show "?f sums ?x" unfolding sums_def
+ using int by (simp add: integrable)
+ then show "?x = suminf ?f" "summable ?f"
+ unfolding sums_iff by auto
+qed
+
+lemma integral_norm_bound:
+ fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
+ shows "integrable M f \<Longrightarrow> norm (integral\<^sup>L M f) \<le> (\<integral>x. norm (f x) \<partial>M)"
+ using nn_integral_eq_integral[of M "\<lambda>x. norm (f x)"]
+ using integral_norm_bound_ennreal[of M f] by (simp add: integral_nonneg_AE)
+
+lemma integral_eq_nn_integral:
+ assumes [measurable]: "f \<in> borel_measurable M"
+ assumes nonneg: "AE x in M. 0 \<le> f x"
+ shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M)"
+proof cases
+ assume *: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = \<infinity>"
+ also have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
+ using nonneg by (intro nn_integral_cong_AE) auto
+ finally have "\<not> integrable M f"
+ by (auto simp: integrable_iff_bounded)
+ then show ?thesis
+ by (simp add: * not_integrable_integral_eq)
+next
+ assume "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
+ then have "integrable M f"
+ by (cases "\<integral>\<^sup>+ x. ennreal (f x) \<partial>M" rule: ennreal_cases)
+ (auto intro!: integrableI_nn_integral_finite assms)
+ from nn_integral_eq_integral[OF this] nonneg show ?thesis
+ by (simp add: integral_nonneg_AE)
+qed
+
+lemma enn2real_nn_integral_eq_integral:
+ assumes eq: "AE x in M. f x = ennreal (g x)" and nn: "AE x in M. 0 \<le> g x"
+ and fin: "(\<integral>\<^sup>+x. f x \<partial>M) < top"
+ and [measurable]: "g \<in> M \<rightarrow>\<^sub>M borel"
+ shows "enn2real (\<integral>\<^sup>+x. f x \<partial>M) = (\<integral>x. g x \<partial>M)"
+proof -
+ have "ennreal (enn2real (\<integral>\<^sup>+x. f x \<partial>M)) = (\<integral>\<^sup>+x. f x \<partial>M)"
+ using fin by (intro ennreal_enn2real) auto
+ also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M)"
+ using eq by (rule nn_integral_cong_AE)
+ also have "\<dots> = (\<integral>x. g x \<partial>M)"
+ proof (rule nn_integral_eq_integral)
+ show "integrable M g"
+ proof (rule integrableI_bounded)
+ have "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M)"
+ using eq nn by (auto intro!: nn_integral_cong_AE elim!: eventually_elim2)
+ also note fin
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (g x)) \<partial>M) < \<infinity>"
+ by simp
+ qed simp
+ qed fact
+ finally show ?thesis
+ using nn by (simp add: integral_nonneg_AE)
+qed
+
+lemma has_bochner_integral_nn_integral:
+ assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> x"
+ assumes "(\<integral>\<^sup>+x. f x \<partial>M) = ennreal x"
+ shows "has_bochner_integral M f x"
+ unfolding has_bochner_integral_iff
+ using assms by (auto simp: assms integral_eq_nn_integral intro: integrableI_nn_integral_finite)
+
+lemma integrableI_simple_bochner_integrable:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "simple_bochner_integrable M f \<Longrightarrow> integrable M f"
+ by (intro integrableI_sequence[where s="\<lambda>_. f"] borel_measurable_simple_function)
+ (auto simp: zero_ennreal_def[symmetric] simple_bochner_integrable.simps)
+
+lemma integrable_induct[consumes 1, case_names base add lim, induct pred: integrable]:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes "integrable M f"
+ assumes base: "\<And>A c. A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> P (\<lambda>x. indicator A x *\<^sub>R c)"
+ assumes add: "\<And>f g. integrable M f \<Longrightarrow> P f \<Longrightarrow> integrable M g \<Longrightarrow> P g \<Longrightarrow> P (\<lambda>x. f x + g x)"
+ assumes lim: "\<And>f s. (\<And>i. integrable M (s i)) \<Longrightarrow> (\<And>i. P (s i)) \<Longrightarrow>
+ (\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x) \<Longrightarrow>
+ (\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)) \<Longrightarrow> integrable M f \<Longrightarrow> P f"
+ shows "P f"
+proof -
+ from \<open>integrable M f\<close> have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>"
+ unfolding integrable_iff_bounded by auto
+ from borel_measurable_implies_sequence_metric[OF f(1)]
+ obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x"
+ "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)"
+ unfolding norm_conv_dist by metis
+
+ { fix f A
+ have [simp]: "P (\<lambda>x. 0)"
+ using base[of "{}" undefined] by simp
+ have "(\<And>i::'b. i \<in> A \<Longrightarrow> integrable M (f i::'a \<Rightarrow> 'b)) \<Longrightarrow>
+ (\<And>i. i \<in> A \<Longrightarrow> P (f i)) \<Longrightarrow> P (\<lambda>x. \<Sum>i\<in>A. f i x)"
+ by (induct A rule: infinite_finite_induct) (auto intro!: add) }
+ note setsum = this
+
+ define s' where [abs_def]: "s' i z = indicator (space M) z *\<^sub>R s i z" for i z
+ then have s'_eq_s: "\<And>i x. x \<in> space M \<Longrightarrow> s' i x = s i x"
+ by simp
+
+ have sf[measurable]: "\<And>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 "\<And>z. {y. s' i z = y \<and> y \<in> s' i ` space M \<and> y \<noteq> 0 \<and> z \<in> space M} =
+ (if z \<in> space M \<and> s' i z \<noteq> 0 then {s' i z} else {})"
+ by (auto simp add: s'_def split: split_indicator)
+ then have "\<And>z. s' i = (\<lambda>z. \<Sum>y\<in>s' i`space M - {0}. indicator {x\<in>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 "(\<integral>\<^sup>+x. norm (s' i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)"
+ using s by (intro nn_integral_mono) (auto simp: s'_eq_s)
+ also have "\<dots> < \<infinity>"
+ using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult)
+ 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 \<in> space M" "s' i x \<noteq> 0"
+ then have "emeasure M {y \<in> space M. s' i y = s' i x} \<le> emeasure M {y \<in> space M. s' i y \<noteq> 0}"
+ by (intro emeasure_mono) auto
+ also have "\<dots> < \<infinity>"
+ using sbi by (auto elim: simple_bochner_integrable.cases simp: less_top)
+ finally have "emeasure M {y \<in> space M. s' i y = s' i x} \<noteq> \<infinity>" by simp }
+ then show "P (s' i)"
+ by (subst s'_eq) (auto intro!: setsum base simp: less_top)
+
+ fix x assume "x \<in> space M" with s show "(\<lambda>i. s' i x) \<longlonglongrightarrow> f x"
+ by (simp add: s'_eq_s)
+ show "norm (s' i x) \<le> 2 * norm (f x)"
+ using \<open>x \<in> space M\<close> s by (simp add: s'_eq_s)
+ qed fact
+qed
+
+lemma integral_eq_zero_AE:
+ "(AE x in M. f x = 0) \<Longrightarrow> integral\<^sup>L M f = 0"
+ using integral_cong_AE[of f M "\<lambda>_. 0"]
+ by (cases "integrable M f") (simp_all add: not_integrable_integral_eq)
+
+lemma integral_nonneg_eq_0_iff_AE:
+ fixes f :: "_ \<Rightarrow> real"
+ assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \<le> f x"
+ shows "integral\<^sup>L M f = 0 \<longleftrightarrow> (AE x in M. f x = 0)"
+proof
+ assume "integral\<^sup>L M f = 0"
+ then have "integral\<^sup>N M f = 0"
+ using nn_integral_eq_integral[OF f nonneg] by simp
+ then have "AE x in M. ennreal (f x) \<le> 0"
+ by (simp add: nn_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 \<Rightarrow> real"
+ assumes "integrable M f" "integrable M g" "AE x in M. f x \<le> g x"
+ shows "integral\<^sup>L M f \<le> integral\<^sup>L M g"
+proof -
+ have "0 \<le> integral\<^sup>L M (\<lambda>x. g x - f x)"
+ using assms by (intro integral_nonneg_AE integrable_diff assms) auto
+ also have "\<dots> = 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 \<Rightarrow> real"
+ shows "integrable M f \<Longrightarrow> integrable M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x) \<Longrightarrow>
+ integral\<^sup>L M f \<le> integral\<^sup>L M g"
+ by (intro integral_mono_AE) auto
+
+lemma (in finite_measure) integrable_measure:
+ assumes I: "disjoint_family_on X I" "countable I"
+ shows "integrable (count_space I) (\<lambda>i. measure M (X i))"
+proof -
+ have "(\<integral>\<^sup>+i. measure M (X i) \<partial>count_space I) = (\<integral>\<^sup>+i. measure M (if X i \<in> sets M then X i else {}) \<partial>count_space I)"
+ by (auto intro!: nn_integral_cong measure_notin_sets)
+ also have "\<dots> = measure M (\<Union>i\<in>I. if X i \<in> sets M then X i else {})"
+ using I unfolding emeasure_eq_measure[symmetric]
+ by (subst emeasure_UN_countable) (auto simp: disjoint_family_on_def)
+ finally show ?thesis
+ by (auto intro!: integrableI_bounded)
+qed
+
+lemma integrableI_real_bounded:
+ assumes f: "f \<in> borel_measurable M" and ae: "AE x in M. 0 \<le> f x" and fin: "integral\<^sup>N M f < \<infinity>"
+ shows "integrable M f"
+proof (rule integrableI_bounded)
+ have "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) = \<integral>\<^sup>+ x. ennreal (f x) \<partial>M"
+ using ae by (auto intro: nn_integral_cong_AE)
+ also note fin
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>" .
+qed fact
+
+lemma integral_real_bounded:
+ assumes "0 \<le> r" "integral\<^sup>N M f \<le> ennreal r"
+ shows "integral\<^sup>L M f \<le> r"
+proof cases
+ assume [simp]: "integrable M f"
+
+ have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (f x))"
+ by (intro nn_integral_eq_integral[symmetric]) auto
+ also have "\<dots> = integral\<^sup>N M f"
+ by (intro nn_integral_cong) (simp add: max_def ennreal_neg)
+ also have "\<dots> \<le> r"
+ by fact
+ finally have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) \<le> r"
+ using \<open>0 \<le> r\<close> by simp
+
+ moreover have "integral\<^sup>L M f \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))"
+ by (rule integral_mono_AE) auto
+ ultimately show ?thesis
+ by simp
+next
+ assume "\<not> integrable M f" then show ?thesis
+ using \<open>0 \<le> r\<close> by (simp add: not_integrable_integral_eq)
+qed
+
+subsection \<open>Restricted measure spaces\<close>
+
+lemma integrable_restrict_space:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
+ shows "integrable (restrict_space M \<Omega>) f \<longleftrightarrow> integrable M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
+ unfolding integrable_iff_bounded
+ borel_measurable_restrict_space_iff[OF \<Omega>]
+ nn_integral_restrict_space[OF \<Omega>]
+ by (simp add: ac_simps ennreal_indicator ennreal_mult)
+
+lemma integral_restrict_space:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
+ shows "integral\<^sup>L (restrict_space M \<Omega>) f = integral\<^sup>L M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
+proof (rule integral_eq_cases)
+ assume "integrable (restrict_space M \<Omega>) f"
+ then show ?thesis
+ proof induct
+ case (base A c) then show ?case
+ by (simp add: indicator_inter_arith[symmetric] sets_restrict_space_iff
+ emeasure_restrict_space Int_absorb1 measure_restrict_space)
+ next
+ case (add g f) then show ?case
+ by (simp add: scaleR_add_right integrable_restrict_space)
+ next
+ case (lim f s)
+ show ?case
+ proof (rule LIMSEQ_unique)
+ show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> integral\<^sup>L (restrict_space M \<Omega>) f"
+ using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) simp_all
+
+ show "(\<lambda>i. integral\<^sup>L (restrict_space M \<Omega>) (s i)) \<longlonglongrightarrow> (\<integral> x. indicator \<Omega> x *\<^sub>R f x \<partial>M)"
+ unfolding lim
+ using lim
+ by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (indicator \<Omega> x *\<^sub>R f x)"])
+ (auto simp add: space_restrict_space integrable_restrict_space simp del: norm_scaleR
+ split: split_indicator)
+ qed
+ qed
+qed (simp add: integrable_restrict_space)
+
+lemma integral_empty:
+ assumes "space M = {}"
+ shows "integral\<^sup>L M f = 0"
+proof -
+ have "(\<integral> x. f x \<partial>M) = (\<integral> x. 0 \<partial>M)"
+ by(rule integral_cong)(simp_all add: assms)
+ thus ?thesis by simp
+qed
+
+subsection \<open>Measure spaces with an associated density\<close>
+
+lemma integrable_density:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
+ assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ and nn: "AE x in M. 0 \<le> g x"
+ shows "integrable (density M g) f \<longleftrightarrow> integrable M (\<lambda>x. g x *\<^sub>R f x)"
+ unfolding integrable_iff_bounded using nn
+ apply (simp add: nn_integral_density less_top[symmetric])
+ apply (intro arg_cong2[where f="op ="] refl nn_integral_cong_AE)
+ apply (auto simp: ennreal_mult)
+ done
+
+lemma integral_density:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
+ assumes f: "f \<in> borel_measurable M"
+ and g[measurable]: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
+ shows "integral\<^sup>L (density M g) f = integral\<^sup>L M (\<lambda>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 \<in> sets M" by auto
+
+ have int: "integrable M (\<lambda>x. g x * indicator A x)"
+ using g base integrable_density[of "indicator A :: 'a \<Rightarrow> real" M g] by simp
+ then have "integral\<^sup>L M (\<lambda>x. g x * indicator A x) = (\<integral>\<^sup>+ x. ennreal (g x * indicator A x) \<partial>M)"
+ using g by (subst nn_integral_eq_integral) auto
+ also have "\<dots> = (\<integral>\<^sup>+ x. ennreal (g x) * indicator A x \<partial>M)"
+ by (intro nn_integral_cong) (auto split: split_indicator)
+ also have "\<dots> = emeasure (density M g) A"
+ by (rule emeasure_density[symmetric]) auto
+ also have "\<dots> = ennreal (measure (density M g) A)"
+ using base by (auto intro: emeasure_eq_ennreal_measure)
+ also have "\<dots> = integral\<^sup>L (density M g) (indicator A)"
+ using base by simp
+ finally show ?case
+ using base g
+ apply (simp add: int integral_nonneg_AE)
+ apply (subst (asm) ennreal_inj)
+ apply (auto intro!: integral_nonneg_AE)
+ done
+ next
+ case (add f h)
+ then have [measurable]: "f \<in> borel_measurable M" "h \<in> 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 \<in> borel_measurable M" "\<And>i. s i \<in> borel_measurable M"
+ using lim(1,5)[THEN borel_measurable_integrable] by auto
+
+ show ?case
+ proof (rule LIMSEQ_unique)
+ show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. g x *\<^sub>R f x)"
+ proof (rule integral_dominated_convergence)
+ show "integrable M (\<lambda>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. (\<lambda>i. g x *\<^sub>R s i x) \<longlonglongrightarrow> g x *\<^sub>R f x"
+ using lim(3) by (auto intro!: tendsto_scaleR AE_I2[of M])
+ show "\<And>i. AE x in M. norm (g x *\<^sub>R s i x) \<le> 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 "(\<lambda>i. integral\<^sup>L M (\<lambda>x. g x *\<^sub>R s i x)) \<longlonglongrightarrow> integral\<^sup>L (density M g) f"
+ unfolding lim(2)[symmetric]
+ by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
+ (insert lim(3-5), auto)
+ qed
+ qed
+qed (simp add: f g integrable_density)
+
+lemma
+ fixes g :: "'a \<Rightarrow> real"
+ assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "g \<in> borel_measurable M"
+ shows integral_real_density: "integral\<^sup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)"
+ and integrable_real_density: "integrable (density M f) g \<longleftrightarrow> integrable M (\<lambda>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 \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
+ shows "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (AE x in M. 0 \<le> g x) \<Longrightarrow>
+ has_bochner_integral M (\<lambda>x. g x *\<^sub>R f x) x \<Longrightarrow> has_bochner_integral (density M g) f x"
+ by (simp add: has_bochner_integral_iff integrable_density integral_density)
+
+subsection \<open>Distributions\<close>
+
+lemma integrable_distr_eq:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes [measurable]: "g \<in> measurable M N" "f \<in> borel_measurable N"
+ shows "integrable (distr M N g) f \<longleftrightarrow> integrable M (\<lambda>x. f (g x))"
+ unfolding integrable_iff_bounded by (simp_all add: nn_integral_distr)
+
+lemma integrable_distr:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "T \<in> measurable M M' \<Longrightarrow> integrable (distr M M' T) f \<Longrightarrow> integrable M (\<lambda>x. f (T x))"
+ by (subst integrable_distr_eq[symmetric, where g=T])
+ (auto dest: borel_measurable_integrable)
+
+lemma integral_distr:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes g[measurable]: "g \<in> measurable M N" and f: "f \<in> borel_measurable N"
+ shows "integral\<^sup>L (distr M N g) f = integral\<^sup>L M (\<lambda>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 \<in> sets N" by auto
+ from base have int: "integrable (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c)"
+ by (intro integrable_indicator)
+
+ have "integral\<^sup>L (distr M N g) (\<lambda>a. indicator A a *\<^sub>R c) = measure (distr M N g) A *\<^sub>R c"
+ using base by auto
+ also have "\<dots> = measure M (g -` A \<inter> space M) *\<^sub>R c"
+ by (subst measure_distr) auto
+ also have "\<dots> = integral\<^sup>L M (\<lambda>a. indicator (g -` A \<inter> space M) a *\<^sub>R c)"
+ using base by (auto simp: emeasure_distr)
+ also have "\<dots> = integral\<^sup>L M (\<lambda>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 \<in> borel_measurable N" "h \<in> 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 \<in> borel_measurable N" "\<And>i. s i \<in> borel_measurable N"
+ using lim(1,5)[THEN borel_measurable_integrable] by auto
+
+ show ?case
+ proof (rule LIMSEQ_unique)
+ show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. f (g x))"
+ proof (rule integral_dominated_convergence)
+ show "integrable M (\<lambda>x. 2 * norm (f (g x)))"
+ using lim by (auto simp: integrable_distr_eq)
+ show "AE x in M. (\<lambda>i. s i (g x)) \<longlonglongrightarrow> f (g x)"
+ using lim(3) g[THEN measurable_space] by auto
+ show "\<And>i. AE x in M. norm (s i (g x)) \<le> 2 * norm (f (g x))"
+ using lim(4) g[THEN measurable_space] by auto
+ qed auto
+ show "(\<lambda>i. integral\<^sup>L M (\<lambda>x. s i (g x))) \<longlonglongrightarrow> integral\<^sup>L (distr M N g) f"
+ unfolding lim(2)[symmetric]
+ by (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
+ (insert lim(3-5), auto)
+ qed
+ qed
+qed (simp add: f g integrable_distr_eq)
+
+lemma has_bochner_integral_distr:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "f \<in> borel_measurable N \<Longrightarrow> g \<in> measurable M N \<Longrightarrow>
+ has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> has_bochner_integral (distr M N g) f x"
+ by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
+
+subsection \<open>Lebesgue integration on @{const count_space}\<close>
+
+lemma integrable_count_space:
+ fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
+ shows "finite X \<Longrightarrow> integrable (count_space X) f"
+ by (auto simp: nn_integral_count_space integrable_iff_bounded)
+
+lemma measure_count_space[simp]:
+ "B \<subseteq> A \<Longrightarrow> finite B \<Longrightarrow> 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\<in>A. f a \<noteq> 0}"
+ shows "(\<integral>x. f x \<partial>count_space A) = (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. f a)"
+proof -
+ have eq: "\<And>x. x \<in> A \<Longrightarrow> (\<Sum>a | x = a \<and> a \<in> A \<and> f a \<noteq> 0. f a) = (\<Sum>x\<in>{x}. f x)"
+ by (intro setsum.mono_neutral_cong_left) auto
+
+ have "(\<integral>x. f x \<partial>count_space A) = (\<integral>x. (\<Sum>a | a \<in> A \<and> f a \<noteq> 0. indicator {a} x *\<^sub>R f a) \<partial>count_space A)"
+ by (intro integral_cong refl) (simp add: f eq)
+ also have "\<dots> = (\<Sum>a | a \<in> A \<and> f a \<noteq> 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 \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
+ by (subst lebesgue_integral_count_space_finite_support)
+ (auto intro!: setsum.mono_neutral_cong_left)
+
+lemma integrable_count_space_nat_iff:
+ fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
+ shows "integrable (count_space UNIV) f \<longleftrightarrow> summable (\<lambda>x. norm (f x))"
+ by (auto simp add: integrable_iff_bounded nn_integral_count_space_nat ennreal_suminf_neq_top
+ intro: summable_suminf_not_top)
+
+lemma sums_integral_count_space_nat:
+ fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
+ assumes *: "integrable (count_space UNIV) f"
+ shows "f sums (integral\<^sup>L (count_space UNIV) f)"
+proof -
+ let ?f = "\<lambda>n i. indicator {n} i *\<^sub>R f i"
+ have f': "\<And>n i. ?f n i = indicator {n} i *\<^sub>R f n"
+ by (auto simp: fun_eq_iff split: split_indicator)
+
+ have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) sums \<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV"
+ proof (rule sums_integral)
+ show "\<And>i. integrable (count_space UNIV) (?f i)"
+ using * by (intro integrable_mult_indicator) auto
+ show "AE n in count_space UNIV. summable (\<lambda>i. norm (?f i n))"
+ using summable_finite[of "{n}" "\<lambda>i. norm (?f i n)" for n] by simp
+ show "summable (\<lambda>i. \<integral> n. norm (?f i n) \<partial>count_space UNIV)"
+ using * by (subst f') (simp add: integrable_count_space_nat_iff)
+ qed
+ also have "(\<integral> n. (\<Sum>i. ?f i n) \<partial>count_space UNIV) = (\<integral>n. f n \<partial>count_space UNIV)"
+ using suminf_finite[of "{n}" "\<lambda>i. ?f i n" for n] by (auto intro!: integral_cong)
+ also have "(\<lambda>i. \<integral>n. ?f i n \<partial>count_space UNIV) = f"
+ by (subst f') simp
+ finally show ?thesis .
+qed
+
+lemma integral_count_space_nat:
+ fixes f :: "nat \<Rightarrow> _::{banach,second_countable_topology}"
+ shows "integrable (count_space UNIV) f \<Longrightarrow> integral\<^sup>L (count_space UNIV) f = (\<Sum>x. f x)"
+ using sums_integral_count_space_nat by (rule sums_unique)
+
+subsection \<open>Point measure\<close>
+
+lemma lebesgue_integral_point_measure_finite:
+ fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ shows "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow>
+ integral\<^sup>L (point_measure A f) g = (\<Sum>a\<in>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 \<Rightarrow> 'b::{banach, second_countable_topology}" and f :: "'a \<Rightarrow> real"
+ shows "finite A \<Longrightarrow> integrable (point_measure A f) g"
+ unfolding point_measure_def
+ apply (subst density_cong[where f'="\<lambda>x. ennreal (max 0 (f x))"])
+ apply (auto split: split_max simp: ennreal_neg)
+ apply (subst integrable_density)
+ apply (auto simp: AE_count_space integrable_count_space)
+ done
+
+subsection \<open>Lebesgue integration on @{const null_measure}\<close>
+
+lemma has_bochner_integral_null_measure_iff[iff]:
+ "has_bochner_integral (null_measure M) f 0 \<longleftrightarrow> f \<in> borel_measurable M"
+ by (auto simp add: has_bochner_integral.simps simple_bochner_integral_def[abs_def]
+ intro!: exI[of _ "\<lambda>n x. 0"] simple_bochner_integrable.intros)
+
+lemma integrable_null_measure_iff[iff]: "integrable (null_measure M) f \<longleftrightarrow> f \<in> borel_measurable M"
+ by (auto simp add: integrable.simps)
+
+lemma integral_null_measure[simp]: "integral\<^sup>L (null_measure M) f = 0"
+ by (cases "integrable (null_measure M) f")
+ (auto simp add: not_integrable_integral_eq has_bochner_integral_integral_eq)
+
+subsection \<open>Legacy lemmas for the real-valued Lebesgue integral\<close>
+
+lemma real_lebesgue_integral_def:
+ assumes f[measurable]: "integrable M f"
+ shows "integral\<^sup>L M f = enn2real (\<integral>\<^sup>+x. f x \<partial>M) - enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
+proof -
+ have "integral\<^sup>L M f = integral\<^sup>L M (\<lambda>x. max 0 (f x) - max 0 (- f x))"
+ by (auto intro!: arg_cong[where f="integral\<^sup>L M"])
+ also have "\<dots> = integral\<^sup>L M (\<lambda>x. max 0 (f x)) - integral\<^sup>L M (\<lambda>x. max 0 (- f x))"
+ by (intro integral_diff integrable_max integrable_minus integrable_zero f)
+ also have "integral\<^sup>L M (\<lambda>x. max 0 (f x)) = enn2real (\<integral>\<^sup>+x. ennreal (f x) \<partial>M)"
+ by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg)
+ also have "integral\<^sup>L M (\<lambda>x. max 0 (- f x)) = enn2real (\<integral>\<^sup>+x. ennreal (- f x) \<partial>M)"
+ by (subst integral_eq_nn_integral) (auto intro!: arg_cong[where f=enn2real] nn_integral_cong simp: max_def ennreal_neg)
+ finally show ?thesis .
+qed
+
+lemma real_integrable_def:
+ "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
+ (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
+ unfolding integrable_iff_bounded
+proof (safe del: notI)
+ assume *: "(\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M) < \<infinity>"
+ have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
+ by (intro nn_integral_mono) auto
+ also note *
+ finally show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>"
+ by simp
+ have "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) \<partial>M)"
+ by (intro nn_integral_mono) auto
+ also note *
+ finally show "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
+ by simp
+next
+ assume [measurable]: "f \<in> borel_measurable M"
+ assume fin: "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
+ have "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) = (\<integral>\<^sup>+ x. ennreal (f x) + ennreal (- f x) \<partial>M)"
+ by (intro nn_integral_cong) (auto simp: abs_real_def ennreal_neg)
+ also have"\<dots> = (\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) + (\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M)"
+ by (intro nn_integral_add) auto
+ also have "\<dots> < \<infinity>"
+ using fin by (auto simp: less_top)
+ finally show "(\<integral>\<^sup>+ x. norm (f x) \<partial>M) < \<infinity>" .
+qed
+
+lemma integrableD[dest]:
+ assumes "integrable M f"
+ shows "f \<in> borel_measurable M" "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^sup>+ x. ennreal (- f x) \<partial>M) \<noteq> \<infinity>"
+ using assms unfolding real_integrable_def by auto
+
+lemma integrableE:
+ assumes "integrable M f"
+ obtains r q where
+ "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M) = ennreal r"
+ "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M) = ennreal q"
+ "f \<in> borel_measurable M" "integral\<^sup>L M f = r - q"
+ using assms unfolding real_integrable_def real_lebesgue_integral_def[OF assms]
+ by (cases rule: ennreal2_cases[of "(\<integral>\<^sup>+x. ennreal (-f x)\<partial>M)" "(\<integral>\<^sup>+x. ennreal (f x)\<partial>M)"]) auto
+
+lemma integral_monotone_convergence_nonneg:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+ assumes i: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
+ and pos: "\<And>i. AE x in M. 0 \<le> f i x"
+ and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
+ and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x"
+ and u: "u \<in> borel_measurable M"
+ shows "integrable M u"
+ and "integral\<^sup>L M u = x"
+proof -
+ have nn: "AE x in M. \<forall>i. 0 \<le> f i x"
+ using pos unfolding AE_all_countable by auto
+ with lim have u_nn: "AE x in M. 0 \<le> u x"
+ by eventually_elim (auto intro: LIMSEQ_le_const)
+ have [simp]: "0 \<le> x"
+ by (intro LIMSEQ_le_const[OF ilim] allI exI impI integral_nonneg_AE pos)
+ have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = (SUP n. (\<integral>\<^sup>+ x. ennreal (f n x) \<partial>M))"
+ proof (subst nn_integral_monotone_convergence_SUP_AE[symmetric])
+ fix i
+ from mono nn show "AE x in M. ennreal (f i x) \<le> ennreal (f (Suc i) x)"
+ by eventually_elim (auto simp: mono_def)
+ show "(\<lambda>x. ennreal (f i x)) \<in> borel_measurable M"
+ using i by auto
+ next
+ show "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = \<integral>\<^sup>+ x. (SUP i. ennreal (f i x)) \<partial>M"
+ apply (rule nn_integral_cong_AE)
+ using lim mono nn u_nn
+ apply eventually_elim
+ apply (simp add: LIMSEQ_unique[OF _ LIMSEQ_SUP] incseq_def)
+ done
+ qed
+ also have "\<dots> = ennreal x"
+ using mono i nn unfolding nn_integral_eq_integral[OF i pos]
+ by (subst LIMSEQ_unique[OF LIMSEQ_SUP]) (auto simp: mono_def integral_nonneg_AE pos intro!: integral_mono_AE ilim)
+ finally have "(\<integral>\<^sup>+ x. ennreal (u x) \<partial>M) = ennreal x" .
+ moreover have "(\<integral>\<^sup>+ x. ennreal (- u x) \<partial>M) = 0"
+ using u u_nn by (subst nn_integral_0_iff_AE) (auto simp add: ennreal_neg)
+ 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 \<Rightarrow> 'a \<Rightarrow> real"
+ assumes f: "\<And>i. integrable M (f i)" and mono: "AE x in M. mono (\<lambda>n. f n x)"
+ and lim: "AE x in M. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
+ and ilim: "(\<lambda>i. integral\<^sup>L M (f i)) \<longlonglongrightarrow> x"
+ and u: "u \<in> 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: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)"
+ using f by auto
+ have 2: "AE x in M. mono (\<lambda>n. f n x - f 0 x)"
+ using mono by (auto simp: mono_def le_fun_def)
+ have 3: "\<And>n. AE x in M. 0 \<le> f n x - f 0 x"
+ using mono by (auto simp: field_simps mono_def le_fun_def)
+ have 4: "AE x in M. (\<lambda>i. f i x - f 0 x) \<longlonglongrightarrow> u x - f 0 x"
+ using lim by (auto intro!: tendsto_diff)
+ have 5: "(\<lambda>i. (\<integral>x. f i x - f 0 x \<partial>M)) \<longlonglongrightarrow> x - integral\<^sup>L M (f 0)"
+ using f ilim by (auto intro!: tendsto_diff)
+ have 6: "(\<lambda>x. u x - f 0 x) \<in> 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 (\<lambda>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 \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes f[measurable]: "integrable M f"
+ shows "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0"
+proof -
+ have "(\<integral>\<^sup>+x. norm (f x) \<partial>M) = (\<integral>x. norm (f x) \<partial>M)"
+ using f by (intro nn_integral_eq_integral integrable_norm) auto
+ then have "(\<integral>x. norm (f x) \<partial>M) = 0 \<longleftrightarrow> (\<integral>\<^sup>+x. norm (f x) \<partial>M) = 0"
+ by simp
+ also have "\<dots> \<longleftrightarrow> emeasure M {x\<in>space M. ennreal (norm (f x)) \<noteq> 0} = 0"
+ by (intro nn_integral_0_iff) auto
+ finally show ?thesis
+ by simp
+qed
+
+lemma integral_0_iff:
+ fixes f :: "'a \<Rightarrow> real"
+ shows "integrable M f \<Longrightarrow> (\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> emeasure M {x\<in>space M. f x \<noteq> 0} = 0"
+ using integral_norm_eq_0_iff[of M f] by simp
+
+lemma (in finite_measure) integrable_const[intro!, simp]: "integrable M (\<lambda>x. a)"
+ using integrable_indicator[of "space M" M a] by (simp cong: integrable_cong add: less_top[symmetric])
+
+lemma lebesgue_integral_const[simp]:
+ fixes a :: "'a :: {banach, second_countable_topology}"
+ shows "(\<integral>x. a \<partial>M) = measure M (space M) *\<^sub>R a"
+proof -
+ { assume "emeasure M (space M) = \<infinity>" "a \<noteq> 0"
+ then have ?thesis
+ by (auto simp add: not_integrable_integral_eq ennreal_mult_less_top measure_def integrable_iff_bounded) }
+ moreover
+ { assume "a = 0" then have ?thesis by simp }
+ moreover
+ { assume "emeasure M (space M) \<noteq> \<infinity>"
+ interpret finite_measure M
+ proof qed fact
+ have "(\<integral>x. a \<partial>M) = (\<integral>x. indicator (space M) x *\<^sub>R a \<partial>M)"
+ by (intro integral_cong) auto
+ also have "\<dots> = measure M (space M) *\<^sub>R a"
+ by (simp add: less_top[symmetric])
+ finally have ?thesis . }
+ ultimately show ?thesis by blast
+qed
+
+lemma (in finite_measure) integrable_const_bound:
+ fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
+ shows "AE x in M. norm (f x) \<le> B \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> integrable M f"
+ apply (rule integrable_bound[OF integrable_const[of B], of f])
+ apply assumption
+ apply (cases "0 \<le> B")
+ apply auto
+ done
+
+lemma integral_indicator_finite_real:
+ fixes f :: "'a \<Rightarrow> real"
+ assumes [simp]: "finite A"
+ assumes [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
+ assumes finite: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} < \<infinity>"
+ shows "(\<integral>x. f x * indicator A x \<partial>M) = (\<Sum>a\<in>A. f a * measure M {a})"
+proof -
+ have "(\<integral>x. f x * indicator A x \<partial>M) = (\<integral>x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
+ proof (intro integral_cong refl)
+ fix x show "f x * indicator A x = (\<Sum>a\<in>A. f a * indicator {a} x)"
+ by (auto split: split_indicator simp: eq_commute[of x] cong: conj_cong)
+ qed
+ also have "\<dots> = (\<Sum>a\<in>A. f a * measure M {a})"
+ using finite by (subst integral_setsum) (auto simp add: integrable_mult_left_iff)
+ finally show ?thesis .
+qed
+
+lemma (in finite_measure) ennreal_integral_real:
+ assumes [measurable]: "f \<in> borel_measurable M"
+ assumes ae: "AE x in M. f x \<le> ennreal B" "0 \<le> B"
+ shows "ennreal (\<integral>x. enn2real (f x) \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"
+proof (subst nn_integral_eq_integral[symmetric])
+ show "integrable M (\<lambda>x. enn2real (f x))"
+ using ae by (intro integrable_const_bound[where B=B]) (auto simp: enn2real_leI enn2real_nonneg)
+ show "(\<integral>\<^sup>+ x. ennreal (enn2real (f x)) \<partial>M) = integral\<^sup>N M f"
+ using ae by (intro nn_integral_cong_AE) (auto simp: le_less_trans[OF _ ennreal_less_top])
+qed (auto simp: enn2real_nonneg)
+
+lemma (in finite_measure) integral_less_AE:
+ fixes X Y :: "'a \<Rightarrow> real"
+ assumes int: "integrable M X" "integrable M Y"
+ assumes A: "(emeasure M) A \<noteq> 0" "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> X x \<noteq> Y x"
+ assumes gt: "AE x in M. X x \<le> Y x"
+ shows "integral\<^sup>L M X < integral\<^sup>L M Y"
+proof -
+ have "integral\<^sup>L M X \<le> integral\<^sup>L M Y"
+ using gt int by (intro integral_mono_AE) auto
+ moreover
+ have "integral\<^sup>L M X \<noteq> integral\<^sup>L M Y"
+ proof
+ assume eq: "integral\<^sup>L M X = integral\<^sup>L M Y"
+ have "integral\<^sup>L M (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^sup>L M (\<lambda>x. Y x - X x)"
+ using gt int by (intro integral_cong_AE) auto
+ also have "\<dots> = 0"
+ using eq int by simp
+ finally have "(emeasure M) {x \<in> space M. Y x - X x \<noteq> 0} = 0"
+ using int by (simp add: integral_0_iff)
+ moreover
+ have "(\<integral>\<^sup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^sup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
+ using A by (intro nn_integral_mono_AE) auto
+ then have "(emeasure M) A \<le> (emeasure M) {x \<in> space M. Y x - X x \<noteq> 0}"
+ using int A by (simp add: integrable_def)
+ ultimately have "emeasure M A = 0"
+ by simp
+ with \<open>(emeasure M) A \<noteq> 0\<close> show False by auto
+ qed
+ ultimately show ?thesis by auto
+qed
+
+lemma (in finite_measure) integral_less_AE_space:
+ fixes X Y :: "'a \<Rightarrow> real"
+ assumes int: "integrable M X" "integrable M Y"
+ assumes gt: "AE x in M. X x < Y x" "emeasure M (space M) \<noteq> 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 tendsto_integral_at_top:
+ fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
+ assumes [measurable_cong]: "sets M = sets borel" and f[measurable]: "integrable M f"
+ shows "((\<lambda>y. \<integral> x. indicator {.. y} x *\<^sub>R f x \<partial>M) \<longlongrightarrow> \<integral> x. f x \<partial>M) at_top"
+proof (rule tendsto_at_topI_sequentially)
+ fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially"
+ show "(\<lambda>n. \<integral>x. indicator {..X n} x *\<^sub>R f x \<partial>M) \<longlonglongrightarrow> integral\<^sup>L M f"
+ proof (rule integral_dominated_convergence)
+ show "integrable M (\<lambda>x. norm (f x))"
+ by (rule integrable_norm) fact
+ show "AE x in M. (\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x"
+ proof
+ fix x
+ from \<open>filterlim X at_top sequentially\<close>
+ have "eventually (\<lambda>n. x \<le> X n) sequentially"
+ unfolding filterlim_at_top_ge[where c=x] by auto
+ then show "(\<lambda>n. indicator {..X n} x *\<^sub>R f x) \<longlonglongrightarrow> f x"
+ by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono)
+ qed
+ fix n show "AE x in M. norm (indicator {..X n} x *\<^sub>R f x) \<le> norm (f x)"
+ by (auto split: split_indicator)
+ qed auto
+qed
+
+lemma
+ fixes f :: "real \<Rightarrow> real"
+ assumes M: "sets M = sets borel"
+ assumes nonneg: "AE x in M. 0 \<le> f x"
+ assumes borel: "f \<in> borel_measurable borel"
+ assumes int: "\<And>y. integrable M (\<lambda>x. f x * indicator {.. y} x)"
+ assumes conv: "((\<lambda>y. \<integral> x. f x * indicator {.. y} x \<partial>M) \<longlongrightarrow> 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 (\<lambda>n::nat. f x * indicator {..real n} x)"
+ by (auto split: split_indicator intro!: monoI)
+ { fix x have "eventually (\<lambda>n. f x * indicator {..real n} x = f x) sequentially"
+ by (rule eventually_sequentiallyI[of "nat \<lceil>x\<rceil>"])
+ (auto split: split_indicator simp: nat_le_iff ceiling_le_iff) }
+ from filterlim_cong[OF refl refl this]
+ have "AE x in M. (\<lambda>i. f x * indicator {..real i} x) \<longlonglongrightarrow> f x"
+ by simp
+ have "(\<lambda>i. \<integral> x. f x * indicator {..real i} x \<partial>M) \<longlonglongrightarrow> 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 \<in> 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
+
+subsection \<open>Product measure\<close>
+
+lemma (in sigma_finite_measure) borel_measurable_lebesgue_integrable[measurable (raw)]:
+ fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes [measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
+ shows "Measurable.pred N (\<lambda>x. integrable M (f x))"
+proof -
+ have [simp]: "\<And>x. x \<in> space N \<Longrightarrow> integrable M (f x) \<longleftrightarrow> (\<integral>\<^sup>+y. norm (f x y) \<partial>M) < \<infinity>"
+ unfolding integrable_iff_bounded by simp
+ show ?thesis
+ by (simp cong: measurable_cong)
+qed
+
+lemma Collect_subset [simp]: "{x\<in>A. P x} \<subseteq> A" by auto
+
+lemma (in sigma_finite_measure) measurable_measure[measurable (raw)]:
+ "(\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M) \<Longrightarrow>
+ {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M) \<Longrightarrow>
+ (\<lambda>x. measure M (A x)) \<in> borel_measurable N"
+ unfolding measure_def by (intro measurable_emeasure borel_measurable_enn2real) auto
+
+lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:
+ fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f[measurable]: "case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
+ shows "(\<lambda>x. \<integral>y. f x y \<partial>M) \<in> borel_measurable N"
+proof -
+ from borel_measurable_implies_sequence_metric[OF f, of 0] guess s ..
+ then have s: "\<And>i. simple_function (N \<Otimes>\<^sub>M M) (s i)"
+ "\<And>x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> (\<lambda>i. s i (x, y)) \<longlonglongrightarrow> f x y"
+ "\<And>i x y. x \<in> space N \<Longrightarrow> y \<in> space M \<Longrightarrow> norm (s i (x, y)) \<le> 2 * norm (f x y)"
+ by (auto simp: space_pair_measure)
+
+ have [measurable]: "\<And>i. s i \<in> borel_measurable (N \<Otimes>\<^sub>M M)"
+ by (rule borel_measurable_simple_function) fact
+
+ have "\<And>i. s i \<in> measurable (N \<Otimes>\<^sub>M M) (count_space UNIV)"
+ by (rule measurable_simple_function) fact
+
+ define f' where [abs_def]: "f' i x =
+ (if integrable M (f x) then simple_bochner_integral M (\<lambda>y. s i (x, y)) else 0)" for i x
+
+ { fix i x assume "x \<in> space N"
+ then have "simple_bochner_integral M (\<lambda>y. s i (x, y)) =
+ (\<Sum>z\<in>s i ` (space N \<times> space M). measure M {y \<in> 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_neutral_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 \<in> borel_measurable N"
+ unfolding f'_def by (simp_all add: eq cong: measurable_cong if_cong)
+ next
+ fix x assume x: "x \<in> space N"
+ { assume int_f: "integrable M (f x)"
+ have int_2f: "integrable M (\<lambda>y. 2 * norm (f x y))"
+ by (intro integrable_norm integrable_mult_right int_f)
+ have "(\<lambda>i. integral\<^sup>L M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
+ proof (rule integral_dominated_convergence)
+ from int_f show "f x \<in> borel_measurable M" by auto
+ show "\<And>i. (\<lambda>y. s i (x, y)) \<in> borel_measurable M"
+ using x by simp
+ show "AE xa in M. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f x xa"
+ using x s(2) by auto
+ show "\<And>i. AE xa in M. norm (s i (x, xa)) \<le> 2 * norm (f x xa)"
+ using x s(3) by auto
+ qed fact
+ moreover
+ { fix i
+ have "simple_bochner_integrable M (\<lambda>y. s i (x, y))"
+ proof (rule simple_bochner_integrableI_bounded)
+ have "(\<lambda>y. s i (x, y)) ` space M \<subseteq> s i ` (space N \<times> space M)"
+ using x by auto
+ then show "simple_function M (\<lambda>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 "(\<integral>\<^sup>+ y. ennreal (norm (s i (x, y))) \<partial>M) \<le> (\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M)"
+ using x s by (intro nn_integral_mono) auto
+ also have "(\<integral>\<^sup>+ y. 2 * norm (f x y) \<partial>M) < \<infinity>"
+ using int_2f by (simp add: integrable_iff_bounded)
+ finally show "(\<integral>\<^sup>+ xa. ennreal (norm (s i (x, xa))) \<partial>M) < \<infinity>" .
+ qed
+ then have "integral\<^sup>L M (\<lambda>y. s i (x, y)) = simple_bochner_integral M (\<lambda>y. s i (x, y))"
+ by (rule simple_bochner_integrable_eq_integral[symmetric]) }
+ ultimately have "(\<lambda>i. simple_bochner_integral M (\<lambda>y. s i (x, y))) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
+ by simp }
+ then
+ show "(\<lambda>i. f' i x) \<longlonglongrightarrow> integral\<^sup>L M (f x)"
+ unfolding f'_def
+ by (cases "integrable M (f x)") (simp_all add: not_integrable_integral_eq)
+ qed
+qed
+
+lemma (in pair_sigma_finite) integrable_product_swap:
+ fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes "integrable (M1 \<Otimes>\<^sub>M M2) f"
+ shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))"
+proof -
+ interpret Q: pair_sigma_finite M2 M1 ..
+ have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(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 :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^sub>M M2) f"
+proof -
+ interpret Q: pair_sigma_finite M2 M1 ..
+ from Q.integrable_product_swap[of "\<lambda>(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 :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
+proof -
+ have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(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) Fubini_integrable:
+ fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f[measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ and integ1: "integrable M1 (\<lambda>x. \<integral> y. norm (f (x, y)) \<partial>M2)"
+ and integ2: "AE x in M1. integrable M2 (\<lambda>y. f (x, y))"
+ shows "integrable (M1 \<Otimes>\<^sub>M M2) f"
+proof (rule integrableI_bounded)
+ have "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
+ by (simp add: M2.nn_integral_fst [symmetric])
+ also have "\<dots> = (\<integral>\<^sup>+ x. \<bar>\<integral>y. norm (f (x, y)) \<partial>M2\<bar> \<partial>M1)"
+ apply (intro nn_integral_cong_AE)
+ using integ2
+ proof eventually_elim
+ fix x assume "integrable M2 (\<lambda>y. f (x, y))"
+ then have f: "integrable M2 (\<lambda>y. norm (f (x, y)))"
+ by simp
+ then have "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal (LINT y|M2. norm (f (x, y)))"
+ by (rule nn_integral_eq_integral) simp
+ also have "\<dots> = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>"
+ using f by simp
+ finally show "(\<integral>\<^sup>+y. ennreal (norm (f (x, y))) \<partial>M2) = ennreal \<bar>LINT y|M2. norm (f (x, y))\<bar>" .
+ qed
+ also have "\<dots> < \<infinity>"
+ using integ1 by (simp add: integrable_iff_bounded integral_nonneg_AE)
+ finally show "(\<integral>\<^sup>+ p. norm (f p) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>" .
+qed fact
+
+lemma (in pair_sigma_finite) emeasure_pair_measure_finite:
+ assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" and finite: "emeasure (M1 \<Otimes>\<^sub>M M2) A < \<infinity>"
+ shows "AE x in M1. emeasure M2 {y\<in>space M2. (x, y) \<in> A} < \<infinity>"
+proof -
+ from M2.emeasure_pair_measure_alt[OF A] finite
+ have "(\<integral>\<^sup>+ x. emeasure M2 (Pair x -` A) \<partial>M1) \<noteq> \<infinity>"
+ by simp
+ then have "AE x in M1. emeasure M2 (Pair x -` A) \<noteq> \<infinity>"
+ by (rule nn_integral_PInf_AE[rotated]) (intro M2.measurable_emeasure_Pair A)
+ moreover have "\<And>x. x \<in> space M1 \<Longrightarrow> Pair x -` A = {y\<in>space M2. (x, y) \<in> A}"
+ using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)
+ ultimately show ?thesis by (auto simp: less_top)
+qed
+
+lemma (in pair_sigma_finite) AE_integrable_fst':
+ fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f"
+ shows "AE x in M1. integrable M2 (\<lambda>y. f (x, y))"
+proof -
+ have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))"
+ by (rule M2.nn_integral_fst) simp
+ also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) \<noteq> \<infinity>"
+ using f unfolding integrable_iff_bounded by simp
+ finally have "AE x in M1. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
+ by (intro nn_integral_PInf_AE M2.borel_measurable_nn_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]] less_top)
+qed
+
+lemma (in pair_sigma_finite) integrable_fst':
+ fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) f"
+ shows "integrable M1 (\<lambda>x. \<integral>y. f (x, y) \<partial>M2)"
+ unfolding integrable_iff_bounded
+proof
+ show "(\<lambda>x. \<integral> y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+ by (rule M2.borel_measurable_lebesgue_integral) simp
+ have "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) \<le> (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1)"
+ using AE_integrable_fst'[OF f] by (auto intro!: nn_integral_mono_AE integral_norm_bound_ennreal)
+ also have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. norm (f (x, y)) \<partial>M2) \<partial>M1) = (\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2))"
+ by (rule M2.nn_integral_fst) simp
+ also have "(\<integral>\<^sup>+x. norm (f x) \<partial>(M1 \<Otimes>\<^sub>M M2)) < \<infinity>"
+ using f unfolding integrable_iff_bounded by simp
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<integral> y. f (x, y) \<partial>M2)) \<partial>M1) < \<infinity>" .
+qed
+
+lemma (in pair_sigma_finite) integral_fst':
+ fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"
+ shows "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
+using f proof induct
+ case (base A c)
+ have A[measurable]: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" by fact
+
+ have eq: "\<And>x y. x \<in> space M1 \<Longrightarrow> indicator A (x, y) = indicator {y\<in>space M2. (x, y) \<in> A} y"
+ using sets.sets_into_space[OF A] by (auto split: split_indicator simp: space_pair_measure)
+
+ have int_A: "integrable (M1 \<Otimes>\<^sub>M M2) (indicator A :: _ \<Rightarrow> real)"
+ using base by (rule integrable_real_indicator)
+
+ have "(\<integral> x. \<integral> y. indicator A (x, y) *\<^sub>R c \<partial>M2 \<partial>M1) = (\<integral>x. measure M2 {y\<in>space M2. (x, y) \<in> A} *\<^sub>R c \<partial>M1)"
+ proof (intro integral_cong_AE, simp, simp)
+ from AE_integrable_fst'[OF int_A] AE_space
+ show "AE x in M1. (\<integral>y. indicator A (x, y) *\<^sub>R c \<partial>M2) = measure M2 {y\<in>space M2. (x, y) \<in> A} *\<^sub>R c"
+ by eventually_elim (simp add: eq integrable_indicator_iff)
+ qed
+ also have "\<dots> = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c"
+ proof (subst integral_scaleR_left)
+ have "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) =
+ (\<integral>\<^sup>+x. emeasure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1)"
+ using emeasure_pair_measure_finite[OF base]
+ by (intro nn_integral_cong_AE, eventually_elim) (simp add: emeasure_eq_ennreal_measure)
+ also have "\<dots> = emeasure (M1 \<Otimes>\<^sub>M M2) A"
+ using sets.sets_into_space[OF A]
+ by (subst M2.emeasure_pair_measure_alt)
+ (auto intro!: nn_integral_cong arg_cong[where f="emeasure M2"] simp: space_pair_measure)
+ finally have *: "(\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1) = emeasure (M1 \<Otimes>\<^sub>M M2) A" .
+
+ from base * show "integrable M1 (\<lambda>x. measure M2 {y \<in> space M2. (x, y) \<in> A})"
+ by (simp add: integrable_iff_bounded)
+ then have "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) =
+ (\<integral>\<^sup>+x. ennreal (measure M2 {y \<in> space M2. (x, y) \<in> A}) \<partial>M1)"
+ by (rule nn_integral_eq_integral[symmetric]) simp
+ also note *
+ finally show "(\<integral>x. measure M2 {y \<in> space M2. (x, y) \<in> A} \<partial>M1) *\<^sub>R c = measure (M1 \<Otimes>\<^sub>M M2) A *\<^sub>R c"
+ using base by (simp add: emeasure_eq_ennreal_measure)
+ qed
+ also have "\<dots> = (\<integral> a. indicator A a *\<^sub>R c \<partial>(M1 \<Otimes>\<^sub>M M2))"
+ using base by simp
+ finally show ?case .
+next
+ case (add f g)
+ then have [measurable]: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "g \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ by auto
+ have "(\<integral> x. \<integral> y. f (x, y) + g (x, y) \<partial>M2 \<partial>M1) =
+ (\<integral> x. (\<integral> y. f (x, y) \<partial>M2) + (\<integral> y. g (x, y) \<partial>M2) \<partial>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 "\<dots> = (\<integral> x. f x \<partial>(M1 \<Otimes>\<^sub>M M2)) + (\<integral> x. g x \<partial>(M1 \<Otimes>\<^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 \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" "\<And>i. s i \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"
+ by auto
+
+ show ?case
+ proof (rule LIMSEQ_unique)
+ show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"
+ proof (rule integral_dominated_convergence)
+ show "integrable (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. 2 * norm (f x))"
+ using lim(5) by auto
+ qed (insert lim, auto)
+ have "(\<lambda>i. \<integral> x. \<integral> y. s i (x, y) \<partial>M2 \<partial>M1) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1"
+ proof (rule integral_dominated_convergence)
+ have "AE x in M1. \<forall>i. integrable M2 (\<lambda>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. (\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2"
+ proof eventually_elim
+ fix x assume x: "x \<in> space M1" and
+ s: "\<forall>i. integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))"
+ show "(\<lambda>i. \<integral> y. s i (x, y) \<partial>M2) \<longlonglongrightarrow> \<integral> y. f (x, y) \<partial>M2"
+ proof (rule integral_dominated_convergence)
+ show "integrable M2 (\<lambda>y. 2 * norm (f (x, y)))"
+ using f by auto
+ show "AE xa in M2. (\<lambda>i. s i (x, xa)) \<longlonglongrightarrow> f (x, xa)"
+ using x lim(3) by (auto simp: space_pair_measure)
+ show "\<And>i. AE xa in M2. norm (s i (x, xa)) \<le> 2 * norm (f (x, xa))"
+ using x lim(4) by (auto simp: space_pair_measure)
+ qed (insert x, measurable)
+ qed
+ show "integrable M1 (\<lambda>x. (\<integral> y. 2 * norm (f (x, y)) \<partial>M2))"
+ by (intro integrable_mult_right integrable_norm integrable_fst' lim)
+ fix i show "AE x in M1. norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>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 \<in> space M1"
+ and s: "integrable M2 (\<lambda>y. s i (x, y))" and f: "integrable M2 (\<lambda>y. f (x, y))"
+ from s have "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral>\<^sup>+y. norm (s i (x, y)) \<partial>M2)"
+ by (rule integral_norm_bound_ennreal)
+ also have "\<dots> \<le> (\<integral>\<^sup>+y. 2 * norm (f (x, y)) \<partial>M2)"
+ using x lim by (auto intro!: nn_integral_mono simp: space_pair_measure)
+ also have "\<dots> = (\<integral>y. 2 * norm (f (x, y)) \<partial>M2)"
+ using f by (intro nn_integral_eq_integral) auto
+ finally show "norm (\<integral> y. s i (x, y) \<partial>M2) \<le> (\<integral> y. 2 * norm (f (x, y)) \<partial>M2)"
+ by simp
+ qed
+ qed simp_all
+ then show "(\<lambda>i. integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (s i)) \<longlonglongrightarrow> \<integral> x. \<integral> y. f (x, y) \<partial>M2 \<partial>M1"
+ using lim by simp
+ qed
+qed
+
+lemma (in pair_sigma_finite)
+ fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
+ shows AE_integrable_fst: "AE x in M1. integrable M2 (\<lambda>y. f x y)" (is "?AE")
+ and integrable_fst: "integrable M1 (\<lambda>x. \<integral>y. f x y \<partial>M2)" (is "?INT")
+ and integral_fst: "(\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (\<lambda>(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 :: "_ \<Rightarrow> _ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes f[measurable]: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
+ shows AE_integrable_snd: "AE y in M2. integrable M1 (\<lambda>x. f x y)" (is "?AE")
+ and integrable_snd: "integrable M2 (\<lambda>y. \<integral>x. f x y \<partial>M1)" (is "?INT")
+ and integral_snd: "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) (case_prod f)" (is "?EQ")
+proof -
+ interpret Q: pair_sigma_finite M2 M1 ..
+ have Q_int: "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(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 "case_prod f"] by simp
+qed
+
+lemma (in pair_sigma_finite) Fubini_integral:
+ fixes f :: "_ \<Rightarrow> _ \<Rightarrow> _ :: {banach, second_countable_topology}"
+ assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) (case_prod f)"
+ shows "(\<integral>y. (\<integral>x. f x y \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f x y \<partial>M2) \<partial>M1)"
+ unfolding integral_snd[OF assms] integral_fst[OF assms] ..
+
+lemma (in product_sigma_finite) product_integral_singleton:
+ fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ shows "f \<in> borel_measurable (M i) \<Longrightarrow> (\<integral>x. f (x i) \<partial>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 :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
+ and f: "integrable (Pi\<^sub>M (I \<union> J) M) f"
+ shows "integral\<^sup>L (Pi\<^sub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^sub>M J M) \<partial>Pi\<^sub>M I M)"
+proof -
+ interpret I: finite_product_sigma_finite M I by standard fact
+ interpret J: finite_product_sigma_finite M J by standard fact
+ have "finite (I \<union> J)" using fin by auto
+ interpret IJ: finite_product_sigma_finite M "I \<union> J" by standard fact
+ interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
+ let ?M = "merge I J"
+ let ?f = "\<lambda>x. f (?M x)"
+ from f have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
+ by auto
+ have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^sub>M I M \<Otimes>\<^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 \<Otimes>\<^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 :: "_ \<Rightarrow> _::{banach, second_countable_topology}"
+ assumes I: "finite I" "i \<notin> I"
+ and f: "integrable (Pi\<^sub>M (insert i I) M) f"
+ shows "integral\<^sup>L (Pi\<^sub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>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 \<union> {i}) M) f"
+ by simp
+ also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) \<partial>Pi\<^sub>M I M)"
+ using f I by (intro product_integral_fold) auto
+ also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^sub>M I M)"
+ proof (rule integral_cong[OF refl], subst product_integral_singleton[symmetric])
+ fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
+ have f_borel: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
+ using f by auto
+ show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
+ using measurable_comp[OF measurable_component_update f_borel, OF x \<open>i \<notin> I\<close>]
+ unfolding comp_def .
+ from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^sub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>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 \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}"
+ assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+ shows "integrable (Pi\<^sub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f")
+proof (unfold integrable_iff_bounded, intro conjI)
+ interpret finite_product_sigma_finite M I by standard fact
+
+ show "?f \<in> borel_measurable (Pi\<^sub>M I M)"
+ using assms by simp
+ have "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) =
+ (\<integral>\<^sup>+ x. (\<Prod>i\<in>I. ennreal (norm (f i (x i)))) \<partial>Pi\<^sub>M I M)"
+ by (simp add: setprod_norm setprod_ennreal)
+ also have "\<dots> = (\<Prod>i\<in>I. \<integral>\<^sup>+ x. ennreal (norm (f i x)) \<partial>M i)"
+ using assms by (intro product_nn_integral_setprod) auto
+ also have "\<dots> < \<infinity>"
+ using integrable by (simp add: less_top[symmetric] ennreal_setprod_eq_top integrable_iff_bounded)
+ finally show "(\<integral>\<^sup>+ x. ennreal (norm (\<Prod>i\<in>I. f i (x i))) \<partial>Pi\<^sub>M I M) < \<infinity>" .
+qed
+
+lemma (in product_sigma_finite) product_integral_setprod:
+ fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> _::{real_normed_field,banach,second_countable_topology}"
+ assumes "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+ shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>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: "\<And>J. J \<subseteq> insert i I \<Longrightarrow>
+ integrable (Pi\<^sub>M J M) (\<lambda>x. (\<Prod>i\<in>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 standard fact
+ have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))"
+ using \<open>i \<notin> I\<close> 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 \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes borel: "f \<in> borel_measurable N"
+ and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
+ shows "integrable N f \<longleftrightarrow> integrable M f" (is ?P)
+proof -
+ have "f \<in> borel_measurable M"
+ using assms by (auto simp: measurable_def)
+ with assms show ?thesis
+ using assms by (auto simp: integrable_iff_bounded nn_integral_subalgebra)
+qed
+
+lemma integral_subalgebra:
+ fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes borel: "f \<in> borel_measurable N"
+ and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> 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 "(\<integral> a. f a + g a \<partial>N) = integral\<^sup>L M f + integral\<^sup>L M g"
+ by simp
+ also have "\<dots> = (\<integral> a. f a + g a \<partial>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: "\<And>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 "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L N f"
+ apply (rule integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"])
+ using lim
+ apply auto
+ done
+ show "(\<lambda>i. integral\<^sup>L N (s i)) \<longlonglongrightarrow> integral\<^sup>L M f"
+ unfolding lim
+ apply (rule integral_dominated_convergence[where w="\<lambda>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 (open) simple_bochner_integral
+hide_const (open) simple_bochner_integrable
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Borel_Space.thy Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,1915 @@
+(* Title: HOL/Probability/Borel_Space.thy
+ Author: Johannes Hölzl, TU München
+ Author: Armin Heller, TU München
+*)
+
+section \<open>Borel spaces\<close>
+
+theory Borel_Space
+imports
+ Measurable Derivative Ordered_Euclidean_Space Extended_Real_Limits
+begin
+
+lemma sets_Collect_eventually_sequentially[measurable]:
+ "(\<And>i. {x\<in>space M. P x i} \<in> sets M) \<Longrightarrow> {x\<in>space M. eventually (P x) sequentially} \<in> sets M"
+ unfolding eventually_sequentially by simp
+
+lemma topological_basis_trivial: "topological_basis {A. open A}"
+ by (auto simp: topological_basis_def)
+
+lemma open_prod_generated: "open = generate_topology {A \<times> B | A B. open A \<and> open B}"
+proof -
+ have "{A \<times> B :: ('a \<times> 'b) set | A B. open A \<and> open B} = ((\<lambda>(a, b). a \<times> b) ` ({A. open A} \<times> {A. open A}))"
+ by auto
+ then show ?thesis
+ by (auto intro: topological_basis_prod topological_basis_trivial topological_basis_imp_subbasis)
+qed
+
+definition "mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r \<le> s \<longrightarrow> f r \<le> f s"
+
+lemma mono_onI:
+ "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r \<le> s \<Longrightarrow> f r \<le> f s) \<Longrightarrow> mono_on f A"
+ unfolding mono_on_def by simp
+
+lemma mono_onD:
+ "\<lbrakk>mono_on f A; r \<in> A; s \<in> A; r \<le> s\<rbrakk> \<Longrightarrow> f r \<le> f s"
+ unfolding mono_on_def by simp
+
+lemma mono_imp_mono_on: "mono f \<Longrightarrow> mono_on f A"
+ unfolding mono_def mono_on_def by auto
+
+lemma mono_on_subset: "mono_on f A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> mono_on f B"
+ unfolding mono_on_def by auto
+
+definition "strict_mono_on f A \<equiv> \<forall>r s. r \<in> A \<and> s \<in> A \<and> r < s \<longrightarrow> f r < f s"
+
+lemma strict_mono_onI:
+ "(\<And>r s. r \<in> A \<Longrightarrow> s \<in> A \<Longrightarrow> r < s \<Longrightarrow> f r < f s) \<Longrightarrow> strict_mono_on f A"
+ unfolding strict_mono_on_def by simp
+
+lemma strict_mono_onD:
+ "\<lbrakk>strict_mono_on f A; r \<in> A; s \<in> A; r < s\<rbrakk> \<Longrightarrow> f r < f s"
+ unfolding strict_mono_on_def by simp
+
+lemma mono_on_greaterD:
+ assumes "mono_on g A" "x \<in> A" "y \<in> A" "g x > (g (y::_::linorder) :: _ :: linorder)"
+ shows "x > y"
+proof (rule ccontr)
+ assume "\<not>x > y"
+ hence "x \<le> y" by (simp add: not_less)
+ from assms(1-3) and this have "g x \<le> g y" by (rule mono_onD)
+ with assms(4) show False by simp
+qed
+
+lemma strict_mono_inv:
+ fixes f :: "('a::linorder) \<Rightarrow> ('b::linorder)"
+ assumes "strict_mono f" and "surj f" and inv: "\<And>x. g (f x) = x"
+ shows "strict_mono g"
+proof
+ fix x y :: 'b assume "x < y"
+ from \<open>surj f\<close> obtain x' y' where [simp]: "x = f x'" "y = f y'" by blast
+ with \<open>x < y\<close> and \<open>strict_mono f\<close> have "x' < y'" by (simp add: strict_mono_less)
+ with inv show "g x < g y" by simp
+qed
+
+lemma strict_mono_on_imp_inj_on:
+ assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> (_ :: preorder)) A"
+ shows "inj_on f A"
+proof (rule inj_onI)
+ fix x y assume "x \<in> A" "y \<in> A" "f x = f y"
+ thus "x = y"
+ by (cases x y rule: linorder_cases)
+ (auto dest: strict_mono_onD[OF assms, of x y] strict_mono_onD[OF assms, of y x])
+qed
+
+lemma strict_mono_on_leD:
+ assumes "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A" "x \<in> A" "y \<in> A" "x \<le> y"
+ shows "f x \<le> f y"
+proof (insert le_less_linear[of y x], elim disjE)
+ assume "x < y"
+ with assms have "f x < f y" by (rule_tac strict_mono_onD[OF assms(1)]) simp_all
+ thus ?thesis by (rule less_imp_le)
+qed (insert assms, simp)
+
+lemma strict_mono_on_eqD:
+ fixes f :: "(_ :: linorder) \<Rightarrow> (_ :: preorder)"
+ assumes "strict_mono_on f A" "f x = f y" "x \<in> A" "y \<in> A"
+ shows "y = x"
+ using assms by (rule_tac linorder_cases[of x y]) (auto dest: strict_mono_onD)
+
+lemma mono_on_imp_deriv_nonneg:
+ assumes mono: "mono_on f A" and deriv: "(f has_real_derivative D) (at x)"
+ assumes "x \<in> interior A"
+ shows "D \<ge> 0"
+proof (rule tendsto_le_const)
+ let ?A' = "(\<lambda>y. y - x) ` interior A"
+ from deriv show "((\<lambda>h. (f (x + h) - f x) / h) \<longlongrightarrow> D) (at 0)"
+ by (simp add: field_has_derivative_at has_field_derivative_def)
+ from mono have mono': "mono_on f (interior A)" by (rule mono_on_subset) (rule interior_subset)
+
+ show "eventually (\<lambda>h. (f (x + h) - f x) / h \<ge> 0) (at 0)"
+ proof (subst eventually_at_topological, intro exI conjI ballI impI)
+ have "open (interior A)" by simp
+ hence "open (op + (-x) ` interior A)" by (rule open_translation)
+ also have "(op + (-x) ` interior A) = ?A'" by auto
+ finally show "open ?A'" .
+ next
+ from \<open>x \<in> interior A\<close> show "0 \<in> ?A'" by auto
+ next
+ fix h assume "h \<in> ?A'"
+ hence "x + h \<in> interior A" by auto
+ with mono' and \<open>x \<in> interior A\<close> show "(f (x + h) - f x) / h \<ge> 0"
+ by (cases h rule: linorder_cases[of _ 0])
+ (simp_all add: divide_nonpos_neg divide_nonneg_pos mono_onD field_simps)
+ qed
+qed simp
+
+lemma strict_mono_on_imp_mono_on:
+ "strict_mono_on (f :: (_ :: linorder) \<Rightarrow> _ :: preorder) A \<Longrightarrow> mono_on f A"
+ by (rule mono_onI, rule strict_mono_on_leD)
+
+lemma mono_on_ctble_discont:
+ fixes f :: "real \<Rightarrow> real"
+ fixes A :: "real set"
+ assumes "mono_on f A"
+ shows "countable {a\<in>A. \<not> continuous (at a within A) f}"
+proof -
+ have mono: "\<And>x y. x \<in> A \<Longrightarrow> y \<in> A \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<le> f y"
+ using \<open>mono_on f A\<close> by (simp add: mono_on_def)
+ have "\<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}. \<exists>q :: nat \<times> rat.
+ (fst q = 0 \<and> of_rat (snd q) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd q))) \<or>
+ (fst q = 1 \<and> of_rat (snd q) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd q)))"
+ proof (clarsimp simp del: One_nat_def)
+ fix a assume "a \<in> A" assume "\<not> continuous (at a within A) f"
+ thus "\<exists>q1 q2.
+ q1 = 0 \<and> real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2) \<or>
+ q1 = 1 \<and> f a < real_of_rat q2 \<and> (\<forall>x\<in>A. a < x \<longrightarrow> real_of_rat q2 < f x)"
+ proof (auto simp add: continuous_within order_tendsto_iff eventually_at)
+ fix l assume "l < f a"
+ then obtain q2 where q2: "l < of_rat q2" "of_rat q2 < f a"
+ using of_rat_dense by blast
+ assume * [rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> l < f x"
+ from q2 have "real_of_rat q2 < f a \<and> (\<forall>x\<in>A. x < a \<longrightarrow> f x < real_of_rat q2)"
+ proof auto
+ fix x assume "x \<in> A" "x < a"
+ with q2 *[of "a - x"] show "f x < real_of_rat q2"
+ apply (auto simp add: dist_real_def not_less)
+ apply (subgoal_tac "f x \<le> f xa")
+ by (auto intro: mono)
+ qed
+ thus ?thesis by auto
+ next
+ fix u assume "u > f a"
+ then obtain q2 where q2: "f a < of_rat q2" "of_rat q2 < u"
+ using of_rat_dense by blast
+ assume *[rule_format]: "\<forall>d>0. \<exists>x\<in>A. x \<noteq> a \<and> dist x a < d \<and> \<not> u > f x"
+ from q2 have "real_of_rat q2 > f a \<and> (\<forall>x\<in>A. x > a \<longrightarrow> f x > real_of_rat q2)"
+ proof auto
+ fix x assume "x \<in> A" "x > a"
+ with q2 *[of "x - a"] show "f x > real_of_rat q2"
+ apply (auto simp add: dist_real_def)
+ apply (subgoal_tac "f x \<ge> f xa")
+ by (auto intro: mono)
+ qed
+ thus ?thesis by auto
+ qed
+ qed
+ hence "\<exists>g :: real \<Rightarrow> nat \<times> rat . \<forall>a \<in> {a\<in>A. \<not> continuous (at a within A) f}.
+ (fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (\<forall>x \<in> A. x < a \<longrightarrow> f x < of_rat (snd (g a)))) |
+ (fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (\<forall>x \<in> A. x > a \<longrightarrow> f x > of_rat (snd (g a))))"
+ by (rule bchoice)
+ then guess g ..
+ hence g: "\<And>a x. a \<in> A \<Longrightarrow> \<not> continuous (at a within A) f \<Longrightarrow> x \<in> A \<Longrightarrow>
+ (fst (g a) = 0 \<and> of_rat (snd (g a)) < f a \<and> (x < a \<longrightarrow> f x < of_rat (snd (g a)))) |
+ (fst (g a) = 1 \<and> of_rat (snd (g a)) > f a \<and> (x > a \<longrightarrow> f x > of_rat (snd (g a))))"
+ by auto
+ have "inj_on g {a\<in>A. \<not> continuous (at a within A) f}"
+ proof (auto simp add: inj_on_def)
+ fix w z
+ assume 1: "w \<in> A" and 2: "\<not> continuous (at w within A) f" and
+ 3: "z \<in> A" and 4: "\<not> continuous (at z within A) f" and
+ 5: "g w = g z"
+ from g [OF 1 2 3] g [OF 3 4 1] 5
+ show "w = z" by auto
+ qed
+ thus ?thesis
+ by (rule countableI')
+qed
+
+lemma mono_on_ctble_discont_open:
+ fixes f :: "real \<Rightarrow> real"
+ fixes A :: "real set"
+ assumes "open A" "mono_on f A"
+ shows "countable {a\<in>A. \<not>isCont f a}"
+proof -
+ have "{a\<in>A. \<not>isCont f a} = {a\<in>A. \<not>(continuous (at a within A) f)}"
+ by (auto simp add: continuous_within_open [OF _ \<open>open A\<close>])
+ thus ?thesis
+ apply (elim ssubst)
+ by (rule mono_on_ctble_discont, rule assms)
+qed
+
+lemma mono_ctble_discont:
+ fixes f :: "real \<Rightarrow> real"
+ assumes "mono f"
+ shows "countable {a. \<not> isCont f a}"
+using assms mono_on_ctble_discont [of f UNIV] unfolding mono_on_def mono_def by auto
+
+lemma has_real_derivative_imp_continuous_on:
+ assumes "\<And>x. x \<in> A \<Longrightarrow> (f has_real_derivative f' x) (at x)"
+ shows "continuous_on A f"
+ apply (intro differentiable_imp_continuous_on, unfold differentiable_on_def)
+ apply (intro ballI Deriv.differentiableI)
+ apply (rule has_field_derivative_subset[OF assms])
+ apply simp_all
+ done
+
+lemma closure_contains_Sup:
+ fixes S :: "real set"
+ assumes "S \<noteq> {}" "bdd_above S"
+ shows "Sup S \<in> closure S"
+proof-
+ have "Inf (uminus ` S) \<in> closure (uminus ` S)"
+ using assms by (intro closure_contains_Inf) auto
+ also have "Inf (uminus ` S) = -Sup S" by (simp add: Inf_real_def)
+ also have "closure (uminus ` S) = uminus ` closure S"
+ by (rule sym, intro closure_injective_linear_image) (auto intro: linearI)
+ finally show ?thesis by auto
+qed
+
+lemma closed_contains_Sup:
+ fixes S :: "real set"
+ shows "S \<noteq> {} \<Longrightarrow> bdd_above S \<Longrightarrow> closed S \<Longrightarrow> Sup S \<in> S"
+ by (subst closure_closed[symmetric], assumption, rule closure_contains_Sup)
+
+lemma deriv_nonneg_imp_mono:
+ assumes deriv: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
+ assumes nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
+ assumes ab: "a \<le> b"
+ shows "g a \<le> g b"
+proof (cases "a < b")
+ assume "a < b"
+ from deriv have "\<forall>x. x \<ge> a \<and> x \<le> b \<longrightarrow> (g has_real_derivative g' x) (at x)" by simp
+ from MVT2[OF \<open>a < b\<close> this] and deriv
+ obtain \<xi> where \<xi>_ab: "\<xi> > a" "\<xi> < b" and g_ab: "g b - g a = (b - a) * g' \<xi>" by blast
+ from \<xi>_ab ab nonneg have "(b - a) * g' \<xi> \<ge> 0" by simp
+ with g_ab show ?thesis by simp
+qed (insert ab, simp)
+
+lemma continuous_interval_vimage_Int:
+ assumes "continuous_on {a::real..b} g" and mono: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
+ assumes "a \<le> b" "(c::real) \<le> d" "{c..d} \<subseteq> {g a..g b}"
+ obtains c' d' where "{a..b} \<inter> g -` {c..d} = {c'..d'}" "c' \<le> d'" "g c' = c" "g d' = d"
+proof-
+ let ?A = "{a..b} \<inter> g -` {c..d}"
+ from IVT'[of g a c b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
+ obtain c'' where c'': "c'' \<in> ?A" "g c'' = c" by auto
+ from IVT'[of g a d b, OF _ _ \<open>a \<le> b\<close> assms(1)] assms(4,5)
+ obtain d'' where d'': "d'' \<in> ?A" "g d'' = d" by auto
+ hence [simp]: "?A \<noteq> {}" by blast
+
+ define c' where "c' = Inf ?A"
+ define d' where "d' = Sup ?A"
+ have "?A \<subseteq> {c'..d'}" unfolding c'_def d'_def
+ by (intro subsetI) (auto intro: cInf_lower cSup_upper)
+ moreover from assms have "closed ?A"
+ using continuous_on_closed_vimage[of "{a..b}" g] by (subst Int_commute) simp
+ hence c'd'_in_set: "c' \<in> ?A" "d' \<in> ?A" unfolding c'_def d'_def
+ by ((intro closed_contains_Inf closed_contains_Sup, simp_all)[])+
+ hence "{c'..d'} \<subseteq> ?A" using assms
+ by (intro subsetI)
+ (auto intro!: order_trans[of c "g c'" "g x" for x] order_trans[of "g x" "g d'" d for x]
+ intro!: mono)
+ moreover have "c' \<le> d'" using c'd'_in_set(2) unfolding c'_def by (intro cInf_lower) auto
+ moreover have "g c' \<le> c" "g d' \<ge> d"
+ apply (insert c'' d'' c'd'_in_set)
+ apply (subst c''(2)[symmetric])
+ apply (auto simp: c'_def intro!: mono cInf_lower c'') []
+ apply (subst d''(2)[symmetric])
+ apply (auto simp: d'_def intro!: mono cSup_upper d'') []
+ done
+ with c'd'_in_set have "g c' = c" "g d' = d" by auto
+ ultimately show ?thesis using that by blast
+qed
+
+subsection \<open>Generic Borel spaces\<close>
+
+definition (in topological_space) borel :: "'a measure" where
+ "borel = sigma UNIV {S. open S}"
+
+abbreviation "borel_measurable M \<equiv> measurable M borel"
+
+lemma in_borel_measurable:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
+ by (auto simp add: measurable_def borel_def)
+
+lemma in_borel_measurable_borel:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ (\<forall>S \<in> sets borel.
+ f -` S \<inter> space M \<in> sets M)"
+ by (auto simp add: measurable_def borel_def)
+
+lemma space_borel[simp]: "space borel = UNIV"
+ unfolding borel_def by auto
+
+lemma space_in_borel[measurable]: "UNIV \<in> sets borel"
+ unfolding borel_def by auto
+
+lemma sets_borel: "sets borel = sigma_sets UNIV {S. open S}"
+ unfolding borel_def by (rule sets_measure_of) simp
+
+lemma measurable_sets_borel:
+ "\<lbrakk>f \<in> measurable borel M; A \<in> sets M\<rbrakk> \<Longrightarrow> f -` A \<in> sets borel"
+ by (drule (1) measurable_sets) simp
+
+lemma pred_Collect_borel[measurable (raw)]: "Measurable.pred borel P \<Longrightarrow> {x. P x} \<in> sets borel"
+ unfolding borel_def pred_def by auto
+
+lemma borel_open[measurable (raw generic)]:
+ assumes "open A" shows "A \<in> sets borel"
+proof -
+ have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
+ thus ?thesis unfolding borel_def by auto
+qed
+
+lemma borel_closed[measurable (raw generic)]:
+ assumes "closed A" shows "A \<in> sets borel"
+proof -
+ have "space borel - (- A) \<in> sets borel"
+ using assms unfolding closed_def by (blast intro: borel_open)
+ thus ?thesis by simp
+qed
+
+lemma borel_singleton[measurable]:
+ "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets (borel :: 'a::t1_space measure)"
+ unfolding insert_def by (rule sets.Un) auto
+
+lemma borel_comp[measurable]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
+ unfolding Compl_eq_Diff_UNIV by simp
+
+lemma borel_measurable_vimage:
+ fixes f :: "'a \<Rightarrow> 'x::t2_space"
+ assumes borel[measurable]: "f \<in> borel_measurable M"
+ shows "f -` {x} \<inter> space M \<in> sets M"
+ by simp
+
+lemma borel_measurableI:
+ fixes f :: "'a \<Rightarrow> 'x::topological_space"
+ assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+ shows "f \<in> borel_measurable M"
+ unfolding borel_def
+proof (rule measurable_measure_of, simp_all)
+ fix S :: "'x set" assume "open S" thus "f -` S \<inter> space M \<in> sets M"
+ using assms[of S] by simp
+qed
+
+lemma borel_measurable_const:
+ "(\<lambda>x. c) \<in> borel_measurable M"
+ by auto
+
+lemma borel_measurable_indicator:
+ assumes A: "A \<in> sets M"
+ shows "indicator A \<in> borel_measurable M"
+ unfolding indicator_def [abs_def] using A
+ by (auto intro!: measurable_If_set)
+
+lemma borel_measurable_count_space[measurable (raw)]:
+ "f \<in> borel_measurable (count_space S)"
+ unfolding measurable_def by auto
+
+lemma borel_measurable_indicator'[measurable (raw)]:
+ assumes [measurable]: "{x\<in>space M. f x \<in> A x} \<in> sets M"
+ shows "(\<lambda>x. indicator (A x) (f x)) \<in> borel_measurable M"
+ unfolding indicator_def[abs_def]
+ by (auto intro!: measurable_If)
+
+lemma borel_measurable_indicator_iff:
+ "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
+ (is "?I \<in> borel_measurable M \<longleftrightarrow> _")
+proof
+ assume "?I \<in> borel_measurable M"
+ then have "?I -` {1} \<inter> space M \<in> sets M"
+ unfolding measurable_def by auto
+ also have "?I -` {1} \<inter> space M = A \<inter> space M"
+ unfolding indicator_def [abs_def] by auto
+ finally show "A \<inter> space M \<in> sets M" .
+next
+ assume "A \<inter> space M \<in> sets M"
+ moreover have "?I \<in> borel_measurable M \<longleftrightarrow>
+ (indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M"
+ by (intro measurable_cong) (auto simp: indicator_def)
+ ultimately show "?I \<in> borel_measurable M" by auto
+qed
+
+lemma borel_measurable_subalgebra:
+ assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N"
+ shows "f \<in> borel_measurable M"
+ using assms unfolding measurable_def by auto
+
+lemma borel_measurable_restrict_space_iff_ereal:
+ fixes f :: "'a \<Rightarrow> ereal"
+ assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
+ shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
+ (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"
+ by (subst measurable_restrict_space_iff)
+ (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
+
+lemma borel_measurable_restrict_space_iff_ennreal:
+ fixes f :: "'a \<Rightarrow> ennreal"
+ assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
+ shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
+ (\<lambda>x. f x * indicator \<Omega> x) \<in> borel_measurable M"
+ by (subst measurable_restrict_space_iff)
+ (auto simp: indicator_def if_distrib[where f="\<lambda>x. a * x" for a] cong del: if_weak_cong)
+
+lemma borel_measurable_restrict_space_iff:
+ fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+ assumes \<Omega>[measurable, simp]: "\<Omega> \<inter> space M \<in> sets M"
+ shows "f \<in> borel_measurable (restrict_space M \<Omega>) \<longleftrightarrow>
+ (\<lambda>x. indicator \<Omega> x *\<^sub>R f x) \<in> borel_measurable M"
+ by (subst measurable_restrict_space_iff)
+ (auto simp: indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a] ac_simps
+ cong del: if_weak_cong)
+
+lemma cbox_borel[measurable]: "cbox a b \<in> sets borel"
+ by (auto intro: borel_closed)
+
+lemma box_borel[measurable]: "box a b \<in> sets borel"
+ by (auto intro: borel_open)
+
+lemma borel_compact: "compact (A::'a::t2_space set) \<Longrightarrow> A \<in> sets borel"
+ by (auto intro: borel_closed dest!: compact_imp_closed)
+
+lemma borel_sigma_sets_subset:
+ "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
+ using sets.sigma_sets_subset[of A borel] by simp
+
+lemma borel_eq_sigmaI1:
+ fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
+ assumes borel_eq: "borel = sigma UNIV X"
+ assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (F ` A))"
+ assumes F: "\<And>i. i \<in> A \<Longrightarrow> F i \<in> sets borel"
+ shows "borel = sigma UNIV (F ` A)"
+ unfolding borel_def
+proof (intro sigma_eqI antisym)
+ have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
+ unfolding borel_def by simp
+ also have "\<dots> = sigma_sets UNIV X"
+ unfolding borel_eq by simp
+ also have "\<dots> \<subseteq> sigma_sets UNIV (F`A)"
+ using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
+ finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (F`A)" .
+ show "sigma_sets UNIV (F`A) \<subseteq> sigma_sets UNIV {S. open S}"
+ unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
+qed auto
+
+lemma borel_eq_sigmaI2:
+ fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
+ and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
+ assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`B)"
+ assumes X: "\<And>i j. (i, j) \<in> B \<Longrightarrow> G i j \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
+ assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
+ shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
+ using assms
+ by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` B" and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_sigmaI3:
+ fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
+ assumes borel_eq: "borel = sigma UNIV X"
+ assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV ((\<lambda>(i, j). F i j) ` A))"
+ assumes F: "\<And>i j. (i, j) \<in> A \<Longrightarrow> F i j \<in> sets borel"
+ shows "borel = sigma UNIV ((\<lambda>(i, j). F i j) ` A)"
+ using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_sigmaI4:
+ fixes F :: "'i \<Rightarrow> 'a::topological_space set"
+ and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
+ assumes borel_eq: "borel = sigma UNIV ((\<lambda>(i, j). G i j)`A)"
+ assumes X: "\<And>i j. (i, j) \<in> A \<Longrightarrow> G i j \<in> sets (sigma UNIV (range F))"
+ assumes F: "\<And>i. F i \<in> sets borel"
+ shows "borel = sigma UNIV (range F)"
+ using assms by (intro borel_eq_sigmaI1[where X="(\<lambda>(i, j). G i j) ` A" and F=F]) auto
+
+lemma borel_eq_sigmaI5:
+ fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
+ assumes borel_eq: "borel = sigma UNIV (range G)"
+ assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
+ assumes F: "\<And>i j. F i j \<in> sets borel"
+ shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
+ using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma second_countable_borel_measurable:
+ fixes X :: "'a::second_countable_topology set set"
+ assumes eq: "open = generate_topology X"
+ shows "borel = sigma UNIV X"
+ unfolding borel_def
+proof (intro sigma_eqI sigma_sets_eqI)
+ interpret X: sigma_algebra UNIV "sigma_sets UNIV X"
+ by (rule sigma_algebra_sigma_sets) simp
+
+ fix S :: "'a set" assume "S \<in> Collect open"
+ then have "generate_topology X S"
+ by (auto simp: eq)
+ then show "S \<in> sigma_sets UNIV X"
+ proof induction
+ case (UN K)
+ then have K: "\<And>k. k \<in> K \<Longrightarrow> open k"
+ unfolding eq by auto
+ from ex_countable_basis obtain B :: "'a set set" where
+ B: "\<And>b. b \<in> B \<Longrightarrow> open b" "\<And>X. open X \<Longrightarrow> \<exists>b\<subseteq>B. (\<Union>b) = X" and "countable B"
+ by (auto simp: topological_basis_def)
+ from B(2)[OF K] obtain m where m: "\<And>k. k \<in> K \<Longrightarrow> m k \<subseteq> B" "\<And>k. k \<in> K \<Longrightarrow> (\<Union>m k) = k"
+ by metis
+ define U where "U = (\<Union>k\<in>K. m k)"
+ with m have "countable U"
+ by (intro countable_subset[OF _ \<open>countable B\<close>]) auto
+ have "\<Union>U = (\<Union>A\<in>U. A)" by simp
+ also have "\<dots> = \<Union>K"
+ unfolding U_def UN_simps by (simp add: m)
+ finally have "\<Union>U = \<Union>K" .
+
+ have "\<forall>b\<in>U. \<exists>k\<in>K. b \<subseteq> k"
+ using m by (auto simp: U_def)
+ then obtain u where u: "\<And>b. b \<in> U \<Longrightarrow> u b \<in> K" and "\<And>b. b \<in> U \<Longrightarrow> b \<subseteq> u b"
+ by metis
+ then have "(\<Union>b\<in>U. u b) \<subseteq> \<Union>K" "\<Union>U \<subseteq> (\<Union>b\<in>U. u b)"
+ by auto
+ then have "\<Union>K = (\<Union>b\<in>U. u b)"
+ unfolding \<open>\<Union>U = \<Union>K\<close> by auto
+ also have "\<dots> \<in> sigma_sets UNIV X"
+ using u UN by (intro X.countable_UN' \<open>countable U\<close>) auto
+ finally show "\<Union>K \<in> sigma_sets UNIV X" .
+ qed auto
+qed (auto simp: eq intro: generate_topology.Basis)
+
+lemma borel_eq_closed: "borel = sigma UNIV (Collect closed)"
+ unfolding borel_def
+proof (intro sigma_eqI sigma_sets_eqI, safe)
+ fix x :: "'a set" assume "open x"
+ hence "x = UNIV - (UNIV - x)" by auto
+ also have "\<dots> \<in> sigma_sets UNIV (Collect closed)"
+ by (force intro: sigma_sets.Compl simp: \<open>open x\<close>)
+ finally show "x \<in> sigma_sets UNIV (Collect closed)" by simp
+next
+ fix x :: "'a set" assume "closed x"
+ hence "x = UNIV - (UNIV - x)" by auto
+ also have "\<dots> \<in> sigma_sets UNIV (Collect open)"
+ by (force intro: sigma_sets.Compl simp: \<open>closed x\<close>)
+ finally show "x \<in> sigma_sets UNIV (Collect open)" by simp
+qed simp_all
+
+lemma borel_eq_countable_basis:
+ fixes B::"'a::topological_space set set"
+ assumes "countable B"
+ assumes "topological_basis B"
+ shows "borel = sigma UNIV B"
+ unfolding borel_def
+proof (intro sigma_eqI sigma_sets_eqI, safe)
+ interpret countable_basis using assms by unfold_locales
+ fix X::"'a set" assume "open X"
+ from open_countable_basisE[OF this] guess B' . note B' = this
+ then show "X \<in> sigma_sets UNIV B"
+ by (blast intro: sigma_sets_UNION \<open>countable B\<close> countable_subset)
+next
+ fix b assume "b \<in> B"
+ hence "open b" by (rule topological_basis_open[OF assms(2)])
+ thus "b \<in> sigma_sets UNIV (Collect open)" by auto
+qed simp_all
+
+lemma borel_measurable_continuous_on_restrict:
+ fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
+ assumes f: "continuous_on A f"
+ shows "f \<in> borel_measurable (restrict_space borel A)"
+proof (rule borel_measurableI)
+ fix S :: "'b set" assume "open S"
+ with f obtain T where "f -` S \<inter> A = T \<inter> A" "open T"
+ by (metis continuous_on_open_invariant)
+ then show "f -` S \<inter> space (restrict_space borel A) \<in> sets (restrict_space borel A)"
+ by (force simp add: sets_restrict_space space_restrict_space)
+qed
+
+lemma borel_measurable_continuous_on1: "continuous_on UNIV f \<Longrightarrow> f \<in> borel_measurable borel"
+ by (drule borel_measurable_continuous_on_restrict) simp
+
+lemma borel_measurable_continuous_on_if:
+ "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> continuous_on (- A) g \<Longrightarrow>
+ (\<lambda>x. if x \<in> A then f x else g x) \<in> borel_measurable borel"
+ by (auto simp add: measurable_If_restrict_space_iff Collect_neg_eq
+ intro!: borel_measurable_continuous_on_restrict)
+
+lemma borel_measurable_continuous_countable_exceptions:
+ fixes f :: "'a::t1_space \<Rightarrow> 'b::topological_space"
+ assumes X: "countable X"
+ assumes "continuous_on (- X) f"
+ shows "f \<in> borel_measurable borel"
+proof (rule measurable_discrete_difference[OF _ X])
+ have "X \<in> sets borel"
+ by (rule sets.countable[OF _ X]) auto
+ then show "(\<lambda>x. if x \<in> X then undefined else f x) \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on_if assms continuous_intros)
+qed auto
+
+lemma borel_measurable_continuous_on:
+ assumes f: "continuous_on UNIV f" and g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f (g x)) \<in> borel_measurable M"
+ using measurable_comp[OF g borel_measurable_continuous_on1[OF f]] by (simp add: comp_def)
+
+lemma borel_measurable_continuous_on_indicator:
+ fixes f g :: "'a::topological_space \<Rightarrow> 'b::real_normed_vector"
+ shows "A \<in> sets borel \<Longrightarrow> continuous_on A f \<Longrightarrow> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable borel"
+ by (subst borel_measurable_restrict_space_iff[symmetric])
+ (auto intro: borel_measurable_continuous_on_restrict)
+
+lemma borel_measurable_Pair[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
+ assumes f[measurable]: "f \<in> borel_measurable M"
+ assumes g[measurable]: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. (f x, g x)) \<in> borel_measurable M"
+proof (subst borel_eq_countable_basis)
+ let ?B = "SOME B::'b set set. countable B \<and> topological_basis B"
+ let ?C = "SOME B::'c set set. countable B \<and> topological_basis B"
+ let ?P = "(\<lambda>(b, c). b \<times> c) ` (?B \<times> ?C)"
+ show "countable ?P" "topological_basis ?P"
+ by (auto intro!: countable_basis topological_basis_prod is_basis)
+
+ show "(\<lambda>x. (f x, g x)) \<in> measurable M (sigma UNIV ?P)"
+ proof (rule measurable_measure_of)
+ fix S assume "S \<in> ?P"
+ then obtain b c where "b \<in> ?B" "c \<in> ?C" and S: "S = b \<times> c" by auto
+ then have borel: "open b" "open c"
+ by (auto intro: is_basis topological_basis_open)
+ have "(\<lambda>x. (f x, g x)) -` S \<inter> space M = (f -` b \<inter> space M) \<inter> (g -` c \<inter> space M)"
+ unfolding S by auto
+ also have "\<dots> \<in> sets M"
+ using borel by simp
+ finally show "(\<lambda>x. (f x, g x)) -` S \<inter> space M \<in> sets M" .
+ qed auto
+qed
+
+lemma borel_measurable_continuous_Pair:
+ fixes f :: "'a \<Rightarrow> 'b::second_countable_topology" and g :: "'a \<Rightarrow> 'c::second_countable_topology"
+ assumes [measurable]: "f \<in> borel_measurable M"
+ assumes [measurable]: "g \<in> borel_measurable M"
+ assumes H: "continuous_on UNIV (\<lambda>x. H (fst x) (snd x))"
+ shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
+proof -
+ have eq: "(\<lambda>x. H (f x) (g x)) = (\<lambda>x. (\<lambda>x. H (fst x) (snd x)) (f x, g x))" by auto
+ show ?thesis
+ unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
+qed
+
+subsection \<open>Borel spaces on order topologies\<close>
+
+lemma [measurable]:
+ fixes a b :: "'a::linorder_topology"
+ shows lessThan_borel: "{..< a} \<in> sets borel"
+ and greaterThan_borel: "{a <..} \<in> sets borel"
+ and greaterThanLessThan_borel: "{a<..<b} \<in> sets borel"
+ and atMost_borel: "{..a} \<in> sets borel"
+ and atLeast_borel: "{a..} \<in> sets borel"
+ and atLeastAtMost_borel: "{a..b} \<in> sets borel"
+ and greaterThanAtMost_borel: "{a<..b} \<in> sets borel"
+ and atLeastLessThan_borel: "{a..<b} \<in> sets borel"
+ unfolding greaterThanAtMost_def atLeastLessThan_def
+ by (blast intro: borel_open borel_closed open_lessThan open_greaterThan open_greaterThanLessThan
+ closed_atMost closed_atLeast closed_atLeastAtMost)+
+
+lemma borel_Iio:
+ "borel = sigma UNIV (range lessThan :: 'a::{linorder_topology, second_countable_topology} set set)"
+ unfolding second_countable_borel_measurable[OF open_generated_order]
+proof (intro sigma_eqI sigma_sets_eqI)
+ from countable_dense_setE guess D :: "'a set" . note D = this
+
+ interpret L: sigma_algebra UNIV "sigma_sets UNIV (range lessThan)"
+ by (rule sigma_algebra_sigma_sets) simp
+
+ fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan"
+ then obtain y where "A = {y <..} \<or> A = {..< y}"
+ by blast
+ then show "A \<in> sigma_sets UNIV (range lessThan)"
+ proof
+ assume A: "A = {y <..}"
+ show ?thesis
+ proof cases
+ assume "\<forall>x>y. \<exists>d. y < d \<and> d < x"
+ with D(2)[of "{y <..< x}" for x] have "\<forall>x>y. \<exists>d\<in>D. y < d \<and> d < x"
+ by (auto simp: set_eq_iff)
+ then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. y < d}. {..< d})"
+ by (auto simp: A) (metis less_asym)
+ also have "\<dots> \<in> sigma_sets UNIV (range lessThan)"
+ using D(1) by (intro L.Diff L.top L.countable_INT'') auto
+ finally show ?thesis .
+ next
+ assume "\<not> (\<forall>x>y. \<exists>d. y < d \<and> d < x)"
+ then obtain x where "y < x" "\<And>d. y < d \<Longrightarrow> \<not> d < x"
+ by auto
+ then have "A = UNIV - {..< x}"
+ unfolding A by (auto simp: not_less[symmetric])
+ also have "\<dots> \<in> sigma_sets UNIV (range lessThan)"
+ by auto
+ finally show ?thesis .
+ qed
+ qed auto
+qed auto
+
+lemma borel_Ioi:
+ "borel = sigma UNIV (range greaterThan :: 'a::{linorder_topology, second_countable_topology} set set)"
+ unfolding second_countable_borel_measurable[OF open_generated_order]
+proof (intro sigma_eqI sigma_sets_eqI)
+ from countable_dense_setE guess D :: "'a set" . note D = this
+
+ interpret L: sigma_algebra UNIV "sigma_sets UNIV (range greaterThan)"
+ by (rule sigma_algebra_sigma_sets) simp
+
+ fix A :: "'a set" assume "A \<in> range lessThan \<union> range greaterThan"
+ then obtain y where "A = {y <..} \<or> A = {..< y}"
+ by blast
+ then show "A \<in> sigma_sets UNIV (range greaterThan)"
+ proof
+ assume A: "A = {..< y}"
+ show ?thesis
+ proof cases
+ assume "\<forall>x<y. \<exists>d. x < d \<and> d < y"
+ with D(2)[of "{x <..< y}" for x] have "\<forall>x<y. \<exists>d\<in>D. x < d \<and> d < y"
+ by (auto simp: set_eq_iff)
+ then have "A = UNIV - (\<Inter>d\<in>{d\<in>D. d < y}. {d <..})"
+ by (auto simp: A) (metis less_asym)
+ also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)"
+ using D(1) by (intro L.Diff L.top L.countable_INT'') auto
+ finally show ?thesis .
+ next
+ assume "\<not> (\<forall>x<y. \<exists>d. x < d \<and> d < y)"
+ then obtain x where "x < y" "\<And>d. y > d \<Longrightarrow> x \<ge> d"
+ by (auto simp: not_less[symmetric])
+ then have "A = UNIV - {x <..}"
+ unfolding A Compl_eq_Diff_UNIV[symmetric] by auto
+ also have "\<dots> \<in> sigma_sets UNIV (range greaterThan)"
+ by auto
+ finally show ?thesis .
+ qed
+ qed auto
+qed auto
+
+lemma borel_measurableI_less:
+ fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
+ shows "(\<And>y. {x\<in>space M. f x < y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
+ unfolding borel_Iio
+ by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
+
+lemma borel_measurableI_greater:
+ fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
+ shows "(\<And>y. {x\<in>space M. y < f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
+ unfolding borel_Ioi
+ by (rule measurable_measure_of) (auto simp: Int_def conj_commute)
+
+lemma borel_measurableI_le:
+ fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
+ shows "(\<And>y. {x\<in>space M. f x \<le> y} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
+ by (rule borel_measurableI_greater) (auto simp: not_le[symmetric])
+
+lemma borel_measurableI_ge:
+ fixes f :: "'a \<Rightarrow> 'b::{linorder_topology, second_countable_topology}"
+ shows "(\<And>y. {x\<in>space M. y \<le> f x} \<in> sets M) \<Longrightarrow> f \<in> borel_measurable M"
+ by (rule borel_measurableI_less) (auto simp: not_le[symmetric])
+
+lemma borel_measurable_less[measurable]:
+ fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "{w \<in> space M. f w < g w} \<in> sets M"
+proof -
+ have "{w \<in> space M. f w < g w} = (\<lambda>x. (f x, g x)) -` {x. fst x < snd x} \<inter> space M"
+ by auto
+ also have "\<dots> \<in> sets M"
+ by (intro measurable_sets[OF borel_measurable_Pair borel_open, OF assms open_Collect_less]
+ continuous_intros)
+ finally show ?thesis .
+qed
+
+lemma
+ fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, linorder_topology}"
+ assumes f[measurable]: "f \<in> borel_measurable M"
+ assumes g[measurable]: "g \<in> borel_measurable M"
+ shows borel_measurable_le[measurable]: "{w \<in> space M. f w \<le> g w} \<in> sets M"
+ and borel_measurable_eq[measurable]: "{w \<in> space M. f w = g w} \<in> sets M"
+ and borel_measurable_neq: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
+ unfolding eq_iff not_less[symmetric]
+ by measurable
+
+lemma borel_measurable_SUP[measurable (raw)]:
+ fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
+ assumes [simp]: "countable I"
+ assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
+ shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M"
+ by (rule borel_measurableI_greater) (simp add: less_SUP_iff)
+
+lemma borel_measurable_INF[measurable (raw)]:
+ fixes F :: "_ \<Rightarrow> _ \<Rightarrow> _::{complete_linorder, linorder_topology, second_countable_topology}"
+ assumes [simp]: "countable I"
+ assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
+ shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
+ by (rule borel_measurableI_less) (simp add: INF_less_iff)
+
+lemma borel_measurable_cSUP[measurable (raw)]:
+ fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
+ assumes [simp]: "countable I"
+ assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
+ assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_above ((\<lambda>i. F i x) ` I)"
+ shows "(\<lambda>x. SUP i:I. F i x) \<in> borel_measurable M"
+proof cases
+ assume "I = {}" then show ?thesis
+ unfolding \<open>I = {}\<close> image_empty by simp
+next
+ assume "I \<noteq> {}"
+ show ?thesis
+ proof (rule borel_measurableI_le)
+ fix y
+ have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} \<in> sets M"
+ by measurable
+ also have "{x \<in> space M. \<forall>i\<in>I. F i x \<le> y} = {x \<in> space M. (SUP i:I. F i x) \<le> y}"
+ by (simp add: cSUP_le_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong)
+ finally show "{x \<in> space M. (SUP i:I. F i x) \<le> y} \<in> sets M" .
+ qed
+qed
+
+lemma borel_measurable_cINF[measurable (raw)]:
+ fixes F :: "_ \<Rightarrow> _ \<Rightarrow> 'a::{conditionally_complete_linorder, linorder_topology, second_countable_topology}"
+ assumes [simp]: "countable I"
+ assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> borel_measurable M"
+ assumes bdd: "\<And>x. x \<in> space M \<Longrightarrow> bdd_below ((\<lambda>i. F i x) ` I)"
+ shows "(\<lambda>x. INF i:I. F i x) \<in> borel_measurable M"
+proof cases
+ assume "I = {}" then show ?thesis
+ unfolding \<open>I = {}\<close> image_empty by simp
+next
+ assume "I \<noteq> {}"
+ show ?thesis
+ proof (rule borel_measurableI_ge)
+ fix y
+ have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} \<in> sets M"
+ by measurable
+ also have "{x \<in> space M. \<forall>i\<in>I. y \<le> F i x} = {x \<in> space M. y \<le> (INF i:I. F i x)}"
+ by (simp add: le_cINF_iff \<open>I \<noteq> {}\<close> bdd cong: conj_cong)
+ finally show "{x \<in> space M. y \<le> (INF i:I. F i x)} \<in> sets M" .
+ qed
+qed
+
+lemma borel_measurable_lfp[consumes 1, case_names continuity step]:
+ fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
+ assumes "sup_continuous F"
+ assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
+ shows "lfp F \<in> borel_measurable M"
+proof -
+ { fix i have "((F ^^ i) bot) \<in> borel_measurable M"
+ by (induct i) (auto intro!: *) }
+ then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> borel_measurable M"
+ by measurable
+ also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = (SUP i. (F ^^ i) bot)"
+ by auto
+ also have "(SUP i. (F ^^ i) bot) = lfp F"
+ by (rule sup_continuous_lfp[symmetric]) fact
+ finally show ?thesis .
+qed
+
+lemma borel_measurable_gfp[consumes 1, case_names continuity step]:
+ fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_linorder, linorder_topology, second_countable_topology})"
+ assumes "inf_continuous F"
+ assumes *: "\<And>f. f \<in> borel_measurable M \<Longrightarrow> F f \<in> borel_measurable M"
+ shows "gfp F \<in> borel_measurable M"
+proof -
+ { fix i have "((F ^^ i) top) \<in> borel_measurable M"
+ by (induct i) (auto intro!: * simp: bot_fun_def) }
+ then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> borel_measurable M"
+ by measurable
+ also have "(\<lambda>x. INF i. (F ^^ i) top x) = (INF i. (F ^^ i) top)"
+ by auto
+ also have "\<dots> = gfp F"
+ by (rule inf_continuous_gfp[symmetric]) fact
+ finally show ?thesis .
+qed
+
+lemma borel_measurable_max[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. max (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
+ by (rule borel_measurableI_less) simp
+
+lemma borel_measurable_min[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. min (g x) (f x) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
+ by (rule borel_measurableI_greater) simp
+
+lemma borel_measurable_Min[measurable (raw)]:
+ "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Min ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
+proof (induct I rule: finite_induct)
+ case (insert i I) then show ?case
+ by (cases "I = {}") auto
+qed auto
+
+lemma borel_measurable_Max[measurable (raw)]:
+ "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<lambda>x. Max ((\<lambda>i. f i x)`I) :: 'b::{second_countable_topology, linorder_topology}) \<in> borel_measurable M"
+proof (induct I rule: finite_induct)
+ case (insert i I) then show ?case
+ by (cases "I = {}") auto
+qed auto
+
+lemma borel_measurable_sup[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. sup (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
+ unfolding sup_max by measurable
+
+lemma borel_measurable_inf[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. inf (g x) (f x) :: 'b::{lattice, second_countable_topology, linorder_topology}) \<in> borel_measurable M"
+ unfolding inf_min by measurable
+
+lemma [measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
+ assumes "\<And>i. f i \<in> borel_measurable M"
+ shows borel_measurable_liminf: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
+ and borel_measurable_limsup: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
+ unfolding liminf_SUP_INF limsup_INF_SUP using assms by auto
+
+lemma measurable_convergent[measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
+ assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "Measurable.pred M (\<lambda>x. convergent (\<lambda>i. f i x))"
+ unfolding convergent_ereal by measurable
+
+lemma sets_Collect_convergent[measurable]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
+ assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "{x\<in>space M. convergent (\<lambda>i. f i x)} \<in> sets M"
+ by measurable
+
+lemma borel_measurable_lim[measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
+ assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
+proof -
+ have "\<And>x. lim (\<lambda>i. f i x) = (if convergent (\<lambda>i. f i x) then limsup (\<lambda>i. f i x) else (THE i. False))"
+ by (simp add: lim_def convergent_def convergent_limsup_cl)
+ then show ?thesis
+ by simp
+qed
+
+lemma borel_measurable_LIMSEQ_order:
+ fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology}"
+ assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
+ and u: "\<And>i. u i \<in> borel_measurable M"
+ shows "u' \<in> borel_measurable M"
+proof -
+ have "\<And>x. x \<in> space M \<Longrightarrow> u' x = liminf (\<lambda>n. u n x)"
+ using u' by (simp add: lim_imp_Liminf[symmetric])
+ with u show ?thesis by (simp cong: measurable_cong)
+qed
+
+subsection \<open>Borel spaces on topological monoids\<close>
+
+lemma borel_measurable_add[measurable (raw)]:
+ fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, topological_monoid_add}"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
+ using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+
+lemma borel_measurable_setsum[measurable (raw)]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, topological_comm_monoid_add}"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
+proof cases
+ assume "finite S"
+ thus ?thesis using assms by induct auto
+qed simp
+
+lemma borel_measurable_suminf_order[measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{complete_linorder, second_countable_topology, linorder_topology, topological_comm_monoid_add}"
+ assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
+ unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
+
+subsection \<open>Borel spaces on Euclidean spaces\<close>
+
+lemma borel_measurable_inner[measurable (raw)]:
+ fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x \<bullet> g x) \<in> borel_measurable M"
+ using assms
+ by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+
+notation
+ eucl_less (infix "<e" 50)
+
+lemma box_oc: "{x. a <e x \<and> x \<le> b} = {x. a <e x} \<inter> {..b}"
+ and box_co: "{x. a \<le> x \<and> x <e b} = {a..} \<inter> {x. x <e b}"
+ by auto
+
+lemma eucl_ivals[measurable]:
+ fixes a b :: "'a::ordered_euclidean_space"
+ shows "{x. x <e a} \<in> sets borel"
+ and "{x. a <e x} \<in> sets borel"
+ and "{..a} \<in> sets borel"
+ and "{a..} \<in> sets borel"
+ and "{a..b} \<in> sets borel"
+ and "{x. a <e x \<and> x \<le> b} \<in> sets borel"
+ and "{x. a \<le> x \<and> x <e b} \<in> sets borel"
+ unfolding box_oc box_co
+ by (auto intro: borel_open borel_closed)
+
+lemma
+ fixes i :: "'a::{second_countable_topology, real_inner}"
+ shows hafspace_less_borel: "{x. a < x \<bullet> i} \<in> sets borel"
+ and hafspace_greater_borel: "{x. x \<bullet> i < a} \<in> sets borel"
+ and hafspace_less_eq_borel: "{x. a \<le> x \<bullet> i} \<in> sets borel"
+ and hafspace_greater_eq_borel: "{x. x \<bullet> i \<le> a} \<in> sets borel"
+ by simp_all
+
+lemma borel_eq_box:
+ "borel = sigma UNIV (range (\<lambda> (a, b). box a b :: 'a :: euclidean_space set))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI1[OF borel_def])
+ fix M :: "'a set" assume "M \<in> {S. open S}"
+ then have "open M" by simp
+ show "M \<in> ?SIGMA"
+ apply (subst open_UNION_box[OF \<open>open M\<close>])
+ apply (safe intro!: sets.countable_UN' countable_PiE countable_Collect)
+ apply (auto intro: countable_rat)
+ done
+qed (auto simp: box_def)
+
+lemma halfspace_gt_in_halfspace:
+ assumes i: "i \<in> A"
+ shows "{x::'a. a < x \<bullet> i} \<in>
+ sigma_sets UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> A))"
+ (is "?set \<in> ?SIGMA")
+proof -
+ interpret sigma_algebra UNIV ?SIGMA
+ by (intro sigma_algebra_sigma_sets) simp_all
+ have *: "?set = (\<Union>n. UNIV - {x::'a. x \<bullet> i < a + 1 / real (Suc n)})"
+ proof (safe, simp_all add: not_less del: of_nat_Suc)
+ fix x :: 'a assume "a < x \<bullet> i"
+ with reals_Archimedean[of "x \<bullet> i - a"]
+ obtain n where "a + 1 / real (Suc n) < x \<bullet> i"
+ by (auto simp: field_simps)
+ then show "\<exists>n. a + 1 / real (Suc n) \<le> x \<bullet> i"
+ by (blast intro: less_imp_le)
+ next
+ fix x n
+ have "a < a + 1 / real (Suc n)" by auto
+ also assume "\<dots> \<le> x"
+ finally show "a < x" .
+ qed
+ show "?set \<in> ?SIGMA" unfolding *
+ by (auto intro!: Diff sigma_sets_Inter i)
+qed
+
+lemma borel_eq_halfspace_less:
+ "borel = sigma UNIV ((\<lambda>(a, i). {x::'a::euclidean_space. x \<bullet> i < a}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_box])
+ fix a b :: 'a
+ have "box a b = {x\<in>space ?SIGMA. \<forall>i\<in>Basis. a \<bullet> i < x \<bullet> i \<and> x \<bullet> i < b \<bullet> i}"
+ by (auto simp: box_def)
+ also have "\<dots> \<in> sets ?SIGMA"
+ by (intro sets.sets_Collect_conj sets.sets_Collect_finite_All sets.sets_Collect_const)
+ (auto intro!: halfspace_gt_in_halfspace countable_PiE countable_rat)
+ finally show "box a b \<in> sets ?SIGMA" .
+qed auto
+
+lemma borel_eq_halfspace_le:
+ "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. x \<bullet> i \<le> a}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have i: "i \<in> Basis" by auto
+ have *: "{x::'a. x\<bullet>i < a} = (\<Union>n. {x. x\<bullet>i \<le> a - 1/real (Suc n)})"
+ proof (safe, simp_all del: of_nat_Suc)
+ fix x::'a assume *: "x\<bullet>i < a"
+ with reals_Archimedean[of "a - x\<bullet>i"]
+ obtain n where "x \<bullet> i < a - 1 / (real (Suc n))"
+ by (auto simp: field_simps)
+ then show "\<exists>n. x \<bullet> i \<le> a - 1 / (real (Suc n))"
+ by (blast intro: less_imp_le)
+ next
+ fix x::'a and n
+ assume "x\<bullet>i \<le> a - 1 / real (Suc n)"
+ also have "\<dots> < a" by auto
+ finally show "x\<bullet>i < a" .
+ qed
+ show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
+ by (intro sets.countable_UN) (auto intro: i)
+qed auto
+
+lemma borel_eq_halfspace_ge:
+ "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a \<le> x \<bullet> i}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
+ fix a :: real and i :: 'a assume i: "(a, i) \<in> UNIV \<times> Basis"
+ have *: "{x::'a. x\<bullet>i < a} = space ?SIGMA - {x::'a. a \<le> x\<bullet>i}" by auto
+ show "{x. x\<bullet>i < a} \<in> ?SIGMA" unfolding *
+ using i by (intro sets.compl_sets) auto
+qed auto
+
+lemma borel_eq_halfspace_greater:
+ "borel = sigma UNIV ((\<lambda> (a, i). {x::'a::euclidean_space. a < x \<bullet> i}) ` (UNIV \<times> Basis))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
+ fix a :: real and i :: 'a assume "(a, i) \<in> (UNIV \<times> Basis)"
+ then have i: "i \<in> Basis" by auto
+ have *: "{x::'a. x\<bullet>i \<le> a} = space ?SIGMA - {x::'a. a < x\<bullet>i}" by auto
+ show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
+ by (intro sets.compl_sets) (auto intro: i)
+qed auto
+
+lemma borel_eq_atMost:
+ "borel = sigma UNIV (range (\<lambda>a. {..a::'a::ordered_euclidean_space}))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have "i \<in> Basis" by auto
+ then have *: "{x::'a. x\<bullet>i \<le> a} = (\<Union>k::nat. {.. (\<Sum>n\<in>Basis. (if n = i then a else real k)*\<^sub>R n)})"
+ proof (safe, simp_all add: eucl_le[where 'a='a] split: if_split_asm)
+ fix x :: 'a
+ from real_arch_simple[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"] guess k::nat ..
+ then have "\<And>i. i \<in> Basis \<Longrightarrow> x\<bullet>i \<le> real k"
+ by (subst (asm) Max_le_iff) auto
+ then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia \<le> real k"
+ by (auto intro!: exI[of _ k])
+ qed
+ show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA" unfolding *
+ by (intro sets.countable_UN) auto
+qed auto
+
+lemma borel_eq_greaterThan:
+ "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. a <e x}))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have i: "i \<in> Basis" by auto
+ have "{x::'a. x\<bullet>i \<le> a} = UNIV - {x::'a. a < x\<bullet>i}" by auto
+ also have *: "{x::'a. a < x\<bullet>i} =
+ (\<Union>k::nat. {x. (\<Sum>n\<in>Basis. (if n = i then a else -real k) *\<^sub>R n) <e x})" using i
+ proof (safe, simp_all add: eucl_less_def split: if_split_asm)
+ fix x :: 'a
+ from reals_Archimedean2[of "Max ((\<lambda>i. -x\<bullet>i)`Basis)"]
+ guess k::nat .. note k = this
+ { fix i :: 'a assume "i \<in> Basis"
+ then have "-x\<bullet>i < real k"
+ using k by (subst (asm) Max_less_iff) auto
+ then have "- real k < x\<bullet>i" by simp }
+ then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> -real k < x \<bullet> ia"
+ by (auto intro!: exI[of _ k])
+ qed
+ finally show "{x. x\<bullet>i \<le> a} \<in> ?SIGMA"
+ apply (simp only:)
+ apply (intro sets.countable_UN sets.Diff)
+ apply (auto intro: sigma_sets_top)
+ done
+qed auto
+
+lemma borel_eq_lessThan:
+ "borel = sigma UNIV (range (\<lambda>a::'a::ordered_euclidean_space. {x. x <e a}))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
+ fix a :: real and i :: 'a assume "(a, i) \<in> UNIV \<times> Basis"
+ then have i: "i \<in> Basis" by auto
+ have "{x::'a. a \<le> x\<bullet>i} = UNIV - {x::'a. x\<bullet>i < a}" by auto
+ also have *: "{x::'a. x\<bullet>i < a} = (\<Union>k::nat. {x. x <e (\<Sum>n\<in>Basis. (if n = i then a else real k) *\<^sub>R n)})" using \<open>i\<in> Basis\<close>
+ proof (safe, simp_all add: eucl_less_def split: if_split_asm)
+ fix x :: 'a
+ from reals_Archimedean2[of "Max ((\<lambda>i. x\<bullet>i)`Basis)"]
+ guess k::nat .. note k = this
+ { fix i :: 'a assume "i \<in> Basis"
+ then have "x\<bullet>i < real k"
+ using k by (subst (asm) Max_less_iff) auto
+ then have "x\<bullet>i < real k" by simp }
+ then show "\<exists>k::nat. \<forall>ia\<in>Basis. ia \<noteq> i \<longrightarrow> x \<bullet> ia < real k"
+ by (auto intro!: exI[of _ k])
+ qed
+ finally show "{x. a \<le> x\<bullet>i} \<in> ?SIGMA"
+ apply (simp only:)
+ apply (intro sets.countable_UN sets.Diff)
+ apply (auto intro: sigma_sets_top )
+ done
+qed auto
+
+lemma borel_eq_atLeastAtMost:
+ "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} ::'a::ordered_euclidean_space set))"
+ (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+ fix a::'a
+ have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
+ proof (safe, simp_all add: eucl_le[where 'a='a])
+ fix x :: 'a
+ from real_arch_simple[of "Max ((\<lambda>i. - x\<bullet>i)`Basis)"]
+ guess k::nat .. note k = this
+ { fix i :: 'a assume "i \<in> Basis"
+ with k have "- x\<bullet>i \<le> real k"
+ by (subst (asm) Max_le_iff) (auto simp: field_simps)
+ then have "- real k \<le> x\<bullet>i" by simp }
+ then show "\<exists>n::nat. \<forall>i\<in>Basis. - real n \<le> x \<bullet> i"
+ by (auto intro!: exI[of _ k])
+ qed
+ show "{..a} \<in> ?SIGMA" unfolding *
+ by (intro sets.countable_UN)
+ (auto intro!: sigma_sets_top)
+qed auto
+
+lemma borel_set_induct[consumes 1, case_names empty interval compl union]:
+ assumes "A \<in> sets borel"
+ assumes empty: "P {}" and int: "\<And>a b. a \<le> b \<Longrightarrow> P {a..b}" and compl: "\<And>A. A \<in> sets borel \<Longrightarrow> P A \<Longrightarrow> P (-A)" and
+ un: "\<And>f. disjoint_family f \<Longrightarrow> (\<And>i. f i \<in> sets borel) \<Longrightarrow> (\<And>i. P (f i)) \<Longrightarrow> P (\<Union>i::nat. f i)"
+ shows "P (A::real set)"
+proof-
+ let ?G = "range (\<lambda>(a,b). {a..b::real})"
+ have "Int_stable ?G" "?G \<subseteq> Pow UNIV" "A \<in> sigma_sets UNIV ?G"
+ using assms(1) by (auto simp add: borel_eq_atLeastAtMost Int_stable_def)
+ thus ?thesis
+ proof (induction rule: sigma_sets_induct_disjoint)
+ case (union f)
+ from union.hyps(2) have "\<And>i. f i \<in> sets borel" by (auto simp: borel_eq_atLeastAtMost)
+ with union show ?case by (auto intro: un)
+ next
+ case (basic A)
+ then obtain a b where "A = {a .. b}" by auto
+ then show ?case
+ by (cases "a \<le> b") (auto intro: int empty)
+ qed (auto intro: empty compl simp: Compl_eq_Diff_UNIV[symmetric] borel_eq_atLeastAtMost)
+qed
+
+lemma borel_sigma_sets_Ioc: "borel = sigma UNIV (range (\<lambda>(a, b). {a <.. b::real}))"
+proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+ fix i :: real
+ have "{..i} = (\<Union>j::nat. {-j <.. i})"
+ by (auto simp: minus_less_iff reals_Archimedean2)
+ also have "\<dots> \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))"
+ by (intro sets.countable_nat_UN) auto
+ finally show "{..i} \<in> sets (sigma UNIV (range (\<lambda>(i, j). {i<..j})))" .
+qed simp
+
+lemma eucl_lessThan: "{x::real. x <e a} = lessThan a"
+ by (simp add: eucl_less_def lessThan_def)
+
+lemma borel_eq_atLeastLessThan:
+ "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
+ have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
+ fix x :: real
+ have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
+ by (auto simp: move_uminus real_arch_simple)
+ then show "{y. y <e x} \<in> ?SIGMA"
+ by (auto intro: sigma_sets.intros(2-) simp: eucl_lessThan)
+qed auto
+
+lemma borel_measurable_halfspacesI:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ assumes F: "borel = sigma UNIV (F ` (UNIV \<times> Basis))"
+ and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
+ shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a::real. S a i \<in> sets M)"
+proof safe
+ fix a :: real and i :: 'b assume i: "i \<in> Basis" and f: "f \<in> borel_measurable M"
+ then show "S a i \<in> sets M" unfolding assms
+ by (auto intro!: measurable_sets simp: assms(1))
+next
+ assume a: "\<forall>i\<in>Basis. \<forall>a. S a i \<in> sets M"
+ then show "f \<in> borel_measurable M"
+ by (auto intro!: measurable_measure_of simp: S_eq F)
+qed
+
+lemma borel_measurable_iff_halfspace_le:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i \<le> a} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
+
+lemma borel_measurable_iff_halfspace_less:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. f w \<bullet> i < a} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
+
+lemma borel_measurable_iff_halfspace_ge:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ shows "f \<in> borel_measurable M = (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a \<le> f w \<bullet> i} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
+
+lemma borel_measurable_iff_halfspace_greater:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. \<forall>a. {w \<in> space M. a < f w \<bullet> i} \<in> sets M)"
+ by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
+
+lemma borel_measurable_iff_le:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
+ using borel_measurable_iff_halfspace_le[where 'c=real] by simp
+
+lemma borel_measurable_iff_less:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
+ using borel_measurable_iff_halfspace_less[where 'c=real] by simp
+
+lemma borel_measurable_iff_ge:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
+ using borel_measurable_iff_halfspace_ge[where 'c=real]
+ by simp
+
+lemma borel_measurable_iff_greater:
+ "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
+ using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
+
+lemma borel_measurable_euclidean_space:
+ fixes f :: "'a \<Rightarrow> 'c::euclidean_space"
+ shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M)"
+proof safe
+ assume f: "\<forall>i\<in>Basis. (\<lambda>x. f x \<bullet> i) \<in> borel_measurable M"
+ then show "f \<in> borel_measurable M"
+ by (subst borel_measurable_iff_halfspace_le) auto
+qed auto
+
+subsection "Borel measurable operators"
+
+lemma borel_measurable_norm[measurable]: "norm \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_sgn [measurable]: "(sgn::'a::real_normed_vector \<Rightarrow> 'a) \<in> borel_measurable borel"
+ by (rule borel_measurable_continuous_countable_exceptions[where X="{0}"])
+ (auto intro!: continuous_on_sgn continuous_on_id)
+
+lemma borel_measurable_uminus[measurable (raw)]:
+ fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. - g x) \<in> borel_measurable M"
+ by (rule borel_measurable_continuous_on[OF _ g]) (intro continuous_intros)
+
+lemma borel_measurable_diff[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+ using borel_measurable_add [of f M "- g"] assms by (simp add: fun_Compl_def)
+
+lemma borel_measurable_times[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_algebra}"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+ using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+
+lemma borel_measurable_setprod[measurable (raw)]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> 'b::{second_countable_topology, real_normed_field}"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
+proof cases
+ assume "finite S"
+ thus ?thesis using assms by induct auto
+qed simp
+
+lemma borel_measurable_dist[measurable (raw)]:
+ fixes g f :: "'a \<Rightarrow> 'b::{second_countable_topology, metric_space}"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. dist (f x) (g x)) \<in> borel_measurable M"
+ using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+
+lemma borel_measurable_scaleR[measurable (raw)]:
+ fixes g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_normed_vector}"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ shows "(\<lambda>x. f x *\<^sub>R g x) \<in> borel_measurable M"
+ using f g by (rule borel_measurable_continuous_Pair) (intro continuous_intros)
+
+lemma affine_borel_measurable_vector:
+ fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
+ assumes "f \<in> borel_measurable M"
+ shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
+proof (rule borel_measurableI)
+ fix S :: "'x set" assume "open S"
+ show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
+ proof cases
+ assume "b \<noteq> 0"
+ with \<open>open S\<close> have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
+ using open_affinity [of S "inverse b" "- a /\<^sub>R b"]
+ by (auto simp: algebra_simps)
+ hence "?S \<in> sets borel" by auto
+ moreover
+ from \<open>b \<noteq> 0\<close> have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
+ apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
+ ultimately show ?thesis using assms unfolding in_borel_measurable_borel
+ by auto
+ qed simp
+qed
+
+lemma borel_measurable_const_scaleR[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. b *\<^sub>R f x ::'a::real_normed_vector) \<in> borel_measurable M"
+ using affine_borel_measurable_vector[of f M 0 b] by simp
+
+lemma borel_measurable_const_add[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. a + f x ::'a::real_normed_vector) \<in> borel_measurable M"
+ using affine_borel_measurable_vector[of f M a 1] by simp
+
+lemma borel_measurable_inverse[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> 'b::real_normed_div_algebra"
+ assumes f: "f \<in> borel_measurable M"
+ shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
+ apply (rule measurable_compose[OF f])
+ apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"])
+ apply (auto intro!: continuous_on_inverse continuous_on_id)
+ done
+
+lemma borel_measurable_divide[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
+ (\<lambda>x. f x / g x::'b::{second_countable_topology, real_normed_div_algebra}) \<in> borel_measurable M"
+ by (simp add: divide_inverse)
+
+lemma borel_measurable_abs[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
+ unfolding abs_real_def by simp
+
+lemma borel_measurable_nth[measurable (raw)]:
+ "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
+ by (simp add: cart_eq_inner_axis)
+
+lemma convex_measurable:
+ fixes A :: "'a :: euclidean_space set"
+ shows "X \<in> borel_measurable M \<Longrightarrow> X ` space M \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> convex_on A q \<Longrightarrow>
+ (\<lambda>x. q (X x)) \<in> borel_measurable M"
+ by (rule measurable_compose[where f=X and N="restrict_space borel A"])
+ (auto intro!: borel_measurable_continuous_on_restrict convex_on_continuous measurable_restrict_space2)
+
+lemma borel_measurable_ln[measurable (raw)]:
+ assumes f: "f \<in> borel_measurable M"
+ shows "(\<lambda>x. ln (f x :: real)) \<in> borel_measurable M"
+ apply (rule measurable_compose[OF f])
+ apply (rule borel_measurable_continuous_countable_exceptions[of "{0}"])
+ apply (auto intro!: continuous_on_ln continuous_on_id)
+ done
+
+lemma borel_measurable_log[measurable (raw)]:
+ "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log (g x) (f x)) \<in> borel_measurable M"
+ unfolding log_def by auto
+
+lemma borel_measurable_exp[measurable]:
+ "(exp::'a::{real_normed_field,banach}\<Rightarrow>'a) \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_at_imp_continuous_on ballI isCont_exp)
+
+lemma measurable_real_floor[measurable]:
+ "(floor :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
+proof -
+ have "\<And>a x. \<lfloor>x\<rfloor> = a \<longleftrightarrow> (real_of_int a \<le> x \<and> x < real_of_int (a + 1))"
+ by (auto intro: floor_eq2)
+ then show ?thesis
+ by (auto simp: vimage_def measurable_count_space_eq2_countable)
+qed
+
+lemma measurable_real_ceiling[measurable]:
+ "(ceiling :: real \<Rightarrow> int) \<in> measurable borel (count_space UNIV)"
+ unfolding ceiling_def[abs_def] by simp
+
+lemma borel_measurable_real_floor: "(\<lambda>x::real. real_of_int \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
+ by simp
+
+lemma borel_measurable_root [measurable]: "root n \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_sqrt [measurable]: "sqrt \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_power [measurable (raw)]:
+ fixes f :: "_ \<Rightarrow> 'b::{power,real_normed_algebra}"
+ assumes f: "f \<in> borel_measurable M"
+ shows "(\<lambda>x. (f x) ^ n) \<in> borel_measurable M"
+ by (intro borel_measurable_continuous_on [OF _ f] continuous_intros)
+
+lemma borel_measurable_Re [measurable]: "Re \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_Im [measurable]: "Im \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_of_real [measurable]: "(of_real :: _ \<Rightarrow> (_::real_normed_algebra)) \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_sin [measurable]: "(sin :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_cos [measurable]: "(cos :: _ \<Rightarrow> (_::{real_normed_field,banach})) \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_arctan [measurable]: "arctan \<in> borel_measurable borel"
+ by (intro borel_measurable_continuous_on1 continuous_intros)
+
+lemma borel_measurable_complex_iff:
+ "f \<in> borel_measurable M \<longleftrightarrow>
+ (\<lambda>x. Re (f x)) \<in> borel_measurable M \<and> (\<lambda>x. Im (f x)) \<in> borel_measurable M"
+ apply auto
+ apply (subst fun_complex_eq)
+ apply (intro borel_measurable_add)
+ apply auto
+ done
+
+subsection "Borel space on the extended reals"
+
+lemma borel_measurable_ereal[measurable (raw)]:
+ assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
+ using continuous_on_ereal f by (rule borel_measurable_continuous_on) (rule continuous_on_id)
+
+lemma borel_measurable_real_of_ereal[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> ereal"
+ assumes f: "f \<in> borel_measurable M"
+ shows "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M"
+ apply (rule measurable_compose[OF f])
+ apply (rule borel_measurable_continuous_countable_exceptions[of "{\<infinity>, -\<infinity> }"])
+ apply (auto intro: continuous_on_real simp: Compl_eq_Diff_UNIV)
+ done
+
+lemma borel_measurable_ereal_cases:
+ fixes f :: "'a \<Rightarrow> ereal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x)))) \<in> borel_measurable M"
+ shows "(\<lambda>x. H (f x)) \<in> borel_measurable M"
+proof -
+ let ?F = "\<lambda>x. if f x = \<infinity> then H \<infinity> else if f x = - \<infinity> then H (-\<infinity>) else H (ereal (real_of_ereal (f x)))"
+ { fix x have "H (f x) = ?F x" by (cases "f x") auto }
+ with f H show ?thesis by simp
+qed
+
+lemma
+ fixes f :: "'a \<Rightarrow> ereal" assumes f[measurable]: "f \<in> borel_measurable M"
+ shows borel_measurable_ereal_abs[measurable(raw)]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
+ and borel_measurable_ereal_inverse[measurable(raw)]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
+ and borel_measurable_uminus_ereal[measurable(raw)]: "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
+ by (auto simp del: abs_real_of_ereal simp: borel_measurable_ereal_cases[OF f] measurable_If)
+
+lemma borel_measurable_uminus_eq_ereal[simp]:
+ "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
+proof
+ assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
+qed auto
+
+lemma set_Collect_ereal2:
+ fixes f g :: "'a \<Rightarrow> ereal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ assumes H: "{x \<in> space M. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))} \<in> sets M"
+ "{x \<in> space borel. H (-\<infinity>) (ereal x)} \<in> sets borel"
+ "{x \<in> space borel. H (\<infinity>) (ereal x)} \<in> sets borel"
+ "{x \<in> space borel. H (ereal x) (-\<infinity>)} \<in> sets borel"
+ "{x \<in> space borel. H (ereal x) (\<infinity>)} \<in> sets borel"
+ shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
+proof -
+ let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = -\<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))"
+ let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = -\<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x"
+ { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
+ note * = this
+ from assms show ?thesis
+ by (subst *) (simp del: space_borel split del: if_split)
+qed
+
+lemma borel_measurable_ereal_iff:
+ shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+proof
+ assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
+ from borel_measurable_real_of_ereal[OF this]
+ show "f \<in> borel_measurable M" by auto
+qed auto
+
+lemma borel_measurable_erealD[measurable_dest]:
+ "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<Longrightarrow> g \<in> measurable N M \<Longrightarrow> (\<lambda>x. f (g x)) \<in> borel_measurable N"
+ unfolding borel_measurable_ereal_iff by simp
+
+lemma borel_measurable_ereal_iff_real:
+ fixes f :: "'a \<Rightarrow> ereal"
+ shows "f \<in> borel_measurable M \<longleftrightarrow>
+ ((\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
+proof safe
+ assume *: "(\<lambda>x. real_of_ereal (f x)) \<in> borel_measurable M" "f -` {\<infinity>} \<inter> space M \<in> sets M" "f -` {-\<infinity>} \<inter> space M \<in> sets M"
+ have "f -` {\<infinity>} \<inter> space M = {x\<in>space M. f x = \<infinity>}" "f -` {-\<infinity>} \<inter> space M = {x\<in>space M. f x = -\<infinity>}" by auto
+ with * have **: "{x\<in>space M. f x = \<infinity>} \<in> sets M" "{x\<in>space M. f x = -\<infinity>} \<in> sets M" by simp_all
+ let ?f = "\<lambda>x. if f x = \<infinity> then \<infinity> else if f x = -\<infinity> then -\<infinity> else ereal (real_of_ereal (f x))"
+ have "?f \<in> borel_measurable M" using * ** by (intro measurable_If) auto
+ also have "?f = f" by (auto simp: fun_eq_iff ereal_real)
+ finally show "f \<in> borel_measurable M" .
+qed simp_all
+
+lemma borel_measurable_ereal_iff_Iio:
+ "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
+ by (auto simp: borel_Iio measurable_iff_measure_of)
+
+lemma borel_measurable_ereal_iff_Ioi:
+ "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
+ by (auto simp: borel_Ioi measurable_iff_measure_of)
+
+lemma vimage_sets_compl_iff:
+ "f -` A \<inter> space M \<in> sets M \<longleftrightarrow> f -` (- A) \<inter> space M \<in> sets M"
+proof -
+ { fix A assume "f -` A \<inter> space M \<in> sets M"
+ moreover have "f -` (- A) \<inter> space M = space M - f -` A \<inter> space M" by auto
+ ultimately have "f -` (- A) \<inter> space M \<in> sets M" by auto }
+ from this[of A] this[of "-A"] show ?thesis
+ by (metis double_complement)
+qed
+
+lemma borel_measurable_iff_Iic_ereal:
+ "(f::'a\<Rightarrow>ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
+ unfolding borel_measurable_ereal_iff_Ioi vimage_sets_compl_iff[where A="{a <..}" for a] by simp
+
+lemma borel_measurable_iff_Ici_ereal:
+ "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
+ unfolding borel_measurable_ereal_iff_Iio vimage_sets_compl_iff[where A="{..< a}" for a] by simp
+
+lemma borel_measurable_ereal2:
+ fixes f g :: "'a \<Rightarrow> ereal"
+ assumes f: "f \<in> borel_measurable M"
+ assumes g: "g \<in> borel_measurable M"
+ assumes H: "(\<lambda>x. H (ereal (real_of_ereal (f x))) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M"
+ "(\<lambda>x. H (-\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M"
+ "(\<lambda>x. H (\<infinity>) (ereal (real_of_ereal (g x)))) \<in> borel_measurable M"
+ "(\<lambda>x. H (ereal (real_of_ereal (f x))) (-\<infinity>)) \<in> borel_measurable M"
+ "(\<lambda>x. H (ereal (real_of_ereal (f x))) (\<infinity>)) \<in> borel_measurable M"
+ shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
+proof -
+ let ?G = "\<lambda>y x. if g x = \<infinity> then H y \<infinity> else if g x = - \<infinity> then H y (-\<infinity>) else H y (ereal (real_of_ereal (g x)))"
+ let ?F = "\<lambda>x. if f x = \<infinity> then ?G \<infinity> x else if f x = - \<infinity> then ?G (-\<infinity>) x else ?G (ereal (real_of_ereal (f x))) x"
+ { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
+ note * = this
+ from assms show ?thesis unfolding * by simp
+qed
+
+lemma [measurable(raw)]:
+ fixes f :: "'a \<Rightarrow> ereal"
+ assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+ shows borel_measurable_ereal_add: "(\<lambda>x. f x + g x) \<in> borel_measurable M"
+ and borel_measurable_ereal_times: "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+ by (simp_all add: borel_measurable_ereal2)
+
+lemma [measurable(raw)]:
+ fixes f g :: "'a \<Rightarrow> ereal"
+ assumes "f \<in> borel_measurable M"
+ assumes "g \<in> borel_measurable M"
+ shows borel_measurable_ereal_diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+ and borel_measurable_ereal_divide: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
+ using assms by (simp_all add: minus_ereal_def divide_ereal_def)
+
+lemma borel_measurable_ereal_setsum[measurable (raw)]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
+ using assms by (induction S rule: infinite_finite_induct) auto
+
+lemma borel_measurable_ereal_setprod[measurable (raw)]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
+ using assms by (induction S rule: infinite_finite_induct) auto
+
+lemma borel_measurable_extreal_suminf[measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
+ assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
+ unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
+
+subsection "Borel space on the extended non-negative reals"
+
+text \<open> @{type ennreal} is a topological monoid, so no rules for plus are required, also all order
+ statements are usually done on type classes. \<close>
+
+lemma measurable_enn2ereal[measurable]: "enn2ereal \<in> borel \<rightarrow>\<^sub>M borel"
+ by (intro borel_measurable_continuous_on1 continuous_on_enn2ereal)
+
+lemma measurable_e2ennreal[measurable]: "e2ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+ by (intro borel_measurable_continuous_on1 continuous_on_e2ennreal)
+
+lemma borel_measurable_enn2real[measurable (raw)]:
+ "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. enn2real (f x)) \<in> M \<rightarrow>\<^sub>M borel"
+ unfolding enn2real_def[abs_def] by measurable
+
+definition [simp]: "is_borel f M \<longleftrightarrow> f \<in> borel_measurable M"
+
+lemma is_borel_transfer[transfer_rule]: "rel_fun (rel_fun op = pcr_ennreal) op = is_borel is_borel"
+ unfolding is_borel_def[abs_def]
+proof (safe intro!: rel_funI ext dest!: rel_fun_eq_pcr_ennreal[THEN iffD1])
+ fix f and M :: "'a measure"
+ show "f \<in> borel_measurable M" if f: "enn2ereal \<circ> f \<in> borel_measurable M"
+ using measurable_compose[OF f measurable_e2ennreal] by simp
+qed simp
+
+context
+ includes ennreal.lifting
+begin
+
+lemma measurable_ennreal[measurable]: "ennreal \<in> borel \<rightarrow>\<^sub>M borel"
+ unfolding is_borel_def[symmetric]
+ by transfer simp
+
+lemma borel_measurable_ennreal_iff[simp]:
+ assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+ shows "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel \<longleftrightarrow> f \<in> M \<rightarrow>\<^sub>M borel"
+proof safe
+ assume "(\<lambda>x. ennreal (f x)) \<in> M \<rightarrow>\<^sub>M borel"
+ then have "(\<lambda>x. enn2real (ennreal (f x))) \<in> M \<rightarrow>\<^sub>M borel"
+ by measurable
+ then show "f \<in> M \<rightarrow>\<^sub>M borel"
+ by (rule measurable_cong[THEN iffD1, rotated]) auto
+qed measurable
+
+lemma borel_measurable_times_ennreal[measurable (raw)]:
+ fixes f g :: "'a \<Rightarrow> ennreal"
+ shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x * g x) \<in> M \<rightarrow>\<^sub>M borel"
+ unfolding is_borel_def[symmetric] by transfer simp
+
+lemma borel_measurable_inverse_ennreal[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> ennreal"
+ shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. inverse (f x)) \<in> M \<rightarrow>\<^sub>M borel"
+ unfolding is_borel_def[symmetric] by transfer simp
+
+lemma borel_measurable_divide_ennreal[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> ennreal"
+ shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x / g x) \<in> M \<rightarrow>\<^sub>M borel"
+ unfolding divide_ennreal_def by simp
+
+lemma borel_measurable_minus_ennreal[measurable (raw)]:
+ fixes f :: "'a \<Rightarrow> ennreal"
+ shows "f \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> g \<in> M \<rightarrow>\<^sub>M borel \<Longrightarrow> (\<lambda>x. f x - g x) \<in> M \<rightarrow>\<^sub>M borel"
+ unfolding is_borel_def[symmetric] by transfer simp
+
+lemma borel_measurable_setprod_ennreal[measurable (raw)]:
+ fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ennreal"
+ assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
+ shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
+ using assms by (induction S rule: infinite_finite_induct) auto
+
+end
+
+hide_const (open) is_borel
+
+subsection \<open>LIMSEQ is borel measurable\<close>
+
+lemma borel_measurable_LIMSEQ_real:
+ fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+ assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) \<longlonglongrightarrow> u' x"
+ and u: "\<And>i. u i \<in> borel_measurable M"
+ shows "u' \<in> borel_measurable M"
+proof -
+ have "\<And>x. x \<in> space M \<Longrightarrow> liminf (\<lambda>n. ereal (u n x)) = ereal (u' x)"
+ using u' by (simp add: lim_imp_Liminf)
+ moreover from u have "(\<lambda>x. liminf (\<lambda>n. ereal (u n x))) \<in> borel_measurable M"
+ by auto
+ ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
+qed
+
+lemma borel_measurable_LIMSEQ_metric:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: metric_space"
+ assumes [measurable]: "\<And>i. f i \<in> borel_measurable M"
+ assumes lim: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. f i x) \<longlonglongrightarrow> g x"
+ shows "g \<in> borel_measurable M"
+ unfolding borel_eq_closed
+proof (safe intro!: measurable_measure_of)
+ fix A :: "'b set" assume "closed A"
+
+ have [measurable]: "(\<lambda>x. infdist (g x) A) \<in> borel_measurable M"
+ proof (rule borel_measurable_LIMSEQ_real)
+ show "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. infdist (f i x) A) \<longlonglongrightarrow> infdist (g x) A"
+ by (intro tendsto_infdist lim)
+ show "\<And>i. (\<lambda>x. infdist (f i x) A) \<in> borel_measurable M"
+ by (intro borel_measurable_continuous_on[where f="\<lambda>x. infdist x A"]
+ continuous_at_imp_continuous_on ballI continuous_infdist continuous_ident) auto
+ qed
+
+ show "g -` A \<inter> space M \<in> sets M"
+ proof cases
+ assume "A \<noteq> {}"
+ then have "\<And>x. infdist x A = 0 \<longleftrightarrow> x \<in> A"
+ using \<open>closed A\<close> by (simp add: in_closed_iff_infdist_zero)
+ then have "g -` A \<inter> space M = {x\<in>space M. infdist (g x) A = 0}"
+ by auto
+ also have "\<dots> \<in> sets M"
+ by measurable
+ finally show ?thesis .
+ qed simp
+qed auto
+
+lemma sets_Collect_Cauchy[measurable]:
+ fixes f :: "nat \<Rightarrow> 'a => 'b::{metric_space, second_countable_topology}"
+ assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
+ unfolding metric_Cauchy_iff2 using f by auto
+
+lemma borel_measurable_lim_metric[measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. lim (\<lambda>i. f i x)) \<in> borel_measurable M"
+proof -
+ define u' where "u' x = lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)" for x
+ then have *: "\<And>x. lim (\<lambda>i. f i x) = (if Cauchy (\<lambda>i. f i x) then u' x else (THE x. False))"
+ by (auto simp: lim_def convergent_eq_cauchy[symmetric])
+ have "u' \<in> borel_measurable M"
+ proof (rule borel_measurable_LIMSEQ_metric)
+ fix x
+ have "convergent (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0)"
+ by (cases "Cauchy (\<lambda>i. f i x)")
+ (auto simp add: convergent_eq_cauchy[symmetric] convergent_def)
+ then show "(\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) \<longlonglongrightarrow> u' x"
+ unfolding u'_def
+ by (rule convergent_LIMSEQ_iff[THEN iffD1])
+ qed measurable
+ then show ?thesis
+ unfolding * by measurable
+qed
+
+lemma borel_measurable_suminf[measurable (raw)]:
+ fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+ assumes f[measurable]: "\<And>i. f i \<in> borel_measurable M"
+ shows "(\<lambda>x. suminf (\<lambda>i. f i x)) \<in> borel_measurable M"
+ unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
+
+lemma Collect_closed_imp_pred_borel: "closed {x. P x} \<Longrightarrow> Measurable.pred borel P"
+ by (simp add: pred_def)
+
+(* Proof by Jeremy Avigad and Luke Serafin *)
+lemma isCont_borel_pred[measurable]:
+ fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
+ shows "Measurable.pred borel (isCont f)"
+proof (subst measurable_cong)
+ let ?I = "\<lambda>j. inverse(real (Suc j))"
+ show "isCont f x = (\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i)" for x
+ unfolding continuous_at_eps_delta
+ proof safe
+ fix i assume "\<forall>e>0. \<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e"
+ moreover have "0 < ?I i / 2"
+ by simp
+ ultimately obtain d where d: "0 < d" "\<And>y. dist x y < d \<Longrightarrow> dist (f y) (f x) < ?I i / 2"
+ by (metis dist_commute)
+ then obtain j where j: "?I j < d"
+ by (metis reals_Archimedean)
+
+ show "\<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i"
+ proof (safe intro!: exI[where x=j])
+ fix y z assume *: "dist x y < ?I j" "dist x z < ?I j"
+ have "dist (f y) (f z) \<le> dist (f y) (f x) + dist (f z) (f x)"
+ by (rule dist_triangle2)
+ also have "\<dots> < ?I i / 2 + ?I i / 2"
+ by (intro add_strict_mono d less_trans[OF _ j] *)
+ also have "\<dots> \<le> ?I i"
+ by (simp add: field_simps of_nat_Suc)
+ finally show "dist (f y) (f z) \<le> ?I i"
+ by simp
+ qed
+ next
+ fix e::real assume "0 < e"
+ then obtain n where n: "?I n < e"
+ by (metis reals_Archimedean)
+ assume "\<forall>i. \<exists>j. \<forall>y z. dist x y < ?I j \<and> dist x z < ?I j \<longrightarrow> dist (f y) (f z) \<le> ?I i"
+ from this[THEN spec, of "Suc n"]
+ obtain j where j: "\<And>y z. dist x y < ?I j \<Longrightarrow> dist x z < ?I j \<Longrightarrow> dist (f y) (f z) \<le> ?I (Suc n)"
+ by auto
+
+ show "\<exists>d>0. \<forall>y. dist y x < d \<longrightarrow> dist (f y) (f x) < e"
+ proof (safe intro!: exI[of _ "?I j"])
+ fix y assume "dist y x < ?I j"
+ then have "dist (f y) (f x) \<le> ?I (Suc n)"
+ by (intro j) (auto simp: dist_commute)
+ also have "?I (Suc n) < ?I n"
+ by simp
+ also note n
+ finally show "dist (f y) (f x) < e" .
+ qed simp
+ qed
+qed (intro pred_intros_countable closed_Collect_all closed_Collect_le open_Collect_less
+ Collect_closed_imp_pred_borel closed_Collect_imp open_Collect_conj continuous_intros)
+
+lemma isCont_borel:
+ fixes f :: "'b::metric_space \<Rightarrow> 'a::metric_space"
+ shows "{x. isCont f x} \<in> sets borel"
+ by simp
+
+lemma is_real_interval:
+ assumes S: "is_interval S"
+ shows "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or> S = {a<..} \<or> S = {a..} \<or>
+ S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}"
+ using S unfolding is_interval_1 by (blast intro: interval_cases)
+
+lemma real_interval_borel_measurable:
+ assumes "is_interval (S::real set)"
+ shows "S \<in> sets borel"
+proof -
+ from assms is_real_interval have "\<exists>a b::real. S = {} \<or> S = UNIV \<or> S = {..<b} \<or> S = {..b} \<or>
+ S = {a<..} \<or> S = {a..} \<or> S = {a<..<b} \<or> S = {a<..b} \<or> S = {a..<b} \<or> S = {a..b}" by auto
+ then guess a ..
+ then guess b ..
+ thus ?thesis
+ by auto
+qed
+
+lemma borel_measurable_mono_on_fnc:
+ fixes f :: "real \<Rightarrow> real" and A :: "real set"
+ assumes "mono_on f A"
+ shows "f \<in> borel_measurable (restrict_space borel A)"
+ apply (rule measurable_restrict_countable[OF mono_on_ctble_discont[OF assms]])
+ apply (auto intro!: image_eqI[where x="{x}" for x] simp: sets_restrict_space)
+ apply (auto simp add: sets_restrict_restrict_space continuous_on_eq_continuous_within
+ cong: measurable_cong_sets
+ intro!: borel_measurable_continuous_on_restrict intro: continuous_within_subset)
+ done
+
+lemma borel_measurable_mono:
+ fixes f :: "real \<Rightarrow> real"
+ shows "mono f \<Longrightarrow> f \<in> borel_measurable borel"
+ using borel_measurable_mono_on_fnc[of f UNIV] by (simp add: mono_def mono_on_def)
+
+no_notation
+ eucl_less (infix "<e" 50)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Caratheodory.thy Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,891 @@
+(* Title: HOL/Probability/Caratheodory.thy
+ Author: Lawrence C Paulson
+ Author: Johannes Hölzl, TU München
+*)
+
+section \<open>Caratheodory Extension Theorem\<close>
+
+theory Caratheodory
+ imports Measure_Space
+begin
+
+text \<open>
+ Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
+\<close>
+
+lemma suminf_ennreal_2dimen:
+ fixes f:: "nat \<times> nat \<Rightarrow> ennreal"
+ assumes "\<And>m. g m = (\<Sum>n. f (m,n))"
+ shows "(\<Sum>i. f (prod_decode i)) = suminf g"
+proof -
+ have g_def: "g = (\<lambda>m. (\<Sum>n. f (m,n)))"
+ using assms by (simp add: fun_eq_iff)
+ have reindex: "\<And>B. (\<Sum>x\<in>B. f (prod_decode x)) = setsum f (prod_decode ` B)"
+ by (simp add: setsum.reindex[OF inj_prod_decode] comp_def)
+ have "(SUP n. \<Sum>i<n. f (prod_decode i)) = (SUP p : UNIV \<times> UNIV. \<Sum>i<fst p. \<Sum>n<snd p. f (i, n))"
+ proof (intro SUP_eq; clarsimp simp: setsum.cartesian_product reindex)
+ fix n
+ let ?M = "\<lambda>f. Suc (Max (f ` prod_decode ` {..<n}))"
+ { fix a b x assume "x < n" and [symmetric]: "(a, b) = prod_decode x"
+ then have "a < ?M fst" "b < ?M snd"
+ by (auto intro!: Max_ge le_imp_less_Suc image_eqI) }
+ then have "setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<?M fst} \<times> {..<?M snd})"
+ by (auto intro!: setsum_mono3)
+ then show "\<exists>a b. setsum f (prod_decode ` {..<n}) \<le> setsum f ({..<a} \<times> {..<b})" by auto
+ next
+ fix a b
+ let ?M = "prod_decode ` {..<Suc (Max (prod_encode ` ({..<a} \<times> {..<b})))}"
+ { fix a' b' assume "a' < a" "b' < b" then have "(a', b') \<in> ?M"
+ by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) }
+ then have "setsum f ({..<a} \<times> {..<b}) \<le> setsum f ?M"
+ by (auto intro!: setsum_mono3)
+ then show "\<exists>n. setsum f ({..<a} \<times> {..<b}) \<le> setsum f (prod_decode ` {..<n})"
+ by auto
+ qed
+ also have "\<dots> = (SUP p. \<Sum>i<p. \<Sum>n. f (i, n))"
+ unfolding suminf_setsum[OF summableI, symmetric]
+ by (simp add: suminf_eq_SUP SUP_pair setsum.commute[of _ "{..< fst _}"])
+ finally show ?thesis unfolding g_def
+ by (simp add: suminf_eq_SUP)
+qed
+
+subsection \<open>Characterizations of Measures\<close>
+
+definition outer_measure_space where
+ "outer_measure_space M f \<longleftrightarrow> positive M f \<and> increasing M f \<and> countably_subadditive M f"
+
+subsubsection \<open>Lambda Systems\<close>
+
+definition lambda_system :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set set"
+where
+ "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
+
+lemma (in algebra) lambda_system_eq:
+ "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}"
+proof -
+ have [simp]: "\<And>l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l"
+ by (metis Int_Diff Int_absorb1 Int_commute sets_into_space)
+ show ?thesis
+ by (auto simp add: lambda_system_def) (metis Int_commute)+
+qed
+
+lemma (in algebra) lambda_system_empty: "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
+ by (auto simp add: positive_def lambda_system_eq)
+
+lemma lambda_system_sets: "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
+ by (simp add: lambda_system_def)
+
+lemma (in algebra) lambda_system_Compl:
+ fixes f:: "'a set \<Rightarrow> ennreal"
+ assumes x: "x \<in> lambda_system \<Omega> M f"
+ shows "\<Omega> - x \<in> lambda_system \<Omega> M f"
+proof -
+ have "x \<subseteq> \<Omega>"
+ by (metis sets_into_space lambda_system_sets x)
+ hence "\<Omega> - (\<Omega> - x) = x"
+ by (metis double_diff equalityE)
+ with x show ?thesis
+ by (force simp add: lambda_system_def ac_simps)
+qed
+
+lemma (in algebra) lambda_system_Int:
+ fixes f:: "'a set \<Rightarrow> ennreal"
+ assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+ shows "x \<inter> y \<in> lambda_system \<Omega> M f"
+proof -
+ from xl yl show ?thesis
+ proof (auto simp add: positive_def lambda_system_eq Int)
+ fix u
+ assume x: "x \<in> M" and y: "y \<in> M" and u: "u \<in> M"
+ and fx: "\<forall>z\<in>M. f (z \<inter> x) + f (z - x) = f z"
+ and fy: "\<forall>z\<in>M. f (z \<inter> y) + f (z - y) = f z"
+ have "u - x \<inter> y \<in> M"
+ by (metis Diff Diff_Int Un u x y)
+ moreover
+ have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast
+ moreover
+ have "u - x \<inter> y - y = u - y" by blast
+ ultimately
+ have ey: "f (u - x \<inter> y) = f (u \<inter> y - x) + f (u - y)" using fy
+ by force
+ have "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y)
+ = (f (u \<inter> (x \<inter> y)) + f (u \<inter> y - x)) + f (u - y)"
+ by (simp add: ey ac_simps)
+ also have "... = (f ((u \<inter> y) \<inter> x) + f (u \<inter> y - x)) + f (u - y)"
+ by (simp add: Int_ac)
+ also have "... = f (u \<inter> y) + f (u - y)"
+ using fx [THEN bspec, of "u \<inter> y"] Int y u
+ by force
+ also have "... = f u"
+ by (metis fy u)
+ finally show "f (u \<inter> (x \<inter> y)) + f (u - x \<inter> y) = f u" .
+ qed
+qed
+
+lemma (in algebra) lambda_system_Un:
+ fixes f:: "'a set \<Rightarrow> ennreal"
+ assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+ shows "x \<union> y \<in> lambda_system \<Omega> M f"
+proof -
+ have "(\<Omega> - x) \<inter> (\<Omega> - y) \<in> M"
+ by (metis Diff_Un Un compl_sets lambda_system_sets xl yl)
+ moreover
+ have "x \<union> y = \<Omega> - ((\<Omega> - x) \<inter> (\<Omega> - y))"
+ by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+
+ ultimately show ?thesis
+ by (metis lambda_system_Compl lambda_system_Int xl yl)
+qed
+
+lemma (in algebra) lambda_system_algebra:
+ "positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)"
+ apply (auto simp add: algebra_iff_Un)
+ apply (metis lambda_system_sets set_mp sets_into_space)
+ apply (metis lambda_system_empty)
+ apply (metis lambda_system_Compl)
+ apply (metis lambda_system_Un)
+ done
+
+lemma (in algebra) lambda_system_strong_additive:
+ assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
+ and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+ shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
+proof -
+ have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
+ moreover
+ have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
+ moreover
+ have "(z \<inter> (x \<union> y)) \<in> M"
+ by (metis Int Un lambda_system_sets xl yl z)
+ ultimately show ?thesis using xl yl
+ by (simp add: lambda_system_eq)
+qed
+
+lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
+proof (auto simp add: additive_def)
+ fix x and y
+ assume disj: "x \<inter> y = {}"
+ and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+ hence "x \<in> M" "y \<in> M" by (blast intro: lambda_system_sets)+
+ thus "f (x \<union> y) = f x + f y"
+ using lambda_system_strong_additive [OF top disj xl yl]
+ by (simp add: Un)
+qed
+
+lemma lambda_system_increasing: "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
+ by (simp add: increasing_def lambda_system_def)
+
+lemma lambda_system_positive: "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
+ by (simp add: positive_def lambda_system_def)
+
+lemma (in algebra) lambda_system_strong_sum:
+ fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ennreal"
+ assumes f: "positive M f" and a: "a \<in> M"
+ and A: "range A \<subseteq> lambda_system \<Omega> M f"
+ and disj: "disjoint_family A"
+ shows "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))"
+proof (induct n)
+ case 0 show ?case using f by (simp add: positive_def)
+next
+ case (Suc n)
+ have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
+ by (force simp add: disjoint_family_on_def neq_iff)
+ have 3: "A n \<in> lambda_system \<Omega> M f" using A
+ by blast
+ interpret l: algebra \<Omega> "lambda_system \<Omega> M f"
+ using f by (rule lambda_system_algebra)
+ have 4: "UNION {0..<n} A \<in> lambda_system \<Omega> M f"
+ using A l.UNION_in_sets by simp
+ from Suc.hyps show ?case
+ by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
+qed
+
+lemma (in sigma_algebra) lambda_system_caratheodory:
+ assumes oms: "outer_measure_space M f"
+ and A: "range A \<subseteq> lambda_system \<Omega> M f"
+ and disj: "disjoint_family A"
+ shows "(\<Union>i. A i) \<in> lambda_system \<Omega> M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
+proof -
+ have pos: "positive M f" and inc: "increasing M f"
+ and csa: "countably_subadditive M f"
+ by (metis oms outer_measure_space_def)+
+ have sa: "subadditive M f"
+ by (metis countably_subadditive_subadditive csa pos)
+ have A': "\<And>S. A`S \<subseteq> (lambda_system \<Omega> M f)" using A
+ by auto
+ interpret ls: algebra \<Omega> "lambda_system \<Omega> M f"
+ using pos by (rule lambda_system_algebra)
+ have A'': "range A \<subseteq> M"
+ by (metis A image_subset_iff lambda_system_sets)
+
+ have U_in: "(\<Union>i. A i) \<in> M"
+ by (metis A'' countable_UN)
+ have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (A i))"
+ proof (rule antisym)
+ show "f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))"
+ using csa[unfolded countably_subadditive_def] A'' disj U_in by auto
+ have dis: "\<And>N. disjoint_family_on A {..<N}" by (intro disjoint_family_on_mono[OF _ disj]) auto
+ show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
+ using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis] A''
+ by (intro suminf_le_const[OF summableI]) (auto intro!: increasingD[OF inc] countable_UN)
+ qed
+ have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a"
+ if a [iff]: "a \<in> M" for a
+ proof (rule antisym)
+ have "range (\<lambda>i. a \<inter> A i) \<subseteq> M" using A''
+ by blast
+ moreover
+ have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
+ by (auto simp add: disjoint_family_on_def)
+ moreover
+ have "a \<inter> (\<Union>i. A i) \<in> M"
+ by (metis Int U_in a)
+ ultimately
+ have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))"
+ using csa[unfolded countably_subadditive_def, rule_format, of "(\<lambda>i. a \<inter> A i)"]
+ by (simp add: o_def)
+ hence "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i))"
+ by (rule add_right_mono)
+ also have "\<dots> \<le> f a"
+ proof (intro ennreal_suminf_bound_add)
+ fix n
+ have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M"
+ by (metis A'' UNION_in_sets)
+ have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
+ by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
+ have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system \<Omega> M f"
+ using ls.UNION_in_sets by (simp add: A)
+ hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))"
+ by (simp add: lambda_system_eq UNION_in)
+ have "f (a - (\<Union>i. A i)) \<le> f (a - (\<Union>i\<in>{0..<n}. A i))"
+ by (blast intro: increasingD [OF inc] UNION_in U_in)
+ thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
+ by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
+ qed
+ finally show "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) \<le> f a"
+ by simp
+ next
+ have "f a \<le> f (a \<inter> (\<Union>i. A i) \<union> (a - (\<Union>i. A i)))"
+ by (blast intro: increasingD [OF inc] U_in)
+ also have "... \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))"
+ by (blast intro: subadditiveD [OF sa] U_in)
+ finally show "f a \<le> f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i))" .
+ qed
+ thus ?thesis
+ by (simp add: lambda_system_eq sums_iff U_eq U_in)
+qed
+
+lemma (in sigma_algebra) caratheodory_lemma:
+ assumes oms: "outer_measure_space M f"
+ defines "L \<equiv> lambda_system \<Omega> M f"
+ shows "measure_space \<Omega> L f"
+proof -
+ have pos: "positive M f"
+ by (metis oms outer_measure_space_def)
+ have alg: "algebra \<Omega> L"
+ using lambda_system_algebra [of f, OF pos]
+ by (simp add: algebra_iff_Un L_def)
+ then
+ have "sigma_algebra \<Omega> L"
+ using lambda_system_caratheodory [OF oms]
+ by (simp add: sigma_algebra_disjoint_iff L_def)
+ moreover
+ have "countably_additive L f" "positive L f"
+ using pos lambda_system_caratheodory [OF oms]
+ by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def)
+ ultimately
+ show ?thesis
+ using pos by (simp add: measure_space_def)
+qed
+
+definition outer_measure :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a set \<Rightarrow> ennreal" where
+ "outer_measure M f X =
+ (INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i)}. \<Sum>i. f (A i))"
+
+lemma (in ring_of_sets) outer_measure_agrees:
+ assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \<in> M"
+ shows "outer_measure M f s = f s"
+ unfolding outer_measure_def
+proof (safe intro!: antisym INF_greatest)
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" and dA: "disjoint_family A" and sA: "s \<subseteq> (\<Union>x. A x)"
+ have inc: "increasing M f"
+ by (metis additive_increasing ca countably_additive_additive posf)
+ have "f s = f (\<Union>i. A i \<inter> s)"
+ using sA by (auto simp: Int_absorb1)
+ also have "\<dots> = (\<Sum>i. f (A i \<inter> s))"
+ using sA dA A s
+ by (intro ca[unfolded countably_additive_def, rule_format, symmetric])
+ (auto simp: Int_absorb1 disjoint_family_on_def)
+ also have "... \<le> (\<Sum>i. f (A i))"
+ using A s by (auto intro!: suminf_le increasingD[OF inc])
+ finally show "f s \<le> (\<Sum>i. f (A i))" .
+next
+ have "(\<Sum>i. f (if i = 0 then s else {})) \<le> f s"
+ using positiveD1[OF posf] by (subst suminf_finite[of "{0}"]) auto
+ with s show "(INF A:{A. range A \<subseteq> M \<and> disjoint_family A \<and> s \<subseteq> UNION UNIV A}. \<Sum>i. f (A i)) \<le> f s"
+ by (intro INF_lower2[of "\<lambda>i. if i = 0 then s else {}"])
+ (auto simp: disjoint_family_on_def)
+qed
+
+lemma outer_measure_empty:
+ "positive M f \<Longrightarrow> {} \<in> M \<Longrightarrow> outer_measure M f {} = 0"
+ unfolding outer_measure_def
+ by (intro antisym INF_lower2[of "\<lambda>_. {}"]) (auto simp: disjoint_family_on_def positive_def)
+
+lemma (in ring_of_sets) positive_outer_measure:
+ assumes "positive M f" shows "positive (Pow \<Omega>) (outer_measure M f)"
+ unfolding positive_def by (auto simp: assms outer_measure_empty)
+
+lemma (in ring_of_sets) increasing_outer_measure: "increasing (Pow \<Omega>) (outer_measure M f)"
+ by (force simp: increasing_def outer_measure_def intro!: INF_greatest intro: INF_lower)
+
+lemma (in ring_of_sets) outer_measure_le:
+ assumes pos: "positive M f" and inc: "increasing M f" and A: "range A \<subseteq> M" and X: "X \<subseteq> (\<Union>i. A i)"
+ shows "outer_measure M f X \<le> (\<Sum>i. f (A i))"
+ unfolding outer_measure_def
+proof (safe intro!: INF_lower2[of "disjointed A"] del: subsetI)
+ show dA: "range (disjointed A) \<subseteq> M"
+ by (auto intro!: A range_disjointed_sets)
+ have "\<forall>n. f (disjointed A n) \<le> f (A n)"
+ by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A)
+ then show "(\<Sum>i. f (disjointed A i)) \<le> (\<Sum>i. f (A i))"
+ by (blast intro!: suminf_le)
+qed (auto simp: X UN_disjointed_eq disjoint_family_disjointed)
+
+lemma (in ring_of_sets) outer_measure_close:
+ "outer_measure M f X < e \<Longrightarrow> \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) < e"
+ unfolding outer_measure_def INF_less_iff by auto
+
+lemma (in ring_of_sets) countably_subadditive_outer_measure:
+ assumes posf: "positive M f" and inc: "increasing M f"
+ shows "countably_subadditive (Pow \<Omega>) (outer_measure M f)"
+proof (simp add: countably_subadditive_def, safe)
+ fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> Pow (\<Omega>)" and sb: "(\<Union>i. A i) \<subseteq> \<Omega>"
+ let ?O = "outer_measure M f"
+ show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n))"
+ proof (rule ennreal_le_epsilon)
+ fix b and e :: real assume "0 < e" "(\<Sum>n. outer_measure M f (A n)) < top"
+ then have *: "\<And>n. outer_measure M f (A n) < outer_measure M f (A n) + e * (1/2)^Suc n"
+ by (auto simp add: less_top dest!: ennreal_suminf_lessD)
+ obtain B
+ where B: "\<And>n. range (B n) \<subseteq> M"
+ and sbB: "\<And>n. A n \<subseteq> (\<Union>i. B n i)"
+ and Ble: "\<And>n. (\<Sum>i. f (B n i)) \<le> ?O (A n) + e * (1/2)^(Suc n)"
+ by (metis less_imp_le outer_measure_close[OF *])
+
+ define C where "C = case_prod B o prod_decode"
+ from B have B_in_M: "\<And>i j. B i j \<in> M"
+ by (rule range_subsetD)
+ then have C: "range C \<subseteq> M"
+ by (auto simp add: C_def split_def)
+ have A_C: "(\<Union>i. A i) \<subseteq> (\<Union>i. C i)"
+ using sbB by (auto simp add: C_def subset_eq) (metis prod.case prod_encode_inverse)
+
+ have "?O (\<Union>i. A i) \<le> ?O (\<Union>i. C i)"
+ using A_C A C by (intro increasing_outer_measure[THEN increasingD]) (auto dest!: sets_into_space)
+ also have "\<dots> \<le> (\<Sum>i. f (C i))"
+ using C by (intro outer_measure_le[OF posf inc]) auto
+ also have "\<dots> = (\<Sum>n. \<Sum>i. f (B n i))"
+ using B_in_M unfolding C_def comp_def by (intro suminf_ennreal_2dimen) auto
+ also have "\<dots> \<le> (\<Sum>n. ?O (A n) + e * (1/2) ^ Suc n)"
+ using B_in_M by (intro suminf_le suminf_nonneg allI Ble) auto
+ also have "... = (\<Sum>n. ?O (A n)) + (\<Sum>n. ennreal e * ennreal ((1/2) ^ Suc n))"
+ using \<open>0 < e\<close> by (subst suminf_add[symmetric])
+ (auto simp del: ennreal_suminf_cmult simp add: ennreal_mult[symmetric])
+ also have "\<dots> = (\<Sum>n. ?O (A n)) + e"
+ unfolding ennreal_suminf_cmult
+ by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
+ finally show "?O (\<Union>i. A i) \<le> (\<Sum>n. ?O (A n)) + e" .
+ qed
+qed
+
+lemma (in ring_of_sets) outer_measure_space_outer_measure:
+ "positive M f \<Longrightarrow> increasing M f \<Longrightarrow> outer_measure_space (Pow \<Omega>) (outer_measure M f)"
+ by (simp add: outer_measure_space_def
+ positive_outer_measure increasing_outer_measure countably_subadditive_outer_measure)
+
+lemma (in ring_of_sets) algebra_subset_lambda_system:
+ assumes posf: "positive M f" and inc: "increasing M f"
+ and add: "additive M f"
+ shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (outer_measure M f)"
+proof (auto dest: sets_into_space
+ simp add: algebra.lambda_system_eq [OF algebra_Pow])
+ fix x s assume x: "x \<in> M" and s: "s \<subseteq> \<Omega>"
+ have [simp]: "\<And>x. x \<in> M \<Longrightarrow> s \<inter> (\<Omega> - x) = s - x" using s
+ by blast
+ have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> outer_measure M f s"
+ unfolding outer_measure_def[of M f s]
+ proof (safe intro!: INF_greatest)
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "disjoint_family A" "range A \<subseteq> M" "s \<subseteq> (\<Union>i. A i)"
+ have "outer_measure M f (s \<inter> x) \<le> (\<Sum>i. f (A i \<inter> x))"
+ unfolding outer_measure_def
+ proof (safe intro!: INF_lower2[of "\<lambda>i. A i \<inter> x"])
+ from A(1) show "disjoint_family (\<lambda>i. A i \<inter> x)"
+ by (rule disjoint_family_on_bisimulation) auto
+ qed (insert x A, auto)
+ moreover
+ have "outer_measure M f (s - x) \<le> (\<Sum>i. f (A i - x))"
+ unfolding outer_measure_def
+ proof (safe intro!: INF_lower2[of "\<lambda>i. A i - x"])
+ from A(1) show "disjoint_family (\<lambda>i. A i - x)"
+ by (rule disjoint_family_on_bisimulation) auto
+ qed (insert x A, auto)
+ ultimately have "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le>
+ (\<Sum>i. f (A i \<inter> x)) + (\<Sum>i. f (A i - x))" by (rule add_mono)
+ also have "\<dots> = (\<Sum>i. f (A i \<inter> x) + f (A i - x))"
+ using A(2) x posf by (subst suminf_add) (auto simp: positive_def)
+ also have "\<dots> = (\<Sum>i. f (A i))"
+ using A x
+ by (subst add[THEN additiveD, symmetric])
+ (auto intro!: arg_cong[where f=suminf] arg_cong[where f=f])
+ finally show "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) \<le> (\<Sum>i. f (A i))" .
+ qed
+ moreover
+ have "outer_measure M f s \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)"
+ proof -
+ have "outer_measure M f s = outer_measure M f ((s \<inter> x) \<union> (s - x))"
+ by (metis Un_Diff_Int Un_commute)
+ also have "... \<le> outer_measure M f (s \<inter> x) + outer_measure M f (s - x)"
+ apply (rule subadditiveD)
+ apply (rule ring_of_sets.countably_subadditive_subadditive [OF ring_of_sets_Pow])
+ apply (simp add: positive_def outer_measure_empty[OF posf])
+ apply (rule countably_subadditive_outer_measure)
+ using s by (auto intro!: posf inc)
+ finally show ?thesis .
+ qed
+ ultimately
+ show "outer_measure M f (s \<inter> x) + outer_measure M f (s - x) = outer_measure M f s"
+ by (rule order_antisym)
+qed
+
+lemma measure_down: "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
+ by (auto simp add: measure_space_def positive_def countably_additive_def subset_eq)
+
+subsection \<open>Caratheodory's theorem\<close>
+
+theorem (in ring_of_sets) caratheodory':
+ assumes posf: "positive M f" and ca: "countably_additive M f"
+ shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+proof -
+ have inc: "increasing M f"
+ by (metis additive_increasing ca countably_additive_additive posf)
+ let ?O = "outer_measure M f"
+ define ls where "ls = lambda_system \<Omega> (Pow \<Omega>) ?O"
+ have mls: "measure_space \<Omega> ls ?O"
+ using sigma_algebra.caratheodory_lemma
+ [OF sigma_algebra_Pow outer_measure_space_outer_measure [OF posf inc]]
+ by (simp add: ls_def)
+ hence sls: "sigma_algebra \<Omega> ls"
+ by (simp add: measure_space_def)
+ have "M \<subseteq> ls"
+ by (simp add: ls_def)
+ (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
+ hence sgs_sb: "sigma_sets (\<Omega>) (M) \<subseteq> ls"
+ using sigma_algebra.sigma_sets_subset [OF sls, of "M"]
+ by simp
+ have "measure_space \<Omega> (sigma_sets \<Omega> M) ?O"
+ by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
+ (simp_all add: sgs_sb space_closed)
+ thus ?thesis using outer_measure_agrees [OF posf ca]
+ by (intro exI[of _ ?O]) auto
+qed
+
+lemma (in ring_of_sets) caratheodory_empty_continuous:
+ assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
+ assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
+ shows "\<exists>\<mu> :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+proof (intro caratheodory' empty_continuous_imp_countably_additive f)
+ show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
+qed (rule cont)
+
+subsection \<open>Volumes\<close>
+
+definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+ "volume M f \<longleftrightarrow>
+ (f {} = 0) \<and> (\<forall>a\<in>M. 0 \<le> f a) \<and>
+ (\<forall>C\<subseteq>M. disjoint C \<longrightarrow> finite C \<longrightarrow> \<Union>C \<in> M \<longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c))"
+
+lemma volumeI:
+ assumes "f {} = 0"
+ assumes "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> f a"
+ assumes "\<And>C. C \<subseteq> M \<Longrightarrow> disjoint C \<Longrightarrow> finite C \<Longrightarrow> \<Union>C \<in> M \<Longrightarrow> f (\<Union>C) = (\<Sum>c\<in>C. f c)"
+ shows "volume M f"
+ using assms by (auto simp: volume_def)
+
+lemma volume_positive:
+ "volume M f \<Longrightarrow> a \<in> M \<Longrightarrow> 0 \<le> f a"
+ by (auto simp: volume_def)
+
+lemma volume_empty:
+ "volume M f \<Longrightarrow> f {} = 0"
+ by (auto simp: volume_def)
+
+lemma volume_finite_additive:
+ assumes "volume M f"
+ assumes A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M" "disjoint_family_on A I" "finite I" "UNION I A \<in> M"
+ shows "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
+proof -
+ have "A`I \<subseteq> M" "disjoint (A`I)" "finite (A`I)" "\<Union>(A`I) \<in> M"
+ using A by (auto simp: disjoint_family_on_disjoint_image)
+ with \<open>volume M f\<close> have "f (\<Union>(A`I)) = (\<Sum>a\<in>A`I. f a)"
+ unfolding volume_def by blast
+ also have "\<dots> = (\<Sum>i\<in>I. f (A i))"
+ proof (subst setsum.reindex_nontrivial)
+ fix i j assume "i \<in> I" "j \<in> I" "i \<noteq> j" "A i = A j"
+ with \<open>disjoint_family_on A I\<close> have "A i = {}"
+ by (auto simp: disjoint_family_on_def)
+ then show "f (A i) = 0"
+ using volume_empty[OF \<open>volume M f\<close>] by simp
+ qed (auto intro: \<open>finite I\<close>)
+ finally show "f (UNION I A) = (\<Sum>i\<in>I. f (A i))"
+ by simp
+qed
+
+lemma (in ring_of_sets) volume_additiveI:
+ assumes pos: "\<And>a. a \<in> M \<Longrightarrow> 0 \<le> \<mu> a"
+ assumes [simp]: "\<mu> {} = 0"
+ assumes add: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> \<mu> (a \<union> b) = \<mu> a + \<mu> b"
+ shows "volume M \<mu>"
+proof (unfold volume_def, safe)
+ fix C assume "finite C" "C \<subseteq> M" "disjoint C"
+ then show "\<mu> (\<Union>C) = setsum \<mu> C"
+ proof (induct C)
+ case (insert c C)
+ from insert(1,2,4,5) have "\<mu> (\<Union>insert c C) = \<mu> c + \<mu> (\<Union>C)"
+ by (auto intro!: add simp: disjoint_def)
+ with insert show ?case
+ by (simp add: disjoint_def)
+ qed simp
+qed fact+
+
+lemma (in semiring_of_sets) extend_volume:
+ assumes "volume M \<mu>"
+ shows "\<exists>\<mu>'. volume generated_ring \<mu>' \<and> (\<forall>a\<in>M. \<mu>' a = \<mu> a)"
+proof -
+ let ?R = generated_ring
+ have "\<forall>a\<in>?R. \<exists>m. \<exists>C\<subseteq>M. a = \<Union>C \<and> finite C \<and> disjoint C \<and> m = (\<Sum>c\<in>C. \<mu> c)"
+ by (auto simp: generated_ring_def)
+ from bchoice[OF this] guess \<mu>' .. note \<mu>'_spec = this
+
+ { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
+ fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
+ assume "\<Union>C = \<Union>D"
+ have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>d\<in>D. \<Sum>c\<in>C. \<mu> (c \<inter> d))"
+ proof (intro setsum.cong refl)
+ fix d assume "d \<in> D"
+ have Un_eq_d: "(\<Union>c\<in>C. c \<inter> d) = d"
+ using \<open>d \<in> D\<close> \<open>\<Union>C = \<Union>D\<close> by auto
+ moreover have "\<mu> (\<Union>c\<in>C. c \<inter> d) = (\<Sum>c\<in>C. \<mu> (c \<inter> d))"
+ proof (rule volume_finite_additive)
+ { fix c assume "c \<in> C" then show "c \<inter> d \<in> M"
+ using C D \<open>d \<in> D\<close> by auto }
+ show "(\<Union>a\<in>C. a \<inter> d) \<in> M"
+ unfolding Un_eq_d using \<open>d \<in> D\<close> D by auto
+ show "disjoint_family_on (\<lambda>a. a \<inter> d) C"
+ using \<open>disjoint C\<close> by (auto simp: disjoint_family_on_def disjoint_def)
+ qed fact+
+ ultimately show "\<mu> d = (\<Sum>c\<in>C. \<mu> (c \<inter> d))" by simp
+ qed }
+ note split_sum = this
+
+ { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
+ fix D assume D: "D \<subseteq> M" "finite D" "disjoint D"
+ assume "\<Union>C = \<Union>D"
+ with split_sum[OF C D] split_sum[OF D C]
+ have "(\<Sum>d\<in>D. \<mu> d) = (\<Sum>c\<in>C. \<mu> c)"
+ by (simp, subst setsum.commute, simp add: ac_simps) }
+ note sum_eq = this
+
+ { fix C assume C: "C \<subseteq> M" "finite C" "disjoint C"
+ then have "\<Union>C \<in> ?R" by (auto simp: generated_ring_def)
+ with \<mu>'_spec[THEN bspec, of "\<Union>C"]
+ obtain D where
+ D: "D \<subseteq> M" "finite D" "disjoint D" "\<Union>C = \<Union>D" and "\<mu>' (\<Union>C) = (\<Sum>d\<in>D. \<mu> d)"
+ by auto
+ with sum_eq[OF C D] have "\<mu>' (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" by simp }
+ note \<mu>' = this
+
+ show ?thesis
+ proof (intro exI conjI ring_of_sets.volume_additiveI[OF generating_ring] ballI)
+ fix a assume "a \<in> M" with \<mu>'[of "{a}"] show "\<mu>' a = \<mu> a"
+ by (simp add: disjoint_def)
+ next
+ fix a assume "a \<in> ?R" then guess Ca .. note Ca = this
+ with \<mu>'[of Ca] \<open>volume M \<mu>\<close>[THEN volume_positive]
+ show "0 \<le> \<mu>' a"
+ by (auto intro!: setsum_nonneg)
+ next
+ show "\<mu>' {} = 0" using \<mu>'[of "{}"] by auto
+ next
+ fix a assume "a \<in> ?R" then guess Ca .. note Ca = this
+ fix b assume "b \<in> ?R" then guess Cb .. note Cb = this
+ assume "a \<inter> b = {}"
+ with Ca Cb have "Ca \<inter> Cb \<subseteq> {{}}" by auto
+ then have C_Int_cases: "Ca \<inter> Cb = {{}} \<or> Ca \<inter> Cb = {}" by auto
+
+ from \<open>a \<inter> b = {}\<close> have "\<mu>' (\<Union>(Ca \<union> Cb)) = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c)"
+ using Ca Cb by (intro \<mu>') (auto intro!: disjoint_union)
+ also have "\<dots> = (\<Sum>c\<in>Ca \<union> Cb. \<mu> c) + (\<Sum>c\<in>Ca \<inter> Cb. \<mu> c)"
+ using C_Int_cases volume_empty[OF \<open>volume M \<mu>\<close>] by (elim disjE) simp_all
+ also have "\<dots> = (\<Sum>c\<in>Ca. \<mu> c) + (\<Sum>c\<in>Cb. \<mu> c)"
+ using Ca Cb by (simp add: setsum.union_inter)
+ also have "\<dots> = \<mu>' a + \<mu>' b"
+ using Ca Cb by (simp add: \<mu>')
+ finally show "\<mu>' (a \<union> b) = \<mu>' a + \<mu>' b"
+ using Ca Cb by simp
+ qed
+qed
+
+subsubsection \<open>Caratheodory on semirings\<close>
+
+theorem (in semiring_of_sets) caratheodory:
+ assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
+ shows "\<exists>\<mu>' :: 'a set \<Rightarrow> ennreal. (\<forall>s \<in> M. \<mu>' s = \<mu> s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>'"
+proof -
+ have "volume M \<mu>"
+ proof (rule volumeI)
+ { fix a assume "a \<in> M" then show "0 \<le> \<mu> a"
+ using pos unfolding positive_def by auto }
+ note p = this
+
+ fix C assume sets_C: "C \<subseteq> M" "\<Union>C \<in> M" and "disjoint C" "finite C"
+ have "\<exists>F'. bij_betw F' {..<card C} C"
+ by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) auto
+ then guess F' .. note F' = this
+ then have F': "C = F' ` {..< card C}" "inj_on F' {..< card C}"
+ by (auto simp: bij_betw_def)
+ { fix i j assume *: "i < card C" "j < card C" "i \<noteq> j"
+ with F' have "F' i \<in> C" "F' j \<in> C" "F' i \<noteq> F' j"
+ unfolding inj_on_def by auto
+ with \<open>disjoint C\<close>[THEN disjointD]
+ have "F' i \<inter> F' j = {}"
+ by auto }
+ note F'_disj = this
+ define F where "F i = (if i < card C then F' i else {})" for i
+ then have "disjoint_family F"
+ using F'_disj by (auto simp: disjoint_family_on_def)
+ moreover from F' have "(\<Union>i. F i) = \<Union>C"
+ by (auto simp add: F_def split: if_split_asm) blast
+ moreover have sets_F: "\<And>i. F i \<in> M"
+ using F' sets_C by (auto simp: F_def)
+ moreover note sets_C
+ ultimately have "\<mu> (\<Union>C) = (\<Sum>i. \<mu> (F i))"
+ using ca[unfolded countably_additive_def, THEN spec, of F] by auto
+ also have "\<dots> = (\<Sum>i<card C. \<mu> (F' i))"
+ proof -
+ have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) sums (\<Sum>i<card C. \<mu> (F' i))"
+ by (rule sums_If_finite_set) auto
+ also have "(\<lambda>i. if i \<in> {..< card C} then \<mu> (F' i) else 0) = (\<lambda>i. \<mu> (F i))"
+ using pos by (auto simp: positive_def F_def)
+ finally show "(\<Sum>i. \<mu> (F i)) = (\<Sum>i<card C. \<mu> (F' i))"
+ by (simp add: sums_iff)
+ qed
+ also have "\<dots> = (\<Sum>c\<in>C. \<mu> c)"
+ using F'(2) by (subst (2) F') (simp add: setsum.reindex)
+ finally show "\<mu> (\<Union>C) = (\<Sum>c\<in>C. \<mu> c)" .
+ next
+ show "\<mu> {} = 0"
+ using \<open>positive M \<mu>\<close> by (rule positiveD1)
+ qed
+ from extend_volume[OF this] obtain \<mu>_r where
+ V: "volume generated_ring \<mu>_r" "\<And>a. a \<in> M \<Longrightarrow> \<mu> a = \<mu>_r a"
+ by auto
+
+ interpret G: ring_of_sets \<Omega> generated_ring
+ by (rule generating_ring)
+
+ have pos: "positive generated_ring \<mu>_r"
+ using V unfolding positive_def by (auto simp: positive_def intro!: volume_positive volume_empty)
+
+ have "countably_additive generated_ring \<mu>_r"
+ proof (rule countably_additiveI)
+ fix A' :: "nat \<Rightarrow> 'a set" assume A': "range A' \<subseteq> generated_ring" "disjoint_family A'"
+ and Un_A: "(\<Union>i. A' i) \<in> generated_ring"
+
+ from generated_ringE[OF Un_A] guess C' . note C' = this
+
+ { fix c assume "c \<in> C'"
+ moreover define A where [abs_def]: "A i = A' i \<inter> c" for i
+ ultimately have A: "range A \<subseteq> generated_ring" "disjoint_family A"
+ and Un_A: "(\<Union>i. A i) \<in> generated_ring"
+ using A' C'
+ by (auto intro!: G.Int G.finite_Union intro: generated_ringI_Basic simp: disjoint_family_on_def)
+ from A C' \<open>c \<in> C'\<close> have UN_eq: "(\<Union>i. A i) = c"
+ by (auto simp: A_def)
+
+ have "\<forall>i::nat. \<exists>f::nat \<Rightarrow> 'a set. \<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j)) \<and> disjoint_family f \<and> \<Union>range f = A i \<and> (\<forall>j. f j \<in> M)"
+ (is "\<forall>i. ?P i")
+ proof
+ fix i
+ from A have Ai: "A i \<in> generated_ring" by auto
+ from generated_ringE[OF this] guess C . note C = this
+
+ have "\<exists>F'. bij_betw F' {..<card C} C"
+ by (rule finite_same_card_bij[OF _ \<open>finite C\<close>]) auto
+ then guess F .. note F = this
+ define f where [abs_def]: "f i = (if i < card C then F i else {})" for i
+ then have f: "bij_betw f {..< card C} C"
+ by (intro bij_betw_cong[THEN iffD1, OF _ F]) auto
+ with C have "\<forall>j. f j \<in> M"
+ by (auto simp: Pi_iff f_def dest!: bij_betw_imp_funcset)
+ moreover
+ from f C have d_f: "disjoint_family_on f {..<card C}"
+ by (intro disjoint_image_disjoint_family_on) (auto simp: bij_betw_def)
+ then have "disjoint_family f"
+ by (auto simp: disjoint_family_on_def f_def)
+ moreover
+ have Ai_eq: "A i = (\<Union>x<card C. f x)"
+ using f C Ai unfolding bij_betw_def by auto
+ then have "\<Union>range f = A i"
+ using f C Ai unfolding bij_betw_def
+ by (auto simp add: f_def cong del: strong_SUP_cong)
+ moreover
+ { have "(\<Sum>j. \<mu>_r (f j)) = (\<Sum>j. if j \<in> {..< card C} then \<mu>_r (f j) else 0)"
+ using volume_empty[OF V(1)] by (auto intro!: arg_cong[where f=suminf] simp: f_def)
+ also have "\<dots> = (\<Sum>j<card C. \<mu>_r (f j))"
+ by (rule sums_If_finite_set[THEN sums_unique, symmetric]) simp
+ also have "\<dots> = \<mu>_r (A i)"
+ using C f[THEN bij_betw_imp_funcset] unfolding Ai_eq
+ by (intro volume_finite_additive[OF V(1) _ d_f, symmetric])
+ (auto simp: Pi_iff Ai_eq intro: generated_ringI_Basic)
+ finally have "\<mu>_r (A i) = (\<Sum>j. \<mu>_r (f j))" .. }
+ ultimately show "?P i"
+ by blast
+ qed
+ from choice[OF this] guess f .. note f = this
+ then have UN_f_eq: "(\<Union>i. case_prod f (prod_decode i)) = (\<Union>i. A i)"
+ unfolding UN_extend_simps surj_prod_decode by (auto simp: set_eq_iff)
+
+ have d: "disjoint_family (\<lambda>i. case_prod f (prod_decode i))"
+ unfolding disjoint_family_on_def
+ proof (intro ballI impI)
+ fix m n :: nat assume "m \<noteq> n"
+ then have neq: "prod_decode m \<noteq> prod_decode n"
+ using inj_prod_decode[of UNIV] by (auto simp: inj_on_def)
+ show "case_prod f (prod_decode m) \<inter> case_prod f (prod_decode n) = {}"
+ proof cases
+ assume "fst (prod_decode m) = fst (prod_decode n)"
+ then show ?thesis
+ using neq f by (fastforce simp: disjoint_family_on_def)
+ next
+ assume neq: "fst (prod_decode m) \<noteq> fst (prod_decode n)"
+ have "case_prod f (prod_decode m) \<subseteq> A (fst (prod_decode m))"
+ "case_prod f (prod_decode n) \<subseteq> A (fst (prod_decode n))"
+ using f[THEN spec, of "fst (prod_decode m)"]
+ using f[THEN spec, of "fst (prod_decode n)"]
+ by (auto simp: set_eq_iff)
+ with f A neq show ?thesis
+ by (fastforce simp: disjoint_family_on_def subset_eq set_eq_iff)
+ qed
+ qed
+ from f have "(\<Sum>n. \<mu>_r (A n)) = (\<Sum>n. \<mu>_r (case_prod f (prod_decode n)))"
+ by (intro suminf_ennreal_2dimen[symmetric] generated_ringI_Basic)
+ (auto split: prod.split)
+ also have "\<dots> = (\<Sum>n. \<mu> (case_prod f (prod_decode n)))"
+ using f V(2) by (auto intro!: arg_cong[where f=suminf] split: prod.split)
+ also have "\<dots> = \<mu> (\<Union>i. case_prod f (prod_decode i))"
+ using f \<open>c \<in> C'\<close> C'
+ by (intro ca[unfolded countably_additive_def, rule_format])
+ (auto split: prod.split simp: UN_f_eq d UN_eq)
+ finally have "(\<Sum>n. \<mu>_r (A' n \<inter> c)) = \<mu> c"
+ using UN_f_eq UN_eq by (simp add: A_def) }
+ note eq = this
+
+ have "(\<Sum>n. \<mu>_r (A' n)) = (\<Sum>n. \<Sum>c\<in>C'. \<mu>_r (A' n \<inter> c))"
+ using C' A'
+ by (subst volume_finite_additive[symmetric, OF V(1)])
+ (auto simp: disjoint_def disjoint_family_on_def
+ intro!: G.Int G.finite_Union arg_cong[where f="\<lambda>X. suminf (\<lambda>i. \<mu>_r (X i))"] ext
+ intro: generated_ringI_Basic)
+ also have "\<dots> = (\<Sum>c\<in>C'. \<Sum>n. \<mu>_r (A' n \<inter> c))"
+ using C' A'
+ by (intro suminf_setsum G.Int G.finite_Union) (auto intro: generated_ringI_Basic)
+ also have "\<dots> = (\<Sum>c\<in>C'. \<mu>_r c)"
+ using eq V C' by (auto intro!: setsum.cong)
+ also have "\<dots> = \<mu>_r (\<Union>C')"
+ using C' Un_A
+ by (subst volume_finite_additive[symmetric, OF V(1)])
+ (auto simp: disjoint_family_on_def disjoint_def
+ intro: generated_ringI_Basic)
+ finally show "(\<Sum>n. \<mu>_r (A' n)) = \<mu>_r (\<Union>i. A' i)"
+ using C' by simp
+ qed
+ from G.caratheodory'[OF \<open>positive generated_ring \<mu>_r\<close> \<open>countably_additive generated_ring \<mu>_r\<close>]
+ guess \<mu>' ..
+ with V show ?thesis
+ unfolding sigma_sets_generated_ring_eq
+ by (intro exI[of _ \<mu>']) (auto intro: generated_ringI_Basic)
+qed
+
+lemma extend_measure_caratheodory:
+ fixes G :: "'i \<Rightarrow> 'a set"
+ assumes M: "M = extend_measure \<Omega> I G \<mu>"
+ assumes "i \<in> I"
+ assumes "semiring_of_sets \<Omega> (G ` I)"
+ assumes empty: "\<And>i. i \<in> I \<Longrightarrow> G i = {} \<Longrightarrow> \<mu> i = 0"
+ assumes inj: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> G i = G j \<Longrightarrow> \<mu> i = \<mu> j"
+ assumes nonneg: "\<And>i. i \<in> I \<Longrightarrow> 0 \<le> \<mu> i"
+ assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>j. A \<in> UNIV \<rightarrow> I \<Longrightarrow> j \<in> I \<Longrightarrow> disjoint_family (G \<circ> A) \<Longrightarrow>
+ (\<Union>i. G (A i)) = G j \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> j"
+ shows "emeasure M (G i) = \<mu> i"
+proof -
+ interpret semiring_of_sets \<Omega> "G ` I"
+ by fact
+ have "\<forall>g\<in>G`I. \<exists>i\<in>I. g = G i"
+ by auto
+ then obtain sel where sel: "\<And>g. g \<in> G ` I \<Longrightarrow> sel g \<in> I" "\<And>g. g \<in> G ` I \<Longrightarrow> G (sel g) = g"
+ by metis
+
+ have "\<exists>\<mu>'. (\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'"
+ proof (rule caratheodory)
+ show "positive (G ` I) (\<lambda>s. \<mu> (sel s))"
+ by (auto simp: positive_def intro!: empty sel nonneg)
+ show "countably_additive (G ` I) (\<lambda>s. \<mu> (sel s))"
+ proof (rule countably_additiveI)
+ fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> G ` I" "disjoint_family A" "(\<Union>i. A i) \<in> G ` I"
+ then show "(\<Sum>i. \<mu> (sel (A i))) = \<mu> (sel (\<Union>i. A i))"
+ by (intro add) (auto simp: sel image_subset_iff_funcset comp_def Pi_iff intro!: sel)
+ qed
+ qed
+ then obtain \<mu>' where \<mu>': "\<forall>s\<in>G ` I. \<mu>' s = \<mu> (sel s)" "measure_space \<Omega> (sigma_sets \<Omega> (G ` I)) \<mu>'"
+ by metis
+
+ show ?thesis
+ proof (rule emeasure_extend_measure[OF M])
+ { fix i assume "i \<in> I" then show "\<mu>' (G i) = \<mu> i"
+ using \<mu>' by (auto intro!: inj sel) }
+ show "G ` I \<subseteq> Pow \<Omega>"
+ by fact
+ then show "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
+ using \<mu>' by (simp_all add: M sets_extend_measure measure_space_def)
+ qed fact
+qed
+
+lemma extend_measure_caratheodory_pair:
+ fixes G :: "'i \<Rightarrow> 'j \<Rightarrow> 'a set"
+ assumes M: "M = extend_measure \<Omega> {(a, b). P a b} (\<lambda>(a, b). G a b) (\<lambda>(a, b). \<mu> a b)"
+ assumes "P i j"
+ assumes semiring: "semiring_of_sets \<Omega> {G a b | a b. P a b}"
+ assumes empty: "\<And>i j. P i j \<Longrightarrow> G i j = {} \<Longrightarrow> \<mu> i j = 0"
+ assumes inj: "\<And>i j k l. P i j \<Longrightarrow> P k l \<Longrightarrow> G i j = G k l \<Longrightarrow> \<mu> i j = \<mu> k l"
+ assumes nonneg: "\<And>i j. P i j \<Longrightarrow> 0 \<le> \<mu> i j"
+ assumes add: "\<And>A::nat \<Rightarrow> 'i. \<And>B::nat \<Rightarrow> 'j. \<And>j k.
+ (\<And>n. P (A n) (B n)) \<Longrightarrow> P j k \<Longrightarrow> disjoint_family (\<lambda>n. G (A n) (B n)) \<Longrightarrow>
+ (\<Union>i. G (A i) (B i)) = G j k \<Longrightarrow> (\<Sum>n. \<mu> (A n) (B n)) = \<mu> j k"
+ shows "emeasure M (G i j) = \<mu> i j"
+proof -
+ have "emeasure M ((\<lambda>(a, b). G a b) (i, j)) = (\<lambda>(a, b). \<mu> a b) (i, j)"
+ proof (rule extend_measure_caratheodory[OF M])
+ show "semiring_of_sets \<Omega> ((\<lambda>(a, b). G a b) ` {(a, b). P a b})"
+ using semiring by (simp add: image_def conj_commute)
+ next
+ fix A :: "nat \<Rightarrow> ('i \<times> 'j)" and j assume "A \<in> UNIV \<rightarrow> {(a, b). P a b}" "j \<in> {(a, b). P a b}"
+ "disjoint_family ((\<lambda>(a, b). G a b) \<circ> A)"
+ "(\<Union>i. case A i of (a, b) \<Rightarrow> G a b) = (case j of (a, b) \<Rightarrow> G a b)"
+ then show "(\<Sum>n. case A n of (a, b) \<Rightarrow> \<mu> a b) = (case j of (a, b) \<Rightarrow> \<mu> a b)"
+ using add[of "\<lambda>i. fst (A i)" "\<lambda>i. snd (A i)" "fst j" "snd j"]
+ by (simp add: split_beta' comp_def Pi_iff)
+ qed (auto split: prod.splits intro: assms)
+ then show ?thesis by simp
+qed
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Complete_Measure.thy Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,307 @@
+(* Title: HOL/Probability/Complete_Measure.thy
+ Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen
+*)
+
+theory Complete_Measure
+ imports Bochner_Integration
+begin
+
+definition
+ "split_completion M A p = (if A \<in> sets M then p = (A, {}) else
+ \<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and> fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets M)"
+
+definition
+ "main_part M A = fst (Eps (split_completion M A))"
+
+definition
+ "null_part M A = snd (Eps (split_completion M A))"
+
+definition completion :: "'a measure \<Rightarrow> 'a measure" where
+ "completion M = measure_of (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }
+ (emeasure M \<circ> main_part M)"
+
+lemma completion_into_space:
+ "{ S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' } \<subseteq> Pow (space M)"
+ using sets.sets_into_space by auto
+
+lemma space_completion[simp]: "space (completion M) = space M"
+ unfolding completion_def using space_measure_of[OF completion_into_space] by simp
+
+lemma completionI:
+ assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+ shows "A \<in> { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+ using assms by auto
+
+lemma completionE:
+ assumes "A \<in> { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+ obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+ using assms by auto
+
+lemma sigma_algebra_completion:
+ "sigma_algebra (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+ (is "sigma_algebra _ ?A")
+ unfolding sigma_algebra_iff2
+proof (intro conjI ballI allI impI)
+ show "?A \<subseteq> Pow (space M)"
+ using sets.sets_into_space by auto
+next
+ show "{} \<in> ?A" by auto
+next
+ let ?C = "space M"
+ fix A assume "A \<in> ?A" from completionE[OF this] guess S N N' .
+ then show "space M - A \<in> ?A"
+ by (intro completionI[of _ "(?C - S) \<inter> (?C - N')" "(?C - S) \<inter> N' \<inter> (?C - N)"]) auto
+next
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> ?A"
+ then have "\<forall>n. \<exists>S N N'. A n = S \<union> N \<and> S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N'"
+ by (auto simp: image_subset_iff)
+ from choice[OF this] guess S ..
+ from choice[OF this] guess N ..
+ from choice[OF this] guess N' ..
+ then show "UNION UNIV A \<in> ?A"
+ using null_sets_UN[of N']
+ by (intro completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) auto
+qed
+
+lemma sets_completion:
+ "sets (completion M) = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+ using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion] by (simp add: completion_def)
+
+lemma sets_completionE:
+ assumes "A \<in> sets (completion M)"
+ obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+ using assms unfolding sets_completion by auto
+
+lemma sets_completionI:
+ assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+ shows "A \<in> sets (completion M)"
+ using assms unfolding sets_completion by auto
+
+lemma sets_completionI_sets[intro, simp]:
+ "A \<in> sets M \<Longrightarrow> A \<in> sets (completion M)"
+ unfolding sets_completion by force
+
+lemma null_sets_completion:
+ assumes "N' \<in> null_sets M" "N \<subseteq> N'" shows "N \<in> sets (completion M)"
+ using assms by (intro sets_completionI[of N "{}" N N']) auto
+
+lemma split_completion:
+ assumes "A \<in> sets (completion M)"
+ shows "split_completion M A (main_part M A, null_part M A)"
+proof cases
+ assume "A \<in> sets M" then show ?thesis
+ by (simp add: split_completion_def[abs_def] main_part_def null_part_def)
+next
+ assume nA: "A \<notin> sets M"
+ show ?thesis
+ unfolding main_part_def null_part_def if_not_P[OF nA]
+ proof (rule someI2_ex)
+ from assms[THEN sets_completionE] guess S N N' . note A = this
+ let ?P = "(S, N - S)"
+ show "\<exists>p. split_completion M A p"
+ unfolding split_completion_def if_not_P[OF nA] using A
+ proof (intro exI conjI)
+ show "A = fst ?P \<union> snd ?P" using A by auto
+ show "snd ?P \<subseteq> N'" using A by auto
+ qed auto
+ qed auto
+qed
+
+lemma
+ assumes "S \<in> sets (completion M)"
+ shows main_part_sets[intro, simp]: "main_part M S \<in> sets M"
+ and main_part_null_part_Un[simp]: "main_part M S \<union> null_part M S = S"
+ and main_part_null_part_Int[simp]: "main_part M S \<inter> null_part M S = {}"
+ using split_completion[OF assms]
+ by (auto simp: split_completion_def split: if_split_asm)
+
+lemma main_part[simp]: "S \<in> sets M \<Longrightarrow> main_part M S = S"
+ using split_completion[of S M]
+ by (auto simp: split_completion_def split: if_split_asm)
+
+lemma null_part:
+ assumes "S \<in> sets (completion M)" shows "\<exists>N. N\<in>null_sets M \<and> null_part M S \<subseteq> N"
+ using split_completion[OF assms] by (auto simp: split_completion_def split: if_split_asm)
+
+lemma null_part_sets[intro, simp]:
+ assumes "S \<in> sets M" shows "null_part M S \<in> sets M" "emeasure M (null_part M S) = 0"
+proof -
+ have S: "S \<in> sets (completion M)" using assms by auto
+ have "S - main_part M S \<in> sets M" using assms by auto
+ moreover
+ from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S]
+ have "S - main_part M S = null_part M S" by auto
+ ultimately show sets: "null_part M S \<in> sets M" by auto
+ from null_part[OF S] guess N ..
+ with emeasure_eq_0[of N _ "null_part M S"] sets
+ show "emeasure M (null_part M S) = 0" by auto
+qed
+
+lemma emeasure_main_part_UN:
+ fixes S :: "nat \<Rightarrow> 'a set"
+ assumes "range S \<subseteq> sets (completion M)"
+ shows "emeasure M (main_part M (\<Union>i. (S i))) = emeasure M (\<Union>i. main_part M (S i))"
+proof -
+ have S: "\<And>i. S i \<in> sets (completion M)" using assms by auto
+ then have UN: "(\<Union>i. S i) \<in> sets (completion M)" by auto
+ have "\<forall>i. \<exists>N. N \<in> null_sets M \<and> null_part M (S i) \<subseteq> N"
+ using null_part[OF S] by auto
+ from choice[OF this] guess N .. note N = this
+ then have UN_N: "(\<Union>i. N i) \<in> null_sets M" by (intro null_sets_UN) auto
+ have "(\<Union>i. S i) \<in> sets (completion M)" using S by auto
+ from null_part[OF this] guess N' .. note N' = this
+ let ?N = "(\<Union>i. N i) \<union> N'"
+ have null_set: "?N \<in> null_sets M" using N' UN_N by (intro null_sets.Un) auto
+ have "main_part M (\<Union>i. S i) \<union> ?N = (main_part M (\<Union>i. S i) \<union> null_part M (\<Union>i. S i)) \<union> ?N"
+ using N' by auto
+ also have "\<dots> = (\<Union>i. main_part M (S i) \<union> null_part M (S i)) \<union> ?N"
+ unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto
+ also have "\<dots> = (\<Union>i. main_part M (S i)) \<union> ?N"
+ using N by auto
+ finally have *: "main_part M (\<Union>i. S i) \<union> ?N = (\<Union>i. main_part M (S i)) \<union> ?N" .
+ have "emeasure M (main_part M (\<Union>i. S i)) = emeasure M (main_part M (\<Union>i. S i) \<union> ?N)"
+ using null_set UN by (intro emeasure_Un_null_set[symmetric]) auto
+ also have "\<dots> = emeasure M ((\<Union>i. main_part M (S i)) \<union> ?N)"
+ unfolding * ..
+ also have "\<dots> = emeasure M (\<Union>i. main_part M (S i))"
+ using null_set S by (intro emeasure_Un_null_set) auto
+ finally show ?thesis .
+qed
+
+lemma emeasure_completion[simp]:
+ assumes S: "S \<in> sets (completion M)" shows "emeasure (completion M) S = emeasure M (main_part M S)"
+proof (subst emeasure_measure_of[OF completion_def completion_into_space])
+ let ?\<mu> = "emeasure M \<circ> main_part M"
+ show "S \<in> sets (completion M)" "?\<mu> S = emeasure M (main_part M S) " using S by simp_all
+ show "positive (sets (completion M)) ?\<mu>"
+ by (simp add: positive_def)
+ show "countably_additive (sets (completion M)) ?\<mu>"
+ proof (intro countably_additiveI)
+ fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets (completion M)" "disjoint_family A"
+ have "disjoint_family (\<lambda>i. main_part M (A i))"
+ proof (intro disjoint_family_on_bisimulation[OF A(2)])
+ fix n m assume "A n \<inter> A m = {}"
+ then have "(main_part M (A n) \<union> null_part M (A n)) \<inter> (main_part M (A m) \<union> null_part M (A m)) = {}"
+ using A by (subst (1 2) main_part_null_part_Un) auto
+ then show "main_part M (A n) \<inter> main_part M (A m) = {}" by auto
+ qed
+ then have "(\<Sum>n. emeasure M (main_part M (A n))) = emeasure M (\<Union>i. main_part M (A i))"
+ using A by (auto intro!: suminf_emeasure)
+ then show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (UNION UNIV A)"
+ by (simp add: completion_def emeasure_main_part_UN[OF A(1)])
+ qed
+qed
+
+lemma emeasure_completion_UN:
+ "range S \<subseteq> sets (completion M) \<Longrightarrow>
+ emeasure (completion M) (\<Union>i::nat. (S i)) = emeasure M (\<Union>i. main_part M (S i))"
+ by (subst emeasure_completion) (auto simp add: emeasure_main_part_UN)
+
+lemma emeasure_completion_Un:
+ assumes S: "S \<in> sets (completion M)" and T: "T \<in> sets (completion M)"
+ shows "emeasure (completion M) (S \<union> T) = emeasure M (main_part M S \<union> main_part M T)"
+proof (subst emeasure_completion)
+ have UN: "(\<Union>i. binary (main_part M S) (main_part M T) i) = (\<Union>i. main_part M (binary S T i))"
+ unfolding binary_def by (auto split: if_split_asm)
+ show "emeasure M (main_part M (S \<union> T)) = emeasure M (main_part M S \<union> main_part M T)"
+ using emeasure_main_part_UN[of "binary S T" M] assms
+ by (simp add: range_binary_eq, simp add: Un_range_binary UN)
+qed (auto intro: S T)
+
+lemma sets_completionI_sub:
+ assumes N: "N' \<in> null_sets M" "N \<subseteq> N'"
+ shows "N \<in> sets (completion M)"
+ using assms by (intro sets_completionI[of _ "{}" N N']) auto
+
+lemma completion_ex_simple_function:
+ assumes f: "simple_function (completion M) f"
+ shows "\<exists>f'. simple_function M f' \<and> (AE x in M. f x = f' x)"
+proof -
+ let ?F = "\<lambda>x. f -` {x} \<inter> space M"
+ have F: "\<And>x. ?F x \<in> sets (completion M)" and fin: "finite (f`space M)"
+ using simple_functionD[OF f] simple_functionD[OF f] by simp_all
+ have "\<forall>x. \<exists>N. N \<in> null_sets M \<and> null_part M (?F x) \<subseteq> N"
+ using F null_part by auto
+ from choice[OF this] obtain N where
+ N: "\<And>x. null_part M (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets M" by auto
+ let ?N = "\<Union>x\<in>f`space M. N x"
+ let ?f' = "\<lambda>x. if x \<in> ?N then undefined else f x"
+ have sets: "?N \<in> null_sets M" using N fin by (intro null_sets.finite_UN) auto
+ show ?thesis unfolding simple_function_def
+ proof (safe intro!: exI[of _ ?f'])
+ have "?f' ` space M \<subseteq> f`space M \<union> {undefined}" by auto
+ from finite_subset[OF this] simple_functionD(1)[OF f]
+ show "finite (?f' ` space M)" by auto
+ next
+ fix x assume "x \<in> space M"
+ have "?f' -` {?f' x} \<inter> space M =
+ (if x \<in> ?N then ?F undefined \<union> ?N
+ else if f x = undefined then ?F (f x) \<union> ?N
+ else ?F (f x) - ?N)"
+ using N(2) sets.sets_into_space by (auto split: if_split_asm simp: null_sets_def)
+ moreover { fix y have "?F y \<union> ?N \<in> sets M"
+ proof cases
+ assume y: "y \<in> f`space M"
+ have "?F y \<union> ?N = (main_part M (?F y) \<union> null_part M (?F y)) \<union> ?N"
+ using main_part_null_part_Un[OF F] by auto
+ also have "\<dots> = main_part M (?F y) \<union> ?N"
+ using y N by auto
+ finally show ?thesis
+ using F sets by auto
+ next
+ assume "y \<notin> f`space M" then have "?F y = {}" by auto
+ then show ?thesis using sets by auto
+ qed }
+ moreover {
+ have "?F (f x) - ?N = main_part M (?F (f x)) \<union> null_part M (?F (f x)) - ?N"
+ using main_part_null_part_Un[OF F] by auto
+ also have "\<dots> = main_part M (?F (f x)) - ?N"
+ using N \<open>x \<in> space M\<close> by auto
+ finally have "?F (f x) - ?N \<in> sets M"
+ using F sets by auto }
+ ultimately show "?f' -` {?f' x} \<inter> space M \<in> sets M" by auto
+ next
+ show "AE x in M. f x = ?f' x"
+ by (rule AE_I', rule sets) auto
+ qed
+qed
+
+lemma completion_ex_borel_measurable:
+ fixes g :: "'a \<Rightarrow> ennreal"
+ assumes g: "g \<in> borel_measurable (completion M)"
+ shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
+proof -
+ from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this
+ from this(1)[THEN completion_ex_simple_function]
+ have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x in M. f i x = f' x)" ..
+ from this[THEN choice] obtain f' where
+ sf: "\<And>i. simple_function M (f' i)" and
+ AE: "\<forall>i. AE x in M. f i x = f' i x" by auto
+ show ?thesis
+ proof (intro bexI)
+ from AE[unfolded AE_all_countable[symmetric]]
+ show "AE x in M. g x = (SUP i. f' i x)" (is "AE x in M. g x = ?f x")
+ proof (elim AE_mp, safe intro!: AE_I2)
+ fix x assume eq: "\<forall>i. f i x = f' i x"
+ moreover have "g x = (SUP i. f i x)"
+ unfolding f by (auto split: split_max)
+ ultimately show "g x = ?f x" by auto
+ qed
+ show "?f \<in> borel_measurable M"
+ using sf[THEN borel_measurable_simple_function] by auto
+ qed
+qed
+
+lemma null_sets_completionI: "N \<in> null_sets M \<Longrightarrow> N \<in> null_sets (completion M)"
+ by (auto simp: null_sets_def)
+
+lemma AE_completion: "(AE x in M. P x) \<Longrightarrow> (AE x in completion M. P x)"
+ unfolding eventually_ae_filter by (auto intro: null_sets_completionI)
+
+lemma null_sets_completion_iff: "N \<in> sets M \<Longrightarrow> N \<in> null_sets (completion M) \<longleftrightarrow> N \<in> null_sets M"
+ by (auto simp: null_sets_def)
+
+lemma AE_completion_iff: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in completion M. P x)"
+ by (simp add: AE_iff_null null_sets_completion_iff)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Embed_Measure.thy Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,398 @@
+(* Title: HOL/Probability/Embed_Measure.thy
+ Author: Manuel Eberl, TU München
+
+ Defines measure embeddings with injective functions, i.e. lifting a measure on some values
+ to a measure on "tagged" values (e.g. embed_measure M Inl lifts a measure on values 'a to a
+ measure on the left part of the sum type 'a + 'b)
+*)
+
+section \<open>Embed Measure Spaces with a Function\<close>
+
+theory Embed_Measure
+imports Binary_Product_Measure
+begin
+
+definition embed_measure :: "'a measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
+ "embed_measure M f = measure_of (f ` space M) {f ` A |A. A \<in> sets M}
+ (\<lambda>A. emeasure M (f -` A \<inter> space M))"
+
+lemma space_embed_measure: "space (embed_measure M f) = f ` space M"
+ unfolding embed_measure_def
+ by (subst space_measure_of) (auto dest: sets.sets_into_space)
+
+lemma sets_embed_measure':
+ assumes inj: "inj_on f (space M)"
+ shows "sets (embed_measure M f) = {f ` A |A. A \<in> sets M}"
+ unfolding embed_measure_def
+proof (intro sigma_algebra.sets_measure_of_eq sigma_algebra_iff2[THEN iffD2] conjI allI ballI impI)
+ fix s assume "s \<in> {f ` A |A. A \<in> sets M}"
+ then obtain s' where s'_props: "s = f ` s'" "s' \<in> sets M" by auto
+ hence "f ` space M - s = f ` (space M - s')" using inj
+ by (auto dest: inj_onD sets.sets_into_space)
+ also have "... \<in> {f ` A |A. A \<in> sets M}" using s'_props by auto
+ finally show "f ` space M - s \<in> {f ` A |A. A \<in> sets M}" .
+next
+ fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> {f ` A |A. A \<in> sets M}"
+ then obtain A' where A': "\<And>i. A i = f ` A' i" "\<And>i. A' i \<in> sets M"
+ by (auto simp: subset_eq choice_iff)
+ then have "(\<Union>x. f ` A' x) = f ` (\<Union>x. A' x)" by blast
+ with A' show "(\<Union>i. A i) \<in> {f ` A |A. A \<in> sets M}"
+ by simp blast
+qed (auto dest: sets.sets_into_space)
+
+lemma the_inv_into_vimage:
+ "inj_on f X \<Longrightarrow> A \<subseteq> X \<Longrightarrow> the_inv_into X f -` A \<inter> (f`X) = f ` A"
+ by (auto simp: the_inv_into_f_f)
+
+lemma sets_embed_eq_vimage_algebra:
+ assumes "inj_on f (space M)"
+ shows "sets (embed_measure M f) = sets (vimage_algebra (f`space M) (the_inv_into (space M) f) M)"
+ by (auto simp: sets_embed_measure'[OF assms] Pi_iff the_inv_into_f_f assms sets_vimage_algebra2 simple_image
+ dest: sets.sets_into_space
+ intro!: image_cong the_inv_into_vimage[symmetric])
+
+lemma sets_embed_measure:
+ assumes inj: "inj f"
+ shows "sets (embed_measure M f) = {f ` A |A. A \<in> sets M}"
+ using assms by (subst sets_embed_measure') (auto intro!: inj_onI dest: injD)
+
+lemma in_sets_embed_measure: "A \<in> sets M \<Longrightarrow> f ` A \<in> sets (embed_measure M f)"
+ unfolding embed_measure_def
+ by (intro in_measure_of) (auto dest: sets.sets_into_space)
+
+lemma measurable_embed_measure1:
+ assumes g: "(\<lambda>x. g (f x)) \<in> measurable M N"
+ shows "g \<in> measurable (embed_measure M f) N"
+ unfolding measurable_def
+proof safe
+ fix A assume "A \<in> sets N"
+ with g have "(\<lambda>x. g (f x)) -` A \<inter> space M \<in> sets M"
+ by (rule measurable_sets)
+ then have "f ` ((\<lambda>x. g (f x)) -` A \<inter> space M) \<in> sets (embed_measure M f)"
+ by (rule in_sets_embed_measure)
+ also have "f ` ((\<lambda>x. g (f x)) -` A \<inter> space M) = g -` A \<inter> space (embed_measure M f)"
+ by (auto simp: space_embed_measure)
+ finally show "g -` A \<inter> space (embed_measure M f) \<in> sets (embed_measure M f)" .
+qed (insert measurable_space[OF assms], auto simp: space_embed_measure)
+
+lemma measurable_embed_measure2':
+ assumes "inj_on f (space M)"
+ shows "f \<in> measurable M (embed_measure M f)"
+proof-
+ {
+ fix A assume A: "A \<in> sets M"
+ also from A have "A = A \<inter> space M" by auto
+ also have "... = f -` f ` A \<inter> space M" using A assms
+ by (auto dest: inj_onD sets.sets_into_space)
+ finally have "f -` f ` A \<inter> space M \<in> sets M" .
+ }
+ thus ?thesis using assms unfolding embed_measure_def
+ by (intro measurable_measure_of) (auto dest: sets.sets_into_space)
+qed
+
+lemma measurable_embed_measure2:
+ assumes [simp]: "inj f" shows "f \<in> measurable M (embed_measure M f)"
+ by (auto simp: inj_vimage_image_eq embed_measure_def
+ intro!: measurable_measure_of dest: sets.sets_into_space)
+
+lemma embed_measure_eq_distr':
+ assumes "inj_on f (space M)"
+ shows "embed_measure M f = distr M (embed_measure M f) f"
+proof-
+ have "distr M (embed_measure M f) f =
+ measure_of (f ` space M) {f ` A |A. A \<in> sets M}
+ (\<lambda>A. emeasure M (f -` A \<inter> space M))" unfolding distr_def
+ by (simp add: space_embed_measure sets_embed_measure'[OF assms])
+ also have "... = embed_measure M f" unfolding embed_measure_def ..
+ finally show ?thesis ..
+qed
+
+lemma embed_measure_eq_distr:
+ "inj f \<Longrightarrow> embed_measure M f = distr M (embed_measure M f) f"
+ by (rule embed_measure_eq_distr') (auto intro!: inj_onI dest: injD)
+
+lemma nn_integral_embed_measure':
+ "inj_on f (space M) \<Longrightarrow> g \<in> borel_measurable (embed_measure M f) \<Longrightarrow>
+ nn_integral (embed_measure M f) g = nn_integral M (\<lambda>x. g (f x))"
+ apply (subst embed_measure_eq_distr', simp)
+ apply (subst nn_integral_distr)
+ apply (simp_all add: measurable_embed_measure2')
+ done
+
+lemma nn_integral_embed_measure:
+ "inj f \<Longrightarrow> g \<in> borel_measurable (embed_measure M f) \<Longrightarrow>
+ nn_integral (embed_measure M f) g = nn_integral M (\<lambda>x. g (f x))"
+ by(erule nn_integral_embed_measure'[OF subset_inj_on]) simp
+
+lemma emeasure_embed_measure':
+ assumes "inj_on f (space M)" "A \<in> sets (embed_measure M f)"
+ shows "emeasure (embed_measure M f) A = emeasure M (f -` A \<inter> space M)"
+ by (subst embed_measure_eq_distr'[OF assms(1)])
+ (simp add: emeasure_distr[OF measurable_embed_measure2'[OF assms(1)] assms(2)])
+
+lemma emeasure_embed_measure:
+ assumes "inj f" "A \<in> sets (embed_measure M f)"
+ shows "emeasure (embed_measure M f) A = emeasure M (f -` A \<inter> space M)"
+ using assms by (intro emeasure_embed_measure') (auto intro!: inj_onI dest: injD)
+
+lemma embed_measure_comp:
+ assumes [simp]: "inj f" "inj g"
+ shows "embed_measure (embed_measure M f) g = embed_measure M (g \<circ> f)"
+proof-
+ have [simp]: "inj (\<lambda>x. g (f x))" by (subst o_def[symmetric]) (auto intro: inj_comp)
+ note measurable_embed_measure2[measurable]
+ have "embed_measure (embed_measure M f) g =
+ distr M (embed_measure (embed_measure M f) g) (g \<circ> f)"
+ by (subst (1 2) embed_measure_eq_distr)
+ (simp_all add: distr_distr sets_embed_measure cong: distr_cong)
+ also have "... = embed_measure M (g \<circ> f)"
+ by (subst (3) embed_measure_eq_distr, simp add: o_def, rule distr_cong)
+ (auto simp: sets_embed_measure o_def image_image[symmetric]
+ intro: inj_comp cong: distr_cong)
+ finally show ?thesis .
+qed
+
+lemma sigma_finite_embed_measure:
+ assumes "sigma_finite_measure M" and inj: "inj f"
+ shows "sigma_finite_measure (embed_measure M f)"
+proof -
+ from assms(1) interpret sigma_finite_measure M .
+ from sigma_finite_countable obtain A where
+ A_props: "countable A" "A \<subseteq> sets M" "\<Union>A = space M" "\<And>X. X\<in>A \<Longrightarrow> emeasure M X \<noteq> \<infinity>" by blast
+ from A_props have "countable (op ` f`A)" by auto
+ moreover
+ from inj and A_props have "op ` f`A \<subseteq> sets (embed_measure M f)"
+ by (auto simp: sets_embed_measure)
+ moreover
+ from A_props and inj have "\<Union>(op ` f`A) = space (embed_measure M f)"
+ by (auto simp: space_embed_measure intro!: imageI)
+ moreover
+ from A_props and inj have "\<forall>a\<in>op ` f ` A. emeasure (embed_measure M f) a \<noteq> \<infinity>"
+ by (intro ballI, subst emeasure_embed_measure)
+ (auto simp: inj_vimage_image_eq intro: in_sets_embed_measure)
+ ultimately show ?thesis by - (standard, blast)
+qed
+
+lemma embed_measure_count_space':
+ "inj_on f A \<Longrightarrow> embed_measure (count_space A) f = count_space (f`A)"
+ apply (subst distr_bij_count_space[of f A "f`A", symmetric])
+ apply (simp add: inj_on_def bij_betw_def)
+ apply (subst embed_measure_eq_distr')
+ apply simp
+ apply(auto 4 3 intro!: measure_eqI imageI simp add: sets_embed_measure' subset_image_iff)
+ apply (subst (1 2) emeasure_distr)
+ apply (auto simp: space_embed_measure sets_embed_measure')
+ done
+
+lemma embed_measure_count_space:
+ "inj f \<Longrightarrow> embed_measure (count_space A) f = count_space (f`A)"
+ by(rule embed_measure_count_space')(erule subset_inj_on, simp)
+
+lemma sets_embed_measure_alt:
+ "inj f \<Longrightarrow> sets (embed_measure M f) = (op`f) ` sets M"
+ by (auto simp: sets_embed_measure)
+
+lemma emeasure_embed_measure_image':
+ assumes "inj_on f (space M)" "X \<in> sets M"
+ shows "emeasure (embed_measure M f) (f`X) = emeasure M X"
+proof-
+ from assms have "emeasure (embed_measure M f) (f`X) = emeasure M (f -` f ` X \<inter> space M)"
+ by (subst emeasure_embed_measure') (auto simp: sets_embed_measure')
+ also from assms have "f -` f ` X \<inter> space M = X" by (auto dest: inj_onD sets.sets_into_space)
+ finally show ?thesis .
+qed
+
+lemma emeasure_embed_measure_image:
+ "inj f \<Longrightarrow> X \<in> sets M \<Longrightarrow> emeasure (embed_measure M f) (f`X) = emeasure M X"
+ by (simp_all add: emeasure_embed_measure in_sets_embed_measure inj_vimage_image_eq)
+
+lemma embed_measure_eq_iff:
+ assumes "inj f"
+ shows "embed_measure A f = embed_measure B f \<longleftrightarrow> A = B" (is "?M = ?N \<longleftrightarrow> _")
+proof
+ from assms have I: "inj (op` f)" by (auto intro: injI dest: injD)
+ assume asm: "?M = ?N"
+ hence "sets (embed_measure A f) = sets (embed_measure B f)" by simp
+ with assms have "sets A = sets B" by (simp only: I inj_image_eq_iff sets_embed_measure_alt)
+ moreover {
+ fix X assume "X \<in> sets A"
+ from asm have "emeasure ?M (f`X) = emeasure ?N (f`X)" by simp
+ with \<open>X \<in> sets A\<close> and \<open>sets A = sets B\<close> and assms
+ have "emeasure A X = emeasure B X" by (simp add: emeasure_embed_measure_image)
+ }
+ ultimately show "A = B" by (rule measure_eqI)
+qed simp
+
+lemma the_inv_into_in_Pi: "inj_on f A \<Longrightarrow> the_inv_into A f \<in> f ` A \<rightarrow> A"
+ by (auto simp: the_inv_into_f_f)
+
+lemma map_prod_image: "map_prod f g ` (A \<times> B) = (f`A) \<times> (g`B)"
+ using map_prod_surj_on[OF refl refl] .
+
+lemma map_prod_vimage: "map_prod f g -` (A \<times> B) = (f-`A) \<times> (g-`B)"
+ by auto
+
+lemma embed_measure_prod:
+ assumes f: "inj f" and g: "inj g" and [simp]: "sigma_finite_measure M" "sigma_finite_measure N"
+ shows "embed_measure M f \<Otimes>\<^sub>M embed_measure N g = embed_measure (M \<Otimes>\<^sub>M N) (\<lambda>(x, y). (f x, g y))"
+ (is "?L = _")
+ unfolding map_prod_def[symmetric]
+proof (rule pair_measure_eqI)
+ have fg[simp]: "\<And>A. inj_on (map_prod f g) A" "\<And>A. inj_on f A" "\<And>A. inj_on g A"
+ using f g by (auto simp: inj_on_def)
+
+ note complete_lattice_class.Sup_insert[simp del] ccSup_insert[simp del] SUP_insert[simp del]
+ ccSUP_insert[simp del]
+ show sets: "sets ?L = sets (embed_measure (M \<Otimes>\<^sub>M N) (map_prod f g))"
+ unfolding map_prod_def[symmetric]
+ apply (simp add: sets_pair_eq_sets_fst_snd sets_embed_eq_vimage_algebra
+ cong: vimage_algebra_cong)
+ apply (subst sets_vimage_Sup_eq[where Y="space (M \<Otimes>\<^sub>M N)"])
+ apply (simp_all add: space_pair_measure[symmetric])
+ apply (auto simp add: the_inv_into_f_f
+ simp del: map_prod_simp
+ del: prod_fun_imageE) []
+ apply auto []
+ apply (subst image_insert)
+ apply simp
+ apply (subst (1 2 3 4 ) vimage_algebra_vimage_algebra_eq)
+ apply (simp_all add: the_inv_into_in_Pi Pi_iff[of snd] Pi_iff[of fst] space_pair_measure)
+ apply (simp_all add: Pi_iff[of snd] Pi_iff[of fst] the_inv_into_in_Pi vimage_algebra_vimage_algebra_eq
+ space_pair_measure[symmetric] map_prod_image[symmetric])
+ apply (intro arg_cong[where f=sets] arg_cong[where f=Sup] arg_cong2[where f=insert] vimage_algebra_cong)
+ apply (auto simp: map_prod_image the_inv_into_f_f
+ simp del: map_prod_simp del: prod_fun_imageE)
+ apply (simp_all add: the_inv_into_f_f space_pair_measure)
+ done
+
+ note measurable_embed_measure2[measurable]
+ fix A B assume AB: "A \<in> sets (embed_measure M f)" "B \<in> sets (embed_measure N g)"
+ moreover have "f -` A \<times> g -` B \<inter> space (M \<Otimes>\<^sub>M N) = (f -` A \<inter> space M) \<times> (g -` B \<inter> space N)"
+ by (auto simp: space_pair_measure)
+ ultimately show "emeasure (embed_measure M f) A * emeasure (embed_measure N g) B =
+ emeasure (embed_measure (M \<Otimes>\<^sub>M N) (map_prod f g)) (A \<times> B)"
+ by (simp add: map_prod_vimage sets[symmetric] emeasure_embed_measure
+ sigma_finite_measure.emeasure_pair_measure_Times)
+qed (insert assms, simp_all add: sigma_finite_embed_measure)
+
+lemma mono_embed_measure:
+ "space M = space M' \<Longrightarrow> sets M \<subseteq> sets M' \<Longrightarrow> sets (embed_measure M f) \<subseteq> sets (embed_measure M' f)"
+ unfolding embed_measure_def
+ apply (subst (1 2) sets_measure_of)
+ apply (blast dest: sets.sets_into_space)
+ apply (blast dest: sets.sets_into_space)
+ apply simp
+ apply (intro sigma_sets_mono')
+ apply safe
+ apply (simp add: subset_eq)
+ apply metis
+ done
+
+lemma density_embed_measure:
+ assumes inj: "inj f" and Mg[measurable]: "g \<in> borel_measurable (embed_measure M f)"
+ shows "density (embed_measure M f) g = embed_measure (density M (g \<circ> f)) f" (is "?M1 = ?M2")
+proof (rule measure_eqI)
+ fix X assume X: "X \<in> sets ?M1"
+ from inj have Mf[measurable]: "f \<in> measurable M (embed_measure M f)"
+ by (rule measurable_embed_measure2)
+ from Mg and X have "emeasure ?M1 X = \<integral>\<^sup>+ x. g x * indicator X x \<partial>embed_measure M f"
+ by (subst emeasure_density) simp_all
+ also from X have "... = \<integral>\<^sup>+ x. g (f x) * indicator X (f x) \<partial>M"
+ by (subst embed_measure_eq_distr[OF inj], subst nn_integral_distr) auto
+ also have "... = \<integral>\<^sup>+ x. g (f x) * indicator (f -` X \<inter> space M) x \<partial>M"
+ by (intro nn_integral_cong) (auto split: split_indicator)
+ also from X have "... = emeasure (density M (g \<circ> f)) (f -` X \<inter> space M)"
+ by (subst emeasure_density) (simp_all add: measurable_comp[OF Mf Mg] measurable_sets[OF Mf])
+ also from X and inj have "... = emeasure ?M2 X"
+ by (subst emeasure_embed_measure) (simp_all add: sets_embed_measure)
+ finally show "emeasure ?M1 X = emeasure ?M2 X" .
+qed (simp_all add: sets_embed_measure inj)
+
+lemma density_embed_measure':
+ assumes inj: "inj f" and inv: "\<And>x. f' (f x) = x" and Mg[measurable]: "g \<in> borel_measurable M"
+ shows "density (embed_measure M f) (g \<circ> f') = embed_measure (density M g) f"
+proof-
+ have "density (embed_measure M f) (g \<circ> f') = embed_measure (density M (g \<circ> f' \<circ> f)) f"
+ by (rule density_embed_measure[OF inj])
+ (rule measurable_comp, rule measurable_embed_measure1, subst measurable_cong,
+ rule inv, rule measurable_ident_sets, simp, rule Mg)
+ also have "density M (g \<circ> f' \<circ> f) = density M g"
+ by (intro density_cong) (subst measurable_cong, simp add: o_def inv, simp_all add: Mg inv)
+ finally show ?thesis .
+qed
+
+lemma inj_on_image_subset_iff:
+ assumes "inj_on f C" "A \<subseteq> C" "B \<subseteq> C"
+ shows "f ` A \<subseteq> f ` B \<longleftrightarrow> A \<subseteq> B"
+proof (intro iffI subsetI)
+ fix x assume A: "f ` A \<subseteq> f ` B" and B: "x \<in> A"
+ from B have "f x \<in> f ` A" by blast
+ with A have "f x \<in> f ` B" by blast
+ then obtain y where "f x = f y" and "y \<in> B" by blast
+ with assms and B have "x = y" by (auto dest: inj_onD)
+ with \<open>y \<in> B\<close> show "x \<in> B" by simp
+qed auto
+
+
+lemma AE_embed_measure':
+ assumes inj: "inj_on f (space M)"
+ shows "(AE x in embed_measure M f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
+proof
+ let ?M = "embed_measure M f"
+ assume "AE x in ?M. P x"
+ then obtain A where A_props: "A \<in> sets ?M" "emeasure ?M A = 0" "{x\<in>space ?M. \<not>P x} \<subseteq> A"
+ by (force elim: AE_E)
+ then obtain A' where A'_props: "A = f ` A'" "A' \<in> sets M" by (auto simp: sets_embed_measure' inj)
+ moreover have B: "{x\<in>space ?M. \<not>P x} = f ` {x\<in>space M. \<not>P (f x)}"
+ by (auto simp: inj space_embed_measure)
+ from A_props(3) have "{x\<in>space M. \<not>P (f x)} \<subseteq> A'"
+ by (subst (asm) B, subst (asm) A'_props, subst (asm) inj_on_image_subset_iff[OF inj])
+ (insert A'_props, auto dest: sets.sets_into_space)
+ moreover from A_props A'_props have "emeasure M A' = 0"
+ by (simp add: emeasure_embed_measure_image' inj)
+ ultimately show "AE x in M. P (f x)" by (intro AE_I)
+next
+ let ?M = "embed_measure M f"
+ assume "AE x in M. P (f x)"
+ then obtain A where A_props: "A \<in> sets M" "emeasure M A = 0" "{x\<in>space M. \<not>P (f x)} \<subseteq> A"
+ by (force elim: AE_E)
+ hence "f`A \<in> sets ?M" "emeasure ?M (f`A) = 0" "{x\<in>space ?M. \<not>P x} \<subseteq> f`A"
+ by (auto simp: space_embed_measure emeasure_embed_measure_image' sets_embed_measure' inj)
+ thus "AE x in ?M. P x" by (intro AE_I)
+qed
+
+lemma AE_embed_measure:
+ assumes inj: "inj f"
+ shows "(AE x in embed_measure M f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
+ using assms by (intro AE_embed_measure') (auto intro!: inj_onI dest: injD)
+
+lemma nn_integral_monotone_convergence_SUP_countable:
+ fixes f :: "'a \<Rightarrow> 'b \<Rightarrow> ennreal"
+ assumes nonempty: "Y \<noteq> {}"
+ and chain: "Complete_Partial_Order.chain op \<le> (f ` Y)"
+ and countable: "countable B"
+ shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space B) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space B))"
+ (is "?lhs = ?rhs")
+proof -
+ let ?f = "(\<lambda>i x. f i (from_nat_into B x) * indicator (to_nat_on B ` B) x)"
+ have "?lhs = \<integral>\<^sup>+ x. (SUP i:Y. f i (from_nat_into B (to_nat_on B x))) \<partial>count_space B"
+ by(rule nn_integral_cong)(simp add: countable)
+ also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. f i (from_nat_into B x)) \<partial>count_space (to_nat_on B ` B)"