move measure theory from HOL-Probability to HOL-Multivariate_Analysis
authorhoelzl
Fri, 05 Aug 2016 18:34:57 +0200
changeset 63626 44ce6b524ff3
parent 63625 1e7c5bbea36d
child 63627 6ddb43c6b711
move measure theory from HOL-Probability to HOL-Multivariate_Analysis
src/HOL/Multivariate_Analysis/Binary_Product_Measure.thy
src/HOL/Multivariate_Analysis/Bochner_Integration.thy
src/HOL/Multivariate_Analysis/Borel_Space.thy
src/HOL/Multivariate_Analysis/Caratheodory.thy
src/HOL/Multivariate_Analysis/Complete_Measure.thy
src/HOL/Multivariate_Analysis/Embed_Measure.thy
src/HOL/Multivariate_Analysis/Finite_Product_Measure.thy
src/HOL/Multivariate_Analysis/Interval_Integral.thy
src/HOL/Multivariate_Analysis/Lebesgue_Integral_Substitution.thy
src/HOL/Multivariate_Analysis/Lebesgue_Measure.thy
src/HOL/Multivariate_Analysis/Measurable.thy
src/HOL/Multivariate_Analysis/Measure_Space.thy
src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy
src/HOL/Multivariate_Analysis/Nonnegative_Lebesgue_Integration.thy
src/HOL/Multivariate_Analysis/Radon_Nikodym.thy
src/HOL/Multivariate_Analysis/Regularity.thy
src/HOL/Multivariate_Analysis/Set_Integral.thy
src/HOL/Multivariate_Analysis/Sigma_Algebra.thy
src/HOL/Multivariate_Analysis/measurable.ML
src/HOL/Probability/Binary_Product_Measure.thy
src/HOL/Probability/Bochner_Integration.thy
src/HOL/Probability/Borel_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Characteristic_Functions.thy
src/HOL/Probability/Complete_Measure.thy
src/HOL/Probability/Embed_Measure.thy
src/HOL/Probability/Finite_Product_Measure.thy
src/HOL/Probability/Giry_Monad.thy
src/HOL/Probability/Independent_Family.thy
src/HOL/Probability/Infinite_Product_Measure.thy
src/HOL/Probability/Information.thy
src/HOL/Probability/Interval_Integral.thy
src/HOL/Probability/Lebesgue_Integral_Substitution.thy
src/HOL/Probability/Lebesgue_Measure.thy
src/HOL/Probability/Measurable.thy
src/HOL/Probability/Measure_Space.thy
src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy
src/HOL/Probability/Probability.thy
src/HOL/Probability/Probability_Measure.thy
src/HOL/Probability/Projective_Family.thy
src/HOL/Probability/Projective_Limit.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/Probability/Regularity.thy
src/HOL/Probability/SPMF.thy
src/HOL/Probability/Set_Integral.thy
src/HOL/Probability/Sigma_Algebra.thy
src/HOL/Probability/measurable.ML
--- /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)"
+    by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1)
+  also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. ?f i x) \<partial>count_space UNIV"
+    by(simp add: nn_integral_count_space_indicator ennreal_indicator[symmetric] SUP_mult_right_ennreal nonempty)
+  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV)"
+  proof(rule nn_integral_monotone_convergence_SUP_nat)
+    show "Complete_Partial_Order.chain op \<le> (?f ` Y)"
+      by(rule chain_imageI[OF chain, unfolded image_image])(auto intro!: le_funI split: split_indicator dest: le_funD)
+  qed fact
+  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B x) \<partial>count_space (to_nat_on B ` B))"
+    by(simp add: nn_integral_count_space_indicator)
+  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B (to_nat_on B x)) \<partial>count_space B)"
+    by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1)
+  also have "\<dots> = ?rhs"
+    by(intro arg_cong2[where f="SUPREMUM"] ext nn_integral_cong_AE)(simp_all add: AE_count_space countable)
+  finally show ?thesis .
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Finite_Product_Measure.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,1199 @@
+(*  Title:      HOL/Probability/Finite_Product_Measure.thy
+    Author:     Johannes Hölzl, TU München
+*)
+
+section \<open>Finite product measures\<close>
+
+theory Finite_Product_Measure
+imports Binary_Product_Measure
+begin
+
+lemma PiE_choice: "(\<exists>f\<in>PiE I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
+  by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1])
+     (force intro: exI[of _ "restrict f I" for f])
+
+lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
+  by auto
+
+subsubsection \<open>More about Function restricted by @{const extensional}\<close>
+
+definition
+  "merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
+
+lemma merge_apply[simp]:
+  "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i"
+  "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i"
+  "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i"
+  "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i"
+  "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I J (x, y) i = undefined"
+  unfolding merge_def by auto
+
+lemma merge_commute:
+  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)"
+  by (force simp: merge_def)
+
+lemma Pi_cancel_merge_range[simp]:
+  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A"
+  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A"
+  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A"
+  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A"
+  by (auto simp: Pi_def)
+
+lemma Pi_cancel_merge[simp]:
+  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
+  "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
+  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
+  "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
+  by (auto simp: Pi_def)
+
+lemma extensional_merge[simp]: "merge I J (x, y) \<in> extensional (I \<union> J)"
+  by (auto simp: extensional_def)
+
+lemma restrict_merge[simp]:
+  "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I"
+  "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
+  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I"
+  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
+  by (auto simp: restrict_def)
+
+lemma split_merge: "P (merge I J (x,y) i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)"
+  unfolding merge_def by auto
+
+lemma PiE_cancel_merge[simp]:
+  "I \<inter> J = {} \<Longrightarrow>
+    merge I J (x, y) \<in> PiE (I \<union> J) B \<longleftrightarrow> x \<in> Pi I B \<and> y \<in> Pi J B"
+  by (auto simp: PiE_def restrict_Pi_cancel)
+
+lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)"
+  unfolding merge_def by (auto simp: fun_eq_iff)
+
+lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I J (x, y) \<in> extensional K"
+  unfolding merge_def extensional_def by auto
+
+lemma merge_restrict[simp]:
+  "merge I J (restrict x I, y) = merge I J (x, y)"
+  "merge I J (x, restrict y J) = merge I J (x, y)"
+  unfolding merge_def by auto
+
+lemma merge_x_x_eq_restrict[simp]:
+  "merge I J (x, x) = restrict x (I \<union> J)"
+  unfolding merge_def by auto
+
+lemma injective_vimage_restrict:
+  assumes J: "J \<subseteq> I"
+  and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}"
+  and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
+  shows "A = B"
+proof  (intro set_eqI)
+  fix x
+  from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto
+  have "J \<inter> (I - J) = {}" by auto
+  show "x \<in> A \<longleftrightarrow> x \<in> B"
+  proof cases
+    assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
+    have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
+      using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
+      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
+    then show "x \<in> A \<longleftrightarrow> x \<in> B"
+      using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
+      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq)
+  qed (insert sets, auto)
+qed
+
+lemma restrict_vimage:
+  "I \<inter> J = {} \<Longrightarrow>
+    (\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \<times> Pi\<^sub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
+  by (auto simp: restrict_Pi_cancel PiE_def)
+
+lemma merge_vimage:
+  "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
+  by (auto simp: restrict_Pi_cancel PiE_def)
+
+subsection \<open>Finite product spaces\<close>
+
+subsubsection \<open>Products\<close>
+
+definition prod_emb where
+  "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
+
+lemma prod_emb_iff:
+  "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
+  unfolding prod_emb_def PiE_def by auto
+
+lemma
+  shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
+    and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B"
+    and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B"
+    and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))"
+    and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))"
+    and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
+  by (auto simp: prod_emb_def)
+
+lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
+    prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))"
+  by (force simp: prod_emb_def PiE_iff if_split_mem2)
+
+lemma prod_emb_PiE_same_index[simp]:
+    "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E"
+  by (auto simp: prod_emb_def PiE_iff)
+
+lemma prod_emb_trans[simp]:
+  "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
+  by (auto simp add: Int_absorb1 prod_emb_def PiE_def)
+
+lemma prod_emb_Pi:
+  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
+  shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))"
+  using assms sets.space_closed
+  by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+
+
+lemma prod_emb_id:
+  "B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
+  by (auto simp: prod_emb_def subset_eq extensional_restrict)
+
+lemma prod_emb_mono:
+  "F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G"
+  by (auto simp: prod_emb_def)
+
+definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
+  "PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
+    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
+    (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
+    (\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
+
+definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
+  "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) `
+    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
+
+abbreviation
+  "Pi\<^sub>M I M \<equiv> PiM I M"
+
+syntax
+  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^sub>M _\<in>_./ _)"  10)
+translations
+  "\<Pi>\<^sub>M x\<in>I. M" == "CONST PiM I (%x. M)"
+
+lemma extend_measure_cong:
+  assumes "\<Omega> = \<Omega>'" "I = I'" "G = G'" "\<And>i. i \<in> I' \<Longrightarrow> \<mu> i = \<mu>' i"
+  shows "extend_measure \<Omega> I G \<mu> = extend_measure \<Omega>' I' G' \<mu>'"
+  unfolding extend_measure_def by (auto simp add: assms)
+
+lemma Pi_cong_sets:
+    "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
+  unfolding Pi_def by auto
+
+lemma PiM_cong:
+  assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
+  shows "PiM I M = PiM J N"
+  unfolding PiM_def
+proof (rule extend_measure_cong, goal_cases)
+  case 1
+  show ?case using assms
+    by (subst assms(1), intro PiE_cong[of J "\<lambda>i. space (M i)" "\<lambda>i. space (N i)"]) simp_all
+next
+  case 2
+  have "\<And>K. K \<subseteq> J \<Longrightarrow> (\<Pi> j\<in>K. sets (M j)) = (\<Pi> j\<in>K. sets (N j))"
+    using assms by (intro Pi_cong_sets) auto
+  thus ?case by (auto simp: assms)
+next
+  case 3
+  show ?case using assms
+    by (intro ext) (auto simp: prod_emb_def dest: PiE_mem)
+next
+  case (4 x)
+  thus ?case using assms
+    by (auto intro!: setprod.cong split: if_split_asm)
+qed
+
+
+lemma prod_algebra_sets_into_space:
+  "prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  by (auto simp: prod_emb_def prod_algebra_def)
+
+lemma prod_algebra_eq_finite:
+  assumes I: "finite I"
+  shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
+proof (intro iffI set_eqI)
+  fix A assume "A \<in> ?L"
+  then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
+    and A: "A = prod_emb I M J (PIE j:J. E j)"
+    by (auto simp: prod_algebra_def)
+  let ?A = "\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i)"
+  have A: "A = ?A"
+    unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto
+  show "A \<in> ?R" unfolding A using J sets.top
+    by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp
+next
+  fix A assume "A \<in> ?R"
+  then obtain X where A: "A = (\<Pi>\<^sub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
+  then have A: "A = prod_emb I M I (\<Pi>\<^sub>E i\<in>I. X i)"
+    by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff)
+  from X I show "A \<in> ?L" unfolding A
+    by (auto simp: prod_algebra_def)
+qed
+
+lemma prod_algebraI:
+  "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
+    \<Longrightarrow> prod_emb I M J (PIE j:J. E j) \<in> prod_algebra I M"
+  by (auto simp: prod_algebra_def)
+
+lemma prod_algebraI_finite:
+  "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M"
+  using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp
+
+lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
+proof (safe intro!: Int_stableI)
+  fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
+  then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
+    by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int)
+qed
+
+lemma prod_algebraE:
+  assumes A: "A \<in> prod_algebra I M"
+  obtains J E where "A = prod_emb I M J (PIE j:J. E j)"
+    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
+  using A by (auto simp: prod_algebra_def)
+
+lemma prod_algebraE_all:
+  assumes A: "A \<in> prod_algebra I M"
+  obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
+proof -
+  from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)"
+    and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
+    by (auto simp: prod_algebra_def)
+  from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)"
+    using sets.sets_into_space by auto
+  then have "A = (\<Pi>\<^sub>E i\<in>I. if i\<in>J then E i else space (M i))"
+    using A J by (auto simp: prod_emb_PiE)
+  moreover have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))"
+    using sets.top E by auto
+  ultimately show ?thesis using that by auto
+qed
+
+lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
+proof (unfold Int_stable_def, safe)
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J E . note A = this
+  fix B assume "B \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess K F . note B = this
+  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter>
+      (if i \<in> K then F i else space (M i)))"
+    unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4)
+      B(5)[THEN sets.sets_into_space]
+    apply (subst (1 2 3) prod_emb_PiE)
+    apply (simp_all add: subset_eq PiE_Int)
+    apply blast
+    apply (intro PiE_cong)
+    apply auto
+    done
+  also have "\<dots> \<in> prod_algebra I M"
+    using A B by (auto intro!: prod_algebraI)
+  finally show "A \<inter> B \<in> prod_algebra I M" .
+qed
+
+lemma prod_algebra_mono:
+  assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
+  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
+  shows "prod_algebra I E \<subseteq> prod_algebra I F"
+proof
+  fix A assume "A \<in> prod_algebra I E"
+  then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
+    and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)"
+    and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)"
+    by (auto simp: prod_algebra_def)
+  moreover
+  from space have "(\<Pi>\<^sub>E i\<in>I. space (E i)) = (\<Pi>\<^sub>E i\<in>I. space (F i))"
+    by (rule PiE_cong)
+  with A have "A = prod_emb I F J (\<Pi>\<^sub>E i\<in>J. G i)"
+    by (simp add: prod_emb_def)
+  moreover
+  from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)"
+    by auto
+  ultimately show "A \<in> prod_algebra I F"
+    apply (simp add: prod_algebra_def image_iff)
+    apply (intro exI[of _ J] exI[of _ G] conjI)
+    apply auto
+    done
+qed
+
+lemma prod_algebra_cong:
+  assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
+  shows "prod_algebra I M = prod_algebra J N"
+proof -
+  have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
+    using sets_eq_imp_space_eq[OF sets] by auto
+  with sets show ?thesis unfolding \<open>I = J\<close>
+    by (intro antisym prod_algebra_mono) auto
+qed
+
+lemma space_in_prod_algebra:
+  "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
+proof cases
+  assume "I = {}" then show ?thesis
+    by (auto simp add: prod_algebra_def image_iff prod_emb_def)
+next
+  assume "I \<noteq> {}"
+  then obtain i where "i \<in> I" by auto
+  then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
+    by (auto simp: prod_emb_def)
+  also have "\<dots> \<in> prod_algebra I M"
+    using \<open>i \<in> I\<close> by (intro prod_algebraI) auto
+  finally show ?thesis .
+qed
+
+lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
+
+lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)"
+  by (auto simp: prod_emb_def space_PiM)
+
+lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow>  (\<exists>i\<in>I. space (M i) = {})"
+  by (auto simp: space_PiM PiE_eq_empty_iff)
+
+lemma undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)"
+  by (auto simp: space_PiM)
+
+lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
+  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
+
+lemma sets_PiM_single: "sets (PiM I M) =
+    sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
+    (is "_ = sigma_sets ?\<Omega> ?R")
+  unfolding sets_PiM
+proof (rule sigma_sets_eqI)
+  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J X . note X = this
+  show "A \<in> sigma_sets ?\<Omega> ?R"
+  proof cases
+    assume "I = {}"
+    with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def)
+    with \<open>I = {}\<close> show ?thesis by (auto intro!: sigma_sets_top)
+  next
+    assume "I \<noteq> {}"
+    with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})"
+      by (auto simp: prod_emb_def)
+    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
+      using X \<open>I \<noteq> {}\<close> by (intro R.finite_INT sigma_sets.Basic) auto
+    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
+  qed
+next
+  fix A assume "A \<in> ?R"
+  then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
+    by auto
+  then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)"
+     by (auto simp: prod_emb_def)
+  also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)"
+    using A by (intro sigma_sets.Basic prod_algebraI) auto
+  finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" .
+qed
+
+lemma sets_PiM_eq_proj:
+  "I \<noteq> {} \<Longrightarrow> sets (PiM I M) = sets (SUP i:I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))"
+  apply (simp add: sets_PiM_single)
+  apply (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"])
+  apply auto []
+  apply auto []
+  apply simp
+  apply (subst SUP_cong[OF refl])
+  apply (rule sets_vimage_algebra2)
+  apply auto []
+  apply (auto intro!: arg_cong2[where f=sigma_sets])
+  done
+
+lemma
+  shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}"
+    and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }"
+  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
+
+lemma sets_PiM_sigma:
+  assumes \<Omega>_cover: "\<And>i. i \<in> I \<Longrightarrow> \<exists>S\<subseteq>E i. countable S \<and> \<Omega> i = \<Union>S"
+  assumes E: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (\<Omega> i)"
+  assumes J: "\<And>j. j \<in> J \<Longrightarrow> finite j" "\<Union>J = I"
+  defines "P \<equiv> {{f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i} | A j. j \<in> J \<and> A \<in> Pi j E}"
+  shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)"
+proof cases
+  assume "I = {}"
+  with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}"
+    by (auto simp: P_def)
+  with \<open>I = {}\<close> show ?thesis
+    by (auto simp add: sets_PiM_empty sigma_sets_empty_eq)
+next
+  let ?F = "\<lambda>i. {(\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega> |A. A \<in> E i}"
+  assume "I \<noteq> {}"
+  then have "sets (Pi\<^sub>M I (\<lambda>i. sigma (\<Omega> i) (E i))) =
+      sets (SUP i:I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<lambda>x. x i) (sigma (\<Omega> i) (E i)))"
+    by (subst sets_PiM_eq_proj) (auto simp: space_measure_of_conv)
+  also have "\<dots> = sets (SUP i:I. sigma (Pi\<^sub>E I \<Omega>) (?F i))"
+    using E by (intro sets_SUP_cong arg_cong[where f=sets] vimage_algebra_sigma) auto
+  also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i))"
+    using \<open>I \<noteq> {}\<close> by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto
+  also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) P)"
+  proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)
+    show "(\<Union>i\<in>I. ?F i) \<subseteq> Pow (Pi\<^sub>E I \<Omega>)" "P \<subseteq> Pow (Pi\<^sub>E I \<Omega>)"
+      by (auto simp: P_def)
+  next
+    interpret P: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
+      by (auto intro!: sigma_algebra_sigma_sets simp: P_def)
+
+    fix Z assume "Z \<in> (\<Union>i\<in>I. ?F i)"
+    then obtain i A where i: "i \<in> I" "A \<in> E i" and Z_def: "Z = (\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega>"
+      by auto
+    from \<open>i \<in> I\<close> J obtain j where j: "i \<in> j" "j \<in> J" "j \<subseteq> I" "finite j"
+      by auto
+    obtain S where S: "\<And>i. i \<in> j \<Longrightarrow> S i \<subseteq> E i" "\<And>i. i \<in> j \<Longrightarrow> countable (S i)"
+      "\<And>i. i \<in> j \<Longrightarrow> \<Omega> i = \<Union>(S i)"
+      by (metis subset_eq \<Omega>_cover \<open>j \<subseteq> I\<close>)
+    define A' where "A' n = n(i := A)" for n
+    then have A'_i: "\<And>n. A' n i = A"
+      by simp
+    { fix n assume "n \<in> Pi\<^sub>E (j - {i}) S"
+      then have "A' n \<in> Pi j E"
+        unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \<open>A \<in> E i\<close> )
+      with \<open>j \<in> J\<close> have "{f \<in> Pi\<^sub>E I \<Omega>. \<forall>i\<in>j. f i \<in> A' n i} \<in> P"
+        by (auto simp: P_def) }
+    note A'_in_P = this
+
+    { fix x assume "x i \<in> A" "x \<in> Pi\<^sub>E I \<Omega>"
+      with S(3) \<open>j \<subseteq> I\<close> have "\<forall>i\<in>j. \<exists>s\<in>S i. x i \<in> s"
+        by (auto simp: PiE_def Pi_def)
+      then obtain s where s: "\<And>i. i \<in> j \<Longrightarrow> s i \<in> S i" "\<And>i. i \<in> j \<Longrightarrow> x i \<in> s i"
+        by metis
+      with \<open>x i \<in> A\<close> have "\<exists>n\<in>PiE (j-{i}) S. \<forall>i\<in>j. x i \<in> A' n i"
+        by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) }
+    then have "Z = (\<Union>n\<in>PiE (j-{i}) S. {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A' n i})"
+      unfolding Z_def
+      by (auto simp add: set_eq_iff ball_conj_distrib \<open>i\<in>j\<close> A'_i dest: bspec[OF _ \<open>i\<in>j\<close>]
+               cong: conj_cong)
+    also have "\<dots> \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
+      using \<open>finite j\<close> S(2)
+      by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P)
+    finally show "Z \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" .
+  next
+    interpret F: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<Union>i\<in>I. ?F i)"
+      by (auto intro!: sigma_algebra_sigma_sets)
+
+    fix b assume "b \<in> P"
+    then obtain A j where b: "b = {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i}" "j \<in> J" "A \<in> Pi j E"
+      by (auto simp: P_def)
+    show "b \<in> sigma_sets (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i)"
+    proof cases
+      assume "j = {}"
+      with b have "b = (\<Pi>\<^sub>E i\<in>I. \<Omega> i)"
+        by auto
+      then show ?thesis
+        by blast
+    next
+      assume "j \<noteq> {}"
+      with J b(2,3) have eq: "b = (\<Inter>i\<in>j. ((\<lambda>x. x i) -` A i \<inter> Pi\<^sub>E I \<Omega>))"
+        unfolding b(1)
+        by (auto simp: PiE_def Pi_def)
+      show ?thesis
+        unfolding eq using \<open>A \<in> Pi j E\<close> \<open>j \<in> J\<close> J(2)
+        by (intro F.finite_INT J \<open>j \<in> J\<close> \<open>j \<noteq> {}\<close> sigma_sets.Basic) blast
+    qed
+  qed
+  finally show "?thesis" .
+qed
+
+lemma sets_PiM_in_sets:
+  assumes space: "space N = (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>i A. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {x\<in>space N. x i \<in> A} \<in> sets N"
+  shows "sets (\<Pi>\<^sub>M i \<in> I. M i) \<subseteq> sets N"
+  unfolding sets_PiM_single space[symmetric]
+  by (intro sets.sigma_sets_subset subsetI) (auto intro: sets)
+
+lemma sets_PiM_cong[measurable_cong]:
+  assumes "I = J" "\<And>i. i \<in> J \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (PiM I M) = sets (PiM J N)"
+  using assms sets_eq_imp_space_eq[OF assms(2)] by (simp add: sets_PiM_single cong: PiE_cong conj_cong)
+
+lemma sets_PiM_I:
+  assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
+  shows "prod_emb I M J (PIE j:J. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
+proof cases
+  assume "J = {}"
+  then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))"
+    by (auto simp: prod_emb_def)
+  then show ?thesis
+    by (auto simp add: sets_PiM intro!: sigma_sets_top)
+next
+  assume "J \<noteq> {}" with assms show ?thesis
+    by (force simp add: sets_PiM prod_algebra_def)
+qed
+
+lemma measurable_PiM:
+  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
+    f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N"
+  shows "f \<in> measurable N (PiM I M)"
+  using sets_PiM prod_algebra_sets_into_space space
+proof (rule measurable_sigma_sets)
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J X .
+  with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto
+qed
+
+lemma measurable_PiM_Collect:
+  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
+    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N"
+  shows "f \<in> measurable N (PiM I M)"
+  using sets_PiM prod_algebra_sets_into_space space
+proof (rule measurable_sigma_sets)
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J X . note X = this
+  then have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}"
+    using space by (auto simp: prod_emb_def del: PiE_I)
+  also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets)
+  finally show "f -` A \<inter> space N \<in> sets N" .
+qed
+
+lemma measurable_PiM_single:
+  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N"
+  shows "f \<in> measurable N (PiM I M)"
+  using sets_PiM_single
+proof (rule measurable_sigma_sets)
+  fix A assume "A \<in> {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+  then obtain B i where "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
+    by auto
+  with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto
+  also have "\<dots> \<in> sets N" using B by (rule sets)
+  finally show "f -` A \<inter> space N \<in> sets N" .
+qed (auto simp: space)
+
+lemma measurable_PiM_single':
+  assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)"
+    and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+  shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)"
+proof (rule measurable_PiM_single)
+  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
+  then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N"
+    by auto
+  then show "{\<omega> \<in> space N. f i \<omega> \<in> A} \<in> sets N"
+    using A f by (auto intro!: measurable_sets)
+qed fact
+
+lemma sets_PiM_I_finite[measurable]:
+  assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
+  shows "(PIE j:I. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
+  using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \<open>finite I\<close> sets by auto
+
+lemma measurable_component_singleton[measurable (raw)]:
+  assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)"
+proof (unfold measurable_def, intro CollectI conjI ballI)
+  fix A assume "A \<in> sets (M i)"
+  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)"
+    using sets.sets_into_space \<open>i \<in> I\<close>
+    by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: if_split_asm)
+  then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)"
+    using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I)
+qed (insert \<open>i \<in> I\<close>, auto simp: space_PiM)
+
+lemma measurable_component_singleton'[measurable_dest]:
+  assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
+  assumes g: "g \<in> measurable L N"
+  assumes i: "i \<in> I"
+  shows "(\<lambda>x. (f (g x)) i) \<in> measurable L (M i)"
+  using measurable_compose[OF measurable_compose[OF g f] measurable_component_singleton, OF i] .
+
+lemma measurable_PiM_component_rev:
+  "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N"
+  by simp
+
+lemma measurable_case_nat[measurable (raw)]:
+  assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N"
+    "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
+  shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N"
+  by (cases i) simp_all
+
+lemma measurable_case_nat'[measurable (raw)]:
+  assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
+  shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
+  using fg[THEN measurable_space]
+  by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split)
+
+lemma measurable_add_dim[measurable]:
+  "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)"
+    (is "?f \<in> measurable ?P ?I")
+proof (rule measurable_PiM_single)
+  fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)"
+  have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} =
+    (if j = i then space (Pi\<^sub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
+    using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM)
+  also have "\<dots> \<in> sets ?P"
+    using A j
+    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
+  finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
+qed (auto simp: space_pair_measure space_PiM PiE_def)
+
+lemma measurable_fun_upd:
+  assumes I: "I = J \<union> {i}"
+  assumes f[measurable]: "f \<in> measurable N (PiM J M)"
+  assumes h[measurable]: "h \<in> measurable N (M i)"
+  shows "(\<lambda>x. (f x) (i := h x)) \<in> measurable N (PiM I M)"
+proof (intro measurable_PiM_single')
+  fix j assume "j \<in> I" then show "(\<lambda>\<omega>. ((f \<omega>)(i := h \<omega>)) j) \<in> measurable N (M j)"
+    unfolding I by (cases "j = i") auto
+next
+  show "(\<lambda>x. (f x)(i := h x)) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+    using I f[THEN measurable_space] h[THEN measurable_space]
+    by (auto simp: space_PiM PiE_iff extensional_def)
+qed
+
+lemma measurable_component_update:
+  "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)"
+  by simp
+
+lemma measurable_merge[measurable]:
+  "merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)"
+    (is "?f \<in> measurable ?P ?U")
+proof (rule measurable_PiM_single)
+  fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J"
+  then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} =
+    (if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)"
+    by (auto simp: merge_def)
+  also have "\<dots> \<in> sets ?P"
+    using A
+    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
+  finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" .
+qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def)
+
+lemma measurable_restrict[measurable (raw)]:
+  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
+  shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)"
+proof (rule measurable_PiM_single)
+  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
+  then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N"
+    by auto
+  then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N"
+    using A X by (auto intro!: measurable_sets)
+qed (insert X, auto simp add: PiE_def dest: measurable_space)
+
+lemma measurable_abs_UNIV:
+  "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)"
+  by (intro measurable_PiM_single) (auto dest: measurable_space)
+
+lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
+  by (intro measurable_restrict measurable_component_singleton) auto
+
+lemma measurable_restrict_subset':
+  assumes "J \<subseteq> L" "\<And>x. x \<in> J \<Longrightarrow> sets (M x) = sets (N x)"
+  shows "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
+proof-
+  from assms(1) have "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
+    by (rule measurable_restrict_subset)
+  also from assms(2) have "measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M) = measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
+    by (intro sets_PiM_cong measurable_cong_sets) simp_all
+  finally show ?thesis .
+qed
+
+lemma measurable_prod_emb[intro, simp]:
+  "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)"
+  unfolding prod_emb_def space_PiM[symmetric]
+  by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
+
+lemma merge_in_prod_emb:
+  assumes "y \<in> space (PiM I M)" "x \<in> X" and X: "X \<in> sets (Pi\<^sub>M J M)" and "J \<subseteq> I"
+  shows "merge J I (x, y) \<in> prod_emb I M J X"
+  using assms sets.sets_into_space[OF X]
+  by (simp add: merge_def prod_emb_def subset_eq space_PiM PiE_def extensional_restrict Pi_iff
+           cong: if_cong restrict_cong)
+     (simp add: extensional_def)
+
+lemma prod_emb_eq_emptyD:
+  assumes J: "J \<subseteq> I" and ne: "space (PiM I M) \<noteq> {}" and X: "X \<in> sets (Pi\<^sub>M J M)"
+    and *: "prod_emb I M J X = {}"
+  shows "X = {}"
+proof safe
+  fix x assume "x \<in> X"
+  obtain \<omega> where "\<omega> \<in> space (PiM I M)"
+    using ne by blast
+  from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto
+qed
+
+lemma sets_in_Pi_aux:
+  "finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
+  {x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)"
+  by (simp add: subset_eq Pi_iff)
+
+lemma sets_in_Pi[measurable (raw)]:
+  "finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow>
+  (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
+  Measurable.pred N (\<lambda>x. f x \<in> Pi I F)"
+  unfolding pred_def
+  by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto
+
+lemma sets_in_extensional_aux:
+  "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)"
+proof -
+  have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)"
+    by (auto simp add: extensional_def space_PiM)
+  then show ?thesis by simp
+qed
+
+lemma sets_in_extensional[measurable (raw)]:
+  "f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)"
+  unfolding pred_def
+  by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto
+
+lemma sets_PiM_I_countable:
+  assumes I: "countable I" and E: "\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i)" shows "Pi\<^sub>E I E \<in> sets (Pi\<^sub>M I M)"
+proof cases
+  assume "I \<noteq> {}"
+  then have "PiE I E = (\<Inter>i\<in>I. prod_emb I M {i} (PiE {i} E))"
+    using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff)
+  also have "\<dots> \<in> sets (PiM I M)"
+    using I \<open>I \<noteq> {}\<close> by (safe intro!: sets.countable_INT' measurable_prod_emb sets_PiM_I_finite E)
+  finally show ?thesis .
+qed (simp add: sets_PiM_empty)
+
+lemma sets_PiM_D_countable:
+  assumes A: "A \<in> PiM I M"
+  shows "\<exists>J\<subseteq>I. \<exists>X\<in>PiM J M. countable J \<and> A = prod_emb I M J X"
+  using A[unfolded sets_PiM_single]
+proof induction
+  case (Basic A)
+  then obtain i X where *: "i \<in> I" "X \<in> sets (M i)" and "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> X}"
+    by auto
+  then have A: "A = prod_emb I M {i} (\<Pi>\<^sub>E _\<in>{i}. X)"
+    by (auto simp: prod_emb_def)
+  then show ?case
+    by (intro exI[of _ "{i}"] conjI bexI[of _ "\<Pi>\<^sub>E _\<in>{i}. X"])
+       (auto intro: countable_finite * sets_PiM_I_finite)
+next
+  case Empty then show ?case
+    by (intro exI[of _ "{}"] conjI bexI[of _ "{}"]) auto
+next
+  case (Compl A)
+  then obtain J X where "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)" "countable J" "A = prod_emb I M J X"
+    by auto
+  then show ?case
+    by (intro exI[of _ J] bexI[of _ "space (PiM J M) - X"] conjI)
+       (auto simp add: space_PiM prod_emb_PiE intro!: sets_PiM_I_countable)
+next
+  case (Union K)
+  obtain J X where J: "\<And>i. J i \<subseteq> I" "\<And>i. countable (J i)" and X: "\<And>i. X i \<in> sets (Pi\<^sub>M (J i) M)"
+    and K: "\<And>i. K i = prod_emb I M (J i) (X i)"
+    by (metis Union.IH)
+  show ?case
+  proof (intro exI[of _ "\<Union>i. J i"] bexI[of _ "\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i)"] conjI)
+    show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" using J by auto
+    with J show "UNION UNIV K = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))"
+      by (simp add: K[abs_def] SUP_upper)
+  qed(auto intro: X)
+qed
+
+lemma measure_eqI_PiM_finite:
+  assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M"
+  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)"
+  assumes A: "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = space (PiM I M)" "\<And>i::nat. P (A i) \<noteq> \<infinity>"
+  shows "P = Q"
+proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
+  show "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. P (A i) \<noteq> \<infinity>"
+    unfolding space_PiM[symmetric] by fact+
+  fix X assume "X \<in> prod_algebra I M"
+  then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"
+    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
+    by (force elim!: prod_algebraE)
+  then show "emeasure P X = emeasure Q X"
+    unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq)
+qed (simp_all add: sets_PiM)
+
+lemma measure_eqI_PiM_infinite:
+  assumes [simp]: "sets P = PiM I M" "sets Q = PiM I M"
+  assumes eq: "\<And>A J. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow>
+    P (prod_emb I M J (Pi\<^sub>E J A)) = Q (prod_emb I M J (Pi\<^sub>E J A))"
+  assumes A: "finite_measure P"
+  shows "P = Q"
+proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
+  interpret finite_measure P by fact
+  define i where "i = (SOME i. i \<in> I)"
+  have i: "I \<noteq> {} \<Longrightarrow> i \<in> I"
+    unfolding i_def by (rule someI_ex) auto
+  define A where "A n =
+    (if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)))"
+    for n :: nat
+  then show "range A \<subseteq> prod_algebra I M"
+    using prod_algebraI[of "{}" I "\<lambda>i. space (M i)" M] by (auto intro!: prod_algebraI i)
+  have "\<And>i. A i = space (PiM I M)"
+    by (auto simp: prod_emb_def space_PiM PiE_iff A_def i ex_in_conv[symmetric] exI)
+  then show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. emeasure P (A i) \<noteq> \<infinity>"
+    by (auto simp: space_PiM)
+next
+  fix X assume X: "X \<in> prod_algebra I M"
+  then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"
+    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
+    by (force elim!: prod_algebraE)
+  then show "emeasure P X = emeasure Q X"
+    by (auto intro!: eq)
+qed (auto simp: sets_PiM)
+
+locale product_sigma_finite =
+  fixes M :: "'i \<Rightarrow> 'a measure"
+  assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
+
+sublocale product_sigma_finite \<subseteq> M?: sigma_finite_measure "M i" for i
+  by (rule sigma_finite_measures)
+
+locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
+  fixes I :: "'i set"
+  assumes finite_index: "finite I"
+
+lemma (in finite_product_sigma_finite) sigma_finite_pairs:
+  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
+    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
+    (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and>
+    (\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)"
+proof -
+  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)"
+    using M.sigma_finite_incseq by metis
+  from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
+  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>"
+    by auto
+  let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k"
+  note space_PiM[simp]
+  show ?thesis
+  proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
+    fix i show "range (F i) \<subseteq> sets (M i)" by fact
+  next
+    fix i k show "emeasure (M i) (F i k) \<noteq> \<infinity>" by fact
+  next
+    fix x assume "x \<in> (\<Union>i. ?F i)" with F(1) show "x \<in> space (PiM I M)"
+      by (auto simp: PiE_def dest!: sets.sets_into_space)
+  next
+    fix f assume "f \<in> space (PiM I M)"
+    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
+    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def PiE_def)
+  next
+    fix i show "?F i \<subseteq> ?F (Suc i)"
+      using \<open>\<And>i. incseq (F i)\<close>[THEN incseq_SucD] by auto
+  qed
+qed
+
+lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
+proof -
+  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ennreal)"
+  have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1"
+  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
+    show "positive (PiM {} M) ?\<mu>"
+      by (auto simp: positive_def)
+    show "countably_additive (PiM {} M) ?\<mu>"
+      by (rule sets.countably_additiveI_finite)
+         (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: )
+  qed (auto simp: prod_emb_def)
+  also have "(prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
+    by (auto simp: prod_emb_def)
+  finally show ?thesis
+    by simp
+qed
+
+lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
+  by (rule measure_eqI) (auto simp add: sets_PiM_empty)
+
+lemma (in product_sigma_finite) emeasure_PiM:
+  "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+proof (induct I arbitrary: A rule: finite_induct)
+  case (insert i I)
+  interpret finite_product_sigma_finite M I by standard fact
+  have "finite (insert i I)" using \<open>finite I\<close> by auto
+  interpret I': finite_product_sigma_finite M "insert i I" by standard fact
+  let ?h = "(\<lambda>(f, y). f(i := y))"
+
+  let ?P = "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h"
+  let ?\<mu> = "emeasure ?P"
+  let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}"
+  let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))"
+
+  have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) =
+    (\<Prod>i\<in>insert i I. emeasure (M i) (A i))"
+  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
+    fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))"
+    then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto
+    let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)"
+    let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^sub>E j\<in>J-{i}. E j)"
+    have "?\<mu> ?p =
+      emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i))"
+      by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
+    also have "?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
+      using J E[rule_format, THEN sets.sets_into_space]
+      by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: if_split_asm)
+    also have "emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
+      emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
+      using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto
+    also have "?p' = (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
+      using J E[rule_format, THEN sets.sets_into_space]
+      by (auto simp: prod_emb_iff PiE_def Pi_iff split: if_split_asm) blast+
+    also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
+      (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
+      using E by (subst insert) (auto intro!: setprod.cong)
+    also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) *
+       emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)"
+      using insert by (auto simp: mult.commute intro!: arg_cong2[where f="op *"] setprod.cong)
+    also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)"
+      using insert(1,2) J E by (intro setprod.mono_neutral_right) auto
+    finally show "?\<mu> ?p = \<dots>" .
+
+    show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))"
+      using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def)
+  next
+    show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>"
+      using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all
+  next
+    show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
+      insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))"
+      using insert by auto
+  qed (auto intro!: setprod.cong)
+  with insert show ?case
+    by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space)
+qed simp
+
+lemma (in product_sigma_finite) PiM_eqI:
+  assumes I[simp]: "finite I" and P: "sets P = PiM I M"
+  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+  shows "P = PiM I M"
+proof -
+  interpret finite_product_sigma_finite M I
+    proof qed fact
+  from sigma_finite_pairs guess C .. note C = this
+  show ?thesis
+  proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric])
+    show "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>M I M) (Pi\<^sub>E I A) = P (Pi\<^sub>E I A)" for A
+      by (simp add: eq emeasure_PiM)
+    define A where "A n = (\<Pi>\<^sub>E i\<in>I. C i n)" for n
+    with C show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>" "(\<Union>i. A i) = space (PiM I M)"
+      by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_setprod_eq_top)
+  qed
+qed
+
+lemma (in product_sigma_finite) sigma_finite:
+  assumes "finite I"
+  shows "sigma_finite_measure (PiM I M)"
+proof
+  interpret finite_product_sigma_finite M I by standard fact
+
+  obtain F where F: "\<And>j. countable (F j)" "\<And>j f. f \<in> F j \<Longrightarrow> f \<in> sets (M j)"
+    "\<And>j f. f \<in> F j \<Longrightarrow> emeasure (M j) f \<noteq> \<infinity>" and
+    in_space: "\<And>j. space (M j) = (\<Union>F j)"
+    using sigma_finite_countable by (metis subset_eq)
+  moreover have "(\<Union>(PiE I ` PiE I F)) = space (Pi\<^sub>M I M)"
+    using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD2])
+  ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets (Pi\<^sub>M I M) \<and> \<Union>A = space (Pi\<^sub>M I M) \<and> (\<forall>a\<in>A. emeasure (Pi\<^sub>M I M) a \<noteq> \<infinity>)"
+    by (intro exI[of _ "PiE I ` PiE I F"])
+       (auto intro!: countable_PiE sets_PiM_I_finite
+             simp: PiE_iff emeasure_PiM finite_index ennreal_setprod_eq_top)
+qed
+
+sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M"
+  using sigma_finite[OF finite_index] .
+
+lemma (in finite_product_sigma_finite) measure_times:
+  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+  using emeasure_PiM[OF finite_index] by auto
+
+lemma (in product_sigma_finite) nn_integral_empty:
+  "0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
+  by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2)
+
+lemma (in product_sigma_finite) distr_merge:
+  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
+  shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
+   (is "?D = ?P")
+proof (rule PiM_eqI)
+  interpret I: finite_product_sigma_finite M I by standard fact
+  interpret J: finite_product_sigma_finite M J by standard fact
+  fix A assume A: "\<And>i. i \<in> I \<union> J \<Longrightarrow> A i \<in> sets (M i)"
+  have *: "(merge I J -` Pi\<^sub>E (I \<union> J) A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)) = PiE I A \<times> PiE J A"
+    using A[THEN sets.sets_into_space] by (auto simp: space_PiM space_pair_measure)
+  from A fin show "emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) (Pi\<^sub>E (I \<union> J) A) =
+      (\<Prod>i\<in>I \<union> J. emeasure (M i) (A i))"
+    by (subst emeasure_distr)
+       (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times setprod.union_disjoint)
+qed (insert fin, simp_all)
+
+lemma (in product_sigma_finite) product_nn_integral_fold:
+  assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
+  and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
+  shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f =
+    (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ 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
+  interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard
+  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] by (simp add: comp_def)
+  show ?thesis
+    apply (subst distr_merge[OF IJ, symmetric])
+    apply (subst nn_integral_distr[OF measurable_merge])
+    apply measurable []
+    apply (subst J.nn_integral_fst[symmetric, OF P_borel])
+    apply simp
+    done
+qed
+
+lemma (in product_sigma_finite) distr_singleton:
+  "distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
+proof (intro measure_eqI[symmetric])
+  interpret I: finite_product_sigma_finite M "{i}" by standard simp
+  fix A assume A: "A \<in> sets (M i)"
+  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M {i} M) = (\<Pi>\<^sub>E i\<in>{i}. A)"
+    using sets.sets_into_space by (auto simp: space_PiM)
+  then show "emeasure (M i) A = emeasure ?D A"
+    using A I.measure_times[of "\<lambda>_. A"]
+    by (simp add: emeasure_distr measurable_component_singleton)
+qed simp
+
+lemma (in product_sigma_finite) product_nn_integral_singleton:
+  assumes f: "f \<in> borel_measurable (M i)"
+  shows "integral\<^sup>N (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>N (M i) f"
+proof -
+  interpret I: finite_product_sigma_finite M "{i}" by standard simp
+  from f show ?thesis
+    apply (subst distr_singleton[symmetric])
+    apply (subst nn_integral_distr[OF measurable_component_singleton])
+    apply simp_all
+    done
+qed
+
+lemma (in product_sigma_finite) product_nn_integral_insert:
+  assumes I[simp]: "finite I" "i \<notin> I"
+    and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
+  shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))"
+proof -
+  interpret I: finite_product_sigma_finite M I by standard auto
+  interpret i: finite_product_sigma_finite M "{i}" by standard auto
+  have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
+    using f by auto
+  show ?thesis
+    unfolding product_nn_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f]
+  proof (rule nn_integral_cong, subst product_nn_integral_singleton[symmetric])
+    fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
+    let ?f = "\<lambda>y. f (x(i := y))"
+    show "?f \<in> borel_measurable (M i)"
+      using measurable_comp[OF measurable_component_update f, OF x \<open>i \<notin> I\<close>]
+      unfolding comp_def .
+    show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)"
+      using x
+      by (auto intro!: nn_integral_cong arg_cong[where f=f]
+               simp add: space_PiM extensional_def PiE_def)
+  qed
+qed
+
+lemma (in product_sigma_finite) product_nn_integral_insert_rev:
+  assumes I[simp]: "finite I" "i \<notin> I"
+    and [measurable]: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
+  shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x(i := y)) \<partial>(Pi\<^sub>M I M)) \<partial>(M i))"
+  apply (subst product_nn_integral_insert[OF assms])
+  apply (rule pair_sigma_finite.Fubini')
+  apply intro_locales []
+  apply (rule sigma_finite[OF I(1)])
+  apply measurable
+  done
+
+lemma (in product_sigma_finite) product_nn_integral_setprod:
+  assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
+  shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))"
+using assms proof (induction I)
+  case (insert i I)
+  note insert.prems[measurable]
+  note \<open>finite I\<close>[intro, simp]
+  interpret I: finite_product_sigma_finite M I by standard auto
+  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 insert by (auto intro!: setprod.cong)
+  have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)"
+    using sets.sets_into_space insert
+    by (intro borel_measurable_setprod_ennreal
+              measurable_comp[OF measurable_component_singleton, unfolded comp_def])
+       auto
+  then show ?case
+    apply (simp add: product_nn_integral_insert[OF insert(1,2)])
+    apply (simp add: insert(2-) * nn_integral_multc)
+    apply (subst nn_integral_cmult)
+    apply (auto simp add: insert(2-))
+    done
+qed (simp add: space_PiM)
+
+lemma (in product_sigma_finite) product_nn_integral_pair:
+  assumes [measurable]: "case_prod f \<in> borel_measurable (M x \<Otimes>\<^sub>M M y)"
+  assumes xy: "x \<noteq> y"
+  shows "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {x, y} M) = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
+proof-
+  interpret psm: pair_sigma_finite "M x" "M y"
+    unfolding pair_sigma_finite_def using sigma_finite_measures by simp_all
+  have "{x, y} = {y, x}" by auto
+  also have "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {y, x} M) = (\<integral>\<^sup>+y. \<integral>\<^sup>+\<sigma>. f (\<sigma> x) y \<partial>PiM {x} M \<partial>M y)"
+    using xy by (subst product_nn_integral_insert_rev) simp_all
+  also have "... = (\<integral>\<^sup>+y. \<integral>\<^sup>+x. f x y \<partial>M x \<partial>M y)"
+    by (intro nn_integral_cong, subst product_nn_integral_singleton) simp_all
+  also have "... = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
+    by (subst psm.nn_integral_snd[symmetric]) simp_all
+  finally show ?thesis .
+qed
+
+lemma (in product_sigma_finite) distr_component:
+  "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
+proof (intro PiM_eqI)
+  fix A assume A: "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
+  then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i"
+    by (auto dest: sets.sets_into_space)
+  with A show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))"
+    by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict)
+qed simp_all
+
+lemma (in product_sigma_finite)
+  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
+  shows emeasure_fold_integral:
+    "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I)
+    and emeasure_fold_measurable:
+    "(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B)
+proof -
+  interpret I: finite_product_sigma_finite M I by standard fact
+  interpret J: finite_product_sigma_finite M J by standard fact
+  interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
+  have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
+    by (intro measurable_sets[OF _ A] measurable_merge assms)
+
+  show ?I
+    apply (subst distr_merge[symmetric, OF IJ])
+    apply (subst emeasure_distr[OF measurable_merge A])
+    apply (subst J.emeasure_pair_measure_alt[OF merge])
+    apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
+    done
+
+  show ?B
+    using IJ.measurable_emeasure_Pair1[OF merge]
+    by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong)
+qed
+
+lemma sets_Collect_single:
+  "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)"
+  by simp
+
+lemma pair_measure_eq_distr_PiM:
+  fixes M1 :: "'a measure" and M2 :: "'a measure"
+  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
+  shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
+    (is "?P = ?D")
+proof (rule pair_measure_eqI[OF assms])
+  interpret B: product_sigma_finite "case_bool M1 M2"
+    unfolding product_sigma_finite_def using assms by (auto split: bool.split)
+  let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)"
+
+  have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)"
+    by auto
+  fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
+  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))"
+    by (simp add: UNIV_bool ac_simps)
+  also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))"
+    using A B by (subst B.emeasure_PiM) (auto split: bool.split)
+  also have "Pi\<^sub>E UNIV (case_bool A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
+    using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space]
+    by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split)
+  finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)"
+    using A B
+      measurable_component_singleton[of True UNIV "case_bool M1 M2"]
+      measurable_component_singleton[of False UNIV "case_bool M1 M2"]
+    by (subst emeasure_distr) (auto simp: measurable_pair_iff)
+qed simp
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Interval_Integral.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,1123 @@
+(*  Title:      HOL/Probability/Interval_Integral.thy
+    Author:     Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU)
+
+Lebesgue integral over an interval (with endpoints possibly +-\<infinity>)
+*)
+
+theory Interval_Integral
+  imports Set_Integral
+begin
+
+lemma continuous_on_vector_derivative:
+  "(\<And>x. x \<in> S \<Longrightarrow> (f has_vector_derivative f' x) (at x within S)) \<Longrightarrow> continuous_on S f"
+  by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous)
+
+lemma has_vector_derivative_weaken:
+  fixes x D and f g s t
+  assumes f: "(f has_vector_derivative D) (at x within t)"
+    and "x \<in> s" "s \<subseteq> t"
+    and "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
+  shows "(g has_vector_derivative D) (at x within s)"
+proof -
+  have "(f has_vector_derivative D) (at x within s) \<longleftrightarrow> (g has_vector_derivative D) (at x within s)"
+    unfolding has_vector_derivative_def has_derivative_iff_norm
+    using assms by (intro conj_cong Lim_cong_within refl) auto
+  then show ?thesis
+    using has_vector_derivative_within_subset[OF f \<open>s \<subseteq> t\<close>] by simp
+qed
+
+definition "einterval a b = {x. a < ereal x \<and> ereal x < b}"
+
+lemma einterval_eq[simp]:
+  shows einterval_eq_Icc: "einterval (ereal a) (ereal b) = {a <..< b}"
+    and einterval_eq_Ici: "einterval (ereal a) \<infinity> = {a <..}"
+    and einterval_eq_Iic: "einterval (- \<infinity>) (ereal b) = {..< b}"
+    and einterval_eq_UNIV: "einterval (- \<infinity>) \<infinity> = UNIV"
+  by (auto simp: einterval_def)
+
+lemma einterval_same: "einterval a a = {}"
+  by (auto simp add: einterval_def)
+
+lemma einterval_iff: "x \<in> einterval a b \<longleftrightarrow> a < ereal x \<and> ereal x < b"
+  by (simp add: einterval_def)
+
+lemma einterval_nonempty: "a < b \<Longrightarrow> \<exists>c. c \<in> einterval a b"
+  by (cases a b rule: ereal2_cases, auto simp: einterval_def intro!: dense gt_ex lt_ex)
+
+lemma open_einterval[simp]: "open (einterval a b)"
+  by (cases a b rule: ereal2_cases)
+     (auto simp: einterval_def intro!: open_Collect_conj open_Collect_less continuous_intros)
+
+lemma borel_einterval[measurable]: "einterval a b \<in> sets borel"
+  unfolding einterval_def by measurable
+
+(*
+    Approximating a (possibly infinite) interval
+*)
+
+lemma filterlim_sup1: "(LIM x F. f x :> G1) \<Longrightarrow> (LIM x F. f x :> (sup G1 G2))"
+ unfolding filterlim_def by (auto intro: le_supI1)
+
+lemma ereal_incseq_approx:
+  fixes a b :: ereal
+  assumes "a < b"
+  obtains X :: "nat \<Rightarrow> real" where
+    "incseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> b"
+proof (cases b)
+  case PInf
+  with \<open>a < b\<close> have "a = -\<infinity> \<or> (\<exists>r. a = ereal r)"
+    by (cases a) auto
+  moreover have "(\<lambda>x. ereal (real (Suc x))) \<longlonglongrightarrow> \<infinity>"
+      apply (subst LIMSEQ_Suc_iff)
+      apply (simp add: Lim_PInfty)
+      using nat_ceiling_le_eq by blast
+  moreover have "\<And>r. (\<lambda>x. ereal (r + real (Suc x))) \<longlonglongrightarrow> \<infinity>"
+    apply (subst LIMSEQ_Suc_iff)
+    apply (subst Lim_PInfty)
+    apply (metis add.commute diff_le_eq nat_ceiling_le_eq ereal_less_eq(3))
+    done
+  ultimately show thesis
+    by (intro that[of "\<lambda>i. real_of_ereal a + Suc i"])
+       (auto simp: incseq_def PInf)
+next
+  case (real b')
+  define d where "d = b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)"
+  with \<open>a < b\<close> have a': "0 < d"
+    by (cases a) (auto simp: real)
+  moreover
+  have "\<And>i r. r < b' \<Longrightarrow> (b' - r) * 1 < (b' - r) * real (Suc (Suc i))"
+    by (intro mult_strict_left_mono) auto
+  with \<open>a < b\<close> a' have "\<And>i. a < ereal (b' - d / real (Suc (Suc i)))"
+    by (cases a) (auto simp: real d_def field_simps)
+  moreover have "(\<lambda>i. b' - d / Suc (Suc i)) \<longlonglongrightarrow> b'"
+    apply (subst filterlim_sequentially_Suc)
+    apply (subst filterlim_sequentially_Suc)
+    apply (rule tendsto_eq_intros)
+    apply (auto intro!: tendsto_divide_0[OF tendsto_const] filterlim_sup1
+                simp: at_infinity_eq_at_top_bot filterlim_real_sequentially)
+    done
+  ultimately show thesis
+    by (intro that[of "\<lambda>i. b' - d / Suc (Suc i)"])
+       (auto simp add: real incseq_def intro!: divide_left_mono)
+qed (insert \<open>a < b\<close>, auto)
+
+lemma ereal_decseq_approx:
+  fixes a b :: ereal
+  assumes "a < b"
+  obtains X :: "nat \<Rightarrow> real" where
+    "decseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> a"
+proof -
+  have "-b < -a" using \<open>a < b\<close> by simp
+  from ereal_incseq_approx[OF this] guess X .
+  then show thesis
+    apply (intro that[of "\<lambda>i. - X i"])
+    apply (auto simp add: uminus_ereal.simps[symmetric] decseq_def incseq_def
+                simp del: uminus_ereal.simps)
+    apply (metis ereal_minus_less_minus ereal_uminus_uminus ereal_Lim_uminus)+
+    done
+qed
+
+lemma einterval_Icc_approximation:
+  fixes a b :: ereal
+  assumes "a < b"
+  obtains u l :: "nat \<Rightarrow> real" where
+    "einterval a b = (\<Union>i. {l i .. u i})"
+    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
+    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
+proof -
+  from dense[OF \<open>a < b\<close>] obtain c where "a < c" "c < b" by safe
+  from ereal_incseq_approx[OF \<open>c < b\<close>] guess u . note u = this
+  from ereal_decseq_approx[OF \<open>a < c\<close>] guess l . note l = this
+  { fix i from less_trans[OF \<open>l i < c\<close> \<open>c < u i\<close>] have "l i < u i" by simp }
+  have "einterval a b = (\<Union>i. {l i .. u i})"
+  proof (auto simp: einterval_iff)
+    fix x assume "a < ereal x" "ereal x < b"
+    have "eventually (\<lambda>i. ereal (l i) < ereal x) sequentially"
+      using l(4) \<open>a < ereal x\<close> by (rule order_tendstoD)
+    moreover
+    have "eventually (\<lambda>i. ereal x < ereal (u i)) sequentially"
+      using u(4) \<open>ereal x< b\<close> by (rule order_tendstoD)
+    ultimately have "eventually (\<lambda>i. l i < x \<and> x < u i) sequentially"
+      by eventually_elim auto
+    then show "\<exists>i. l i \<le> x \<and> x \<le> u i"
+      by (auto intro: less_imp_le simp: eventually_sequentially)
+  next
+    fix x i assume "l i \<le> x" "x \<le> u i"
+    with \<open>a < ereal (l i)\<close> \<open>ereal (u i) < b\<close>
+    show "a < ereal x" "ereal x < b"
+      by (auto simp del: ereal_less_eq(3) simp add: ereal_less_eq(3)[symmetric])
+  qed
+  show thesis
+    by (intro that) fact+
+qed
+
+(* TODO: in this definition, it would be more natural if einterval a b included a and b when
+   they are real. *)
+definition interval_lebesgue_integral :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> 'a::{banach, second_countable_topology}" where
+  "interval_lebesgue_integral M a b f =
+    (if a \<le> b then (LINT x:einterval a b|M. f x) else - (LINT x:einterval b a|M. f x))"
+
+syntax
+  "_ascii_interval_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real measure \<Rightarrow> real \<Rightarrow> real"
+  ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60)
+
+translations
+  "LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\<lambda>x. f)"
+
+definition interval_lebesgue_integrable :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a::{banach, second_countable_topology}) \<Rightarrow> bool" where
+  "interval_lebesgue_integrable M a b f =
+    (if a \<le> b then set_integrable M (einterval a b) f else set_integrable M (einterval b a) f)"
+
+syntax
+  "_ascii_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real"
+  ("(4LBINT _=_.._. _)" [0,60,60,61] 60)
+
+translations
+  "LBINT x=a..b. f" == "CONST interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
+
+(*
+    Basic properties of integration over an interval.
+*)
+
+lemma interval_lebesgue_integral_cong:
+  "a \<le> b \<Longrightarrow> (\<And>x. x \<in> einterval a b \<Longrightarrow> f x = g x) \<Longrightarrow> einterval a b \<in> sets M \<Longrightarrow>
+    interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
+  by (auto intro: set_lebesgue_integral_cong simp: interval_lebesgue_integral_def)
+
+lemma interval_lebesgue_integral_cong_AE:
+  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
+    a \<le> b \<Longrightarrow> AE x \<in> einterval a b in M. f x = g x \<Longrightarrow> einterval a b \<in> sets M \<Longrightarrow>
+    interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
+  by (auto intro: set_lebesgue_integral_cong_AE simp: interval_lebesgue_integral_def)
+
+lemma interval_integrable_mirror:
+  shows "interval_lebesgue_integrable lborel a b (\<lambda>x. f (-x)) \<longleftrightarrow>
+    interval_lebesgue_integrable lborel (-b) (-a) f"
+proof -
+  have *: "indicator (einterval a b) (- x) = (indicator (einterval (-b) (-a)) x :: real)"
+    for a b :: ereal and x :: real
+    by (cases a b rule: ereal2_cases) (auto simp: einterval_def split: split_indicator)
+  show ?thesis
+    unfolding interval_lebesgue_integrable_def
+    using lborel_integrable_real_affine_iff[symmetric, of "-1" "\<lambda>x. indicator (einterval _ _) x *\<^sub>R f x" 0]
+    by (simp add: *)
+qed
+
+lemma interval_lebesgue_integral_add [intro, simp]:
+  fixes M a b f
+  assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g"
+  shows "interval_lebesgue_integrable M a b (\<lambda>x. f x + g x)" and
+    "interval_lebesgue_integral M a b (\<lambda>x. f x + g x) =
+   interval_lebesgue_integral M a b f + interval_lebesgue_integral M a b g"
+using assms by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def
+    field_simps)
+
+lemma interval_lebesgue_integral_diff [intro, simp]:
+  fixes M a b f
+  assumes "interval_lebesgue_integrable M a b f"
+    "interval_lebesgue_integrable M a b g"
+  shows "interval_lebesgue_integrable M a b (\<lambda>x. f x - g x)" and
+    "interval_lebesgue_integral M a b (\<lambda>x. f x - g x) =
+   interval_lebesgue_integral M a b f - interval_lebesgue_integral M a b g"
+using assms by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def
+    field_simps)
+
+lemma interval_lebesgue_integrable_mult_right [intro, simp]:
+  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
+  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
+    interval_lebesgue_integrable M a b (\<lambda>x. c * f x)"
+  by (simp add: interval_lebesgue_integrable_def)
+
+lemma interval_lebesgue_integrable_mult_left [intro, simp]:
+  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
+  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
+    interval_lebesgue_integrable M a b (\<lambda>x. f x * c)"
+  by (simp add: interval_lebesgue_integrable_def)
+
+lemma interval_lebesgue_integrable_divide [intro, simp]:
+  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
+  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
+    interval_lebesgue_integrable M a b (\<lambda>x. f x / c)"
+  by (simp add: interval_lebesgue_integrable_def)
+
+lemma interval_lebesgue_integral_mult_right [simp]:
+  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
+  shows "interval_lebesgue_integral M a b (\<lambda>x. c * f x) =
+    c * interval_lebesgue_integral M a b f"
+  by (simp add: interval_lebesgue_integral_def)
+
+lemma interval_lebesgue_integral_mult_left [simp]:
+  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
+  shows "interval_lebesgue_integral M a b (\<lambda>x. f x * c) =
+    interval_lebesgue_integral M a b f * c"
+  by (simp add: interval_lebesgue_integral_def)
+
+lemma interval_lebesgue_integral_divide [simp]:
+  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
+  shows "interval_lebesgue_integral M a b (\<lambda>x. f x / c) =
+    interval_lebesgue_integral M a b f / c"
+  by (simp add: interval_lebesgue_integral_def)
+
+lemma interval_lebesgue_integral_uminus:
+  "interval_lebesgue_integral M a b (\<lambda>x. - f x) = - interval_lebesgue_integral M a b f"
+  by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def)
+
+lemma interval_lebesgue_integral_of_real:
+  "interval_lebesgue_integral M a b (\<lambda>x. complex_of_real (f x)) =
+    of_real (interval_lebesgue_integral M a b f)"
+  unfolding interval_lebesgue_integral_def
+  by (auto simp add: interval_lebesgue_integral_def set_integral_complex_of_real)
+
+lemma interval_lebesgue_integral_le_eq:
+  fixes a b f
+  assumes "a \<le> b"
+  shows "interval_lebesgue_integral M a b f = (LINT x : einterval a b | M. f x)"
+using assms by (auto simp add: interval_lebesgue_integral_def)
+
+lemma interval_lebesgue_integral_gt_eq:
+  fixes a b f
+  assumes "a > b"
+  shows "interval_lebesgue_integral M a b f = -(LINT x : einterval b a | M. f x)"
+using assms by (auto simp add: interval_lebesgue_integral_def less_imp_le einterval_def)
+
+lemma interval_lebesgue_integral_gt_eq':
+  fixes a b f
+  assumes "a > b"
+  shows "interval_lebesgue_integral M a b f = - interval_lebesgue_integral M b a f"
+using assms by (auto simp add: interval_lebesgue_integral_def less_imp_le einterval_def)
+
+lemma interval_integral_endpoints_same [simp]: "(LBINT x=a..a. f x) = 0"
+  by (simp add: interval_lebesgue_integral_def einterval_same)
+
+lemma interval_integral_endpoints_reverse: "(LBINT x=a..b. f x) = -(LBINT x=b..a. f x)"
+  by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integral_def einterval_same)
+
+lemma interval_integrable_endpoints_reverse:
+  "interval_lebesgue_integrable lborel a b f \<longleftrightarrow>
+    interval_lebesgue_integrable lborel b a f"
+  by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integrable_def einterval_same)
+
+lemma interval_integral_reflect:
+  "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))"
+proof (induct a b rule: linorder_wlog)
+  case (sym a b) then show ?case
+    by (auto simp add: interval_lebesgue_integral_def interval_integrable_endpoints_reverse
+             split: if_split_asm)
+next
+  case (le a b) then show ?case
+    unfolding interval_lebesgue_integral_def
+    by (subst set_integral_reflect)
+       (auto simp: interval_lebesgue_integrable_def einterval_iff
+                   ereal_uminus_le_reorder ereal_uminus_less_reorder not_less
+                   uminus_ereal.simps[symmetric]
+             simp del: uminus_ereal.simps
+             intro!: integral_cong
+             split: split_indicator)
+qed
+
+lemma interval_lebesgue_integral_0_infty:
+  "interval_lebesgue_integrable M 0 \<infinity> f \<longleftrightarrow> set_integrable M {0<..} f"
+  "interval_lebesgue_integral M 0 \<infinity> f = (LINT x:{0<..}|M. f x)"
+  unfolding zero_ereal_def
+  by (auto simp: interval_lebesgue_integral_le_eq interval_lebesgue_integrable_def)
+
+lemma interval_integral_to_infinity_eq: "(LINT x=ereal a..\<infinity> | M. f x) = (LINT x : {a<..} | M. f x)"
+  unfolding interval_lebesgue_integral_def by auto
+
+lemma interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) =
+  (set_integrable M {a<..} f)"
+  unfolding interval_lebesgue_integrable_def by auto
+
+(*
+    Basic properties of integration over an interval wrt lebesgue measure.
+*)
+
+lemma interval_integral_zero [simp]:
+  fixes a b :: ereal
+  shows"LBINT x=a..b. 0 = 0"
+unfolding interval_lebesgue_integral_def einterval_eq
+by simp
+
+lemma interval_integral_const [intro, simp]:
+  fixes a b c :: real
+  shows "interval_lebesgue_integrable lborel a b (\<lambda>x. c)" and "LBINT x=a..b. c = c * (b - a)"
+unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq
+by (auto simp add:  less_imp_le field_simps measure_def)
+
+lemma interval_integral_cong_AE:
+  assumes [measurable]: "f \<in> borel_measurable borel" "g \<in> borel_measurable borel"
+  assumes "AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x"
+  shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
+  using assms
+proof (induct a b rule: linorder_wlog)
+  case (sym a b) then show ?case
+    by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
+next
+  case (le a b) then show ?case
+    by (auto simp: interval_lebesgue_integral_def max_def min_def
+             intro!: set_lebesgue_integral_cong_AE)
+qed
+
+lemma interval_integral_cong:
+  assumes "\<And>x. x \<in> einterval (min a b) (max a b) \<Longrightarrow> f x = g x"
+  shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
+  using assms
+proof (induct a b rule: linorder_wlog)
+  case (sym a b) then show ?case
+    by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
+next
+  case (le a b) then show ?case
+    by (auto simp: interval_lebesgue_integral_def max_def min_def
+             intro!: set_lebesgue_integral_cong)
+qed
+
+lemma interval_lebesgue_integrable_cong_AE:
+    "f \<in> borel_measurable lborel \<Longrightarrow> g \<in> borel_measurable lborel \<Longrightarrow>
+    AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x \<Longrightarrow>
+    interval_lebesgue_integrable lborel a b f = interval_lebesgue_integrable lborel a b g"
+  apply (simp add: interval_lebesgue_integrable_def )
+  apply (intro conjI impI set_integrable_cong_AE)
+  apply (auto simp: min_def max_def)
+  done
+
+lemma interval_integrable_abs_iff:
+  fixes f :: "real \<Rightarrow> real"
+  shows  "f \<in> borel_measurable lborel \<Longrightarrow>
+    interval_lebesgue_integrable lborel a b (\<lambda>x. \<bar>f x\<bar>) = interval_lebesgue_integrable lborel a b f"
+  unfolding interval_lebesgue_integrable_def
+  by (subst (1 2) set_integrable_abs_iff') simp_all
+
+lemma interval_integral_Icc:
+  fixes a b :: real
+  shows "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a..b}. f x)"
+  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
+           simp add: interval_lebesgue_integral_def)
+
+lemma interval_integral_Icc':
+  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {x. a \<le> ereal x \<and> ereal x \<le> b}. f x)"
+  by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"]
+           simp add: interval_lebesgue_integral_def einterval_iff)
+
+lemma interval_integral_Ioc:
+  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a<..b}. f x)"
+  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
+           simp add: interval_lebesgue_integral_def einterval_iff)
+
+(* TODO: other versions as well? *) (* Yes: I need the Icc' version. *)
+lemma interval_integral_Ioc':
+  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {x. a < ereal x \<and> ereal x \<le> b}. f x)"
+  by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"]
+           simp add: interval_lebesgue_integral_def einterval_iff)
+
+lemma interval_integral_Ico:
+  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a..<b}. f x)"
+  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
+           simp add: interval_lebesgue_integral_def einterval_iff)
+
+lemma interval_integral_Ioi:
+  "\<bar>a\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..\<infinity>. f x) = (LBINT x : {real_of_ereal a <..}. f x)"
+  by (auto simp add: interval_lebesgue_integral_def einterval_iff)
+
+lemma interval_integral_Ioo:
+  "a \<le> b \<Longrightarrow> \<bar>a\<bar> < \<infinity> ==> \<bar>b\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {real_of_ereal a <..< real_of_ereal b}. f x)"
+  by (auto simp add: interval_lebesgue_integral_def einterval_iff)
+
+lemma interval_integral_discrete_difference:
+  fixes f :: "real \<Rightarrow> 'b::{banach, second_countable_topology}" and a b :: ereal
+  assumes "countable X"
+  and eq: "\<And>x. a \<le> b \<Longrightarrow> a < x \<Longrightarrow> x < b \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+  and anti_eq: "\<And>x. b \<le> a \<Longrightarrow> b < x \<Longrightarrow> x < a \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+  shows "interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
+  unfolding interval_lebesgue_integral_def
+  apply (intro if_cong refl arg_cong[where f="\<lambda>x. - x"] integral_discrete_difference[of X] assms)
+  apply (auto simp: eq anti_eq einterval_iff split: split_indicator)
+  done
+
+lemma interval_integral_sum:
+  fixes a b c :: ereal
+  assumes integrable: "interval_lebesgue_integrable lborel (min a (min b c)) (max a (max b c)) f"
+  shows "(LBINT x=a..b. f x) + (LBINT x=b..c. f x) = (LBINT x=a..c. f x)"
+proof -
+  let ?I = "\<lambda>a b. LBINT x=a..b. f x"
+  { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
+    then have ord: "a \<le> b" "b \<le> c" "a \<le> c" and f': "set_integrable lborel (einterval a c) f"
+      by (auto simp: interval_lebesgue_integrable_def)
+    then have f: "set_borel_measurable borel (einterval a c) f"
+      by (drule_tac borel_measurable_integrable) simp
+    have "(LBINT x:einterval a c. f x) = (LBINT x:einterval a b \<union> einterval b c. f x)"
+    proof (rule set_integral_cong_set)
+      show "AE x in lborel. (x \<in> einterval a b \<union> einterval b c) = (x \<in> einterval a c)"
+        using AE_lborel_singleton[of "real_of_ereal b"] ord
+        by (cases a b c rule: ereal3_cases) (auto simp: einterval_iff)
+    qed (insert ord, auto intro!: set_borel_measurable_subset[OF f] simp: einterval_iff)
+    also have "\<dots> = (LBINT x:einterval a b. f x) + (LBINT x:einterval b c. f x)"
+      using ord
+      by (intro set_integral_Un_AE) (auto intro!: set_integrable_subset[OF f'] simp: einterval_iff not_less)
+    finally have "?I a b + ?I b c = ?I a c"
+      using ord by (simp add: interval_lebesgue_integral_def)
+  } note 1 = this
+  { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
+    from 1[OF this] have "?I b c + ?I a b = ?I a c"
+      by (metis add.commute)
+  } note 2 = this
+  have 3: "\<And>a b. b \<le> a \<Longrightarrow> (LBINT x=a..b. f x) = - (LBINT x=b..a. f x)"
+    by (rule interval_integral_endpoints_reverse)
+  show ?thesis
+    using integrable
+    by (cases a b b c a c rule: linorder_le_cases[case_product linorder_le_cases linorder_cases])
+       (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3)
+qed
+
+lemma interval_integrable_isCont:
+  fixes a b and f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
+  shows "(\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> isCont f x) \<Longrightarrow>
+    interval_lebesgue_integrable lborel a b f"
+proof (induct a b rule: linorder_wlog)
+  case (le a b) then show ?case
+    by (auto simp: interval_lebesgue_integrable_def
+             intro!: set_integrable_subset[OF borel_integrable_compact[of "{a .. b}"]]
+                     continuous_at_imp_continuous_on)
+qed (auto intro: interval_integrable_endpoints_reverse[THEN iffD1])
+
+lemma interval_integrable_continuous_on:
+  fixes a b :: real and f
+  assumes "a \<le> b" and "continuous_on {a..b} f"
+  shows "interval_lebesgue_integrable lborel a b f"
+using assms unfolding interval_lebesgue_integrable_def apply simp
+  by (rule set_integrable_subset, rule borel_integrable_atLeastAtMost' [of a b], auto)
+
+lemma interval_integral_eq_integral:
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
+  shows "a \<le> b \<Longrightarrow> set_integrable lborel {a..b} f \<Longrightarrow> LBINT x=a..b. f x = integral {a..b} f"
+  by (subst interval_integral_Icc, simp) (rule set_borel_integral_eq_integral)
+
+lemma interval_integral_eq_integral':
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
+  shows "a \<le> b \<Longrightarrow> set_integrable lborel (einterval a b) f \<Longrightarrow> LBINT x=a..b. f x = integral (einterval a b) f"
+  by (subst interval_lebesgue_integral_le_eq, simp) (rule set_borel_integral_eq_integral)
+
+(*
+    General limit approximation arguments
+*)
+
+lemma interval_integral_Icc_approx_nonneg:
+  fixes a b :: ereal
+  assumes "a < b"
+  fixes u l :: "nat \<Rightarrow> real"
+  assumes  approx: "einterval a b = (\<Union>i. {l i .. u i})"
+    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
+    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
+  fixes f :: "real \<Rightarrow> real"
+  assumes f_integrable: "\<And>i. set_integrable lborel {l i..u i} f"
+  assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x"
+  assumes f_measurable: "set_borel_measurable lborel (einterval a b) f"
+  assumes lbint_lim: "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> C"
+  shows
+    "set_integrable lborel (einterval a b) f"
+    "(LBINT x=a..b. f x) = C"
+proof -
+  have 1: "\<And>i. set_integrable lborel {l i..u i} f" by (rule f_integrable)
+  have 2: "AE x in lborel. mono (\<lambda>n. indicator {l n..u n} x *\<^sub>R f x)"
+  proof -
+     from f_nonneg have "AE x in lborel. \<forall>i. l i \<le> x \<longrightarrow> x \<le> u i \<longrightarrow> 0 \<le> f x"
+      by eventually_elim
+         (metis approx(5) approx(6) dual_order.strict_trans1 ereal_less_eq(3) le_less_trans)
+    then show ?thesis
+      apply eventually_elim
+      apply (auto simp: mono_def split: split_indicator)
+      apply (metis approx(3) decseqD order_trans)
+      apply (metis approx(2) incseqD order_trans)
+      done
+  qed
+  have 3: "AE x in lborel. (\<lambda>i. indicator {l i..u i} x *\<^sub>R f x) \<longlonglongrightarrow> indicator (einterval a b) x *\<^sub>R f x"
+  proof -
+    { fix x i assume "l i \<le> x" "x \<le> u i"
+      then have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
+        apply (auto simp: eventually_sequentially intro!: exI[of _ i])
+        apply (metis approx(3) decseqD order_trans)
+        apply (metis approx(2) incseqD order_trans)
+        done
+      then have "eventually (\<lambda>i. f x * indicator {l i..u i} x = f x) sequentially"
+        by eventually_elim auto }
+    then show ?thesis
+      unfolding approx(1) by (auto intro!: AE_I2 Lim_eventually split: split_indicator)
+  qed
+  have 4: "(\<lambda>i. \<integral> x. indicator {l i..u i} x *\<^sub>R f x \<partial>lborel) \<longlonglongrightarrow> C"
+    using lbint_lim by (simp add: interval_integral_Icc approx less_imp_le)
+  have 5: "set_borel_measurable lborel (einterval a b) f" by (rule assms)
+  have "(LBINT x=a..b. f x) = lebesgue_integral lborel (\<lambda>x. indicator (einterval a b) x *\<^sub>R f x)"
+    using assms by (simp add: interval_lebesgue_integral_def less_imp_le)
+  also have "... = C" by (rule integral_monotone_convergence [OF 1 2 3 4 5])
+  finally show "(LBINT x=a..b. f x) = C" .
+
+  show "set_integrable lborel (einterval a b) f"
+    by (rule integrable_monotone_convergence[OF 1 2 3 4 5])
+qed
+
+lemma interval_integral_Icc_approx_integrable:
+  fixes u l :: "nat \<Rightarrow> real" and a b :: ereal
+  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
+  assumes "a < b"
+  assumes  approx: "einterval a b = (\<Union>i. {l i .. u i})"
+    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
+    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
+  assumes f_integrable: "set_integrable lborel (einterval a b) f"
+  shows "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
+proof -
+  have "(\<lambda>i. LBINT x:{l i.. u i}. f x) \<longlonglongrightarrow> (LBINT x:einterval a b. f x)"
+  proof (rule integral_dominated_convergence)
+    show "integrable lborel (\<lambda>x. norm (indicator (einterval a b) x *\<^sub>R f x))"
+      by (rule integrable_norm) fact
+    show "set_borel_measurable lborel (einterval a b) f"
+      using f_integrable by (rule borel_measurable_integrable)
+    then show "\<And>i. set_borel_measurable lborel {l i..u i} f"
+      by (rule set_borel_measurable_subset) (auto simp: approx)
+    show "\<And>i. AE x in lborel. norm (indicator {l i..u i} x *\<^sub>R f x) \<le> norm (indicator (einterval a b) x *\<^sub>R f x)"
+      by (intro AE_I2) (auto simp: approx split: split_indicator)
+    show "AE x in lborel. (\<lambda>i. indicator {l i..u i} x *\<^sub>R f x) \<longlonglongrightarrow> indicator (einterval a b) x *\<^sub>R f x"
+    proof (intro AE_I2 tendsto_intros Lim_eventually)
+      fix x
+      { fix i assume "l i \<le> x" "x \<le> u i"
+        with \<open>incseq u\<close>[THEN incseqD, of i] \<open>decseq l\<close>[THEN decseqD, of i]
+        have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
+          by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) }
+      then show "eventually (\<lambda>xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially"
+        using approx order_tendstoD(2)[OF \<open>l \<longlonglongrightarrow> a\<close>, of x] order_tendstoD(1)[OF \<open>u \<longlonglongrightarrow> b\<close>, of x]
+        by (auto split: split_indicator)
+    qed
+  qed
+  with \<open>a < b\<close> \<open>\<And>i. l i < u i\<close> show ?thesis
+    by (simp add: interval_lebesgue_integral_le_eq[symmetric] interval_integral_Icc less_imp_le)
+qed
+
+(*
+  A slightly stronger version of integral_FTC_atLeastAtMost and related facts,
+  with continuous_on instead of isCont
+
+  TODO: make the older versions corollaries of these (using continuous_at_imp_continuous_on, etc.)
+*)
+
+(*
+TODO: many proofs below require inferences like
+
+  a < ereal x \<Longrightarrow> x < y \<Longrightarrow> a < ereal y
+
+where x and y are real. These should be better automated.
+*)
+
+(*
+    The first Fundamental Theorem of Calculus
+
+    First, for finite intervals, and then two versions for arbitrary intervals.
+*)
+
+lemma interval_integral_FTC_finite:
+  fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: real
+  assumes f: "continuous_on {min a b..max a b} f"
+  assumes F: "\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> (F has_vector_derivative (f x)) (at x within
+    {min a b..max a b})"
+  shows "(LBINT x=a..b. f x) = F b - F a"
+  apply (case_tac "a \<le> b")
+  apply (subst interval_integral_Icc, simp)
+  apply (rule integral_FTC_atLeastAtMost, assumption)
+  apply (metis F max_def min_def)
+  using f apply (simp add: min_absorb1 max_absorb2)
+  apply (subst interval_integral_endpoints_reverse)
+  apply (subst interval_integral_Icc, simp)
+  apply (subst integral_FTC_atLeastAtMost, auto)
+  apply (metis F max_def min_def)
+using f by (simp add: min_absorb2 max_absorb1)
+
+lemma interval_integral_FTC_nonneg:
+  fixes f F :: "real \<Rightarrow> real" and a b :: ereal
+  assumes "a < b"
+  assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV F x :> f x"
+  assumes f: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f x"
+  assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x"
+  assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
+  assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
+  shows
+    "set_integrable lborel (einterval a b) f"
+    "(LBINT x=a..b. f x) = B - A"
+proof -
+  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx = this
+  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
+    by (rule order_less_le_trans, rule approx, force)
+  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
+    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
+  have FTCi: "\<And>i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)"
+    using assms approx apply (intro interval_integral_FTC_finite)
+    apply (auto simp add: less_imp_le min_def max_def
+      has_field_derivative_iff_has_vector_derivative[symmetric])
+    apply (rule continuous_at_imp_continuous_on, auto intro!: f)
+    by (rule DERIV_subset [OF F], auto)
+  have 1: "\<And>i. set_integrable lborel {l i..u i} f"
+  proof -
+    fix i show "set_integrable lborel {l i .. u i} f"
+      using \<open>a < l i\<close> \<open>u i < b\<close>
+      by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI)
+         (auto simp del: ereal_less_eq simp add: ereal_less_eq(3)[symmetric])
+  qed
+  have 2: "set_borel_measurable lborel (einterval a b) f"
+    by (auto simp del: real_scaleR_def intro!: set_borel_measurable_continuous
+             simp: continuous_on_eq_continuous_at einterval_iff f)
+  have 3: "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A"
+    apply (subst FTCi)
+    apply (intro tendsto_intros)
+    using B approx unfolding tendsto_at_iff_sequentially comp_def
+    using tendsto_at_iff_sequentially[where 'a=real]
+    apply (elim allE[of _ "\<lambda>i. ereal (u i)"], auto)
+    using A approx unfolding tendsto_at_iff_sequentially comp_def
+    by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
+  show "(LBINT x=a..b. f x) = B - A"
+    by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
+  show "set_integrable lborel (einterval a b) f"
+    by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
+qed
+
+lemma interval_integral_FTC_integrable:
+  fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: ereal
+  assumes "a < b"
+  assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> (F has_vector_derivative f x) (at x)"
+  assumes f: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f x"
+  assumes f_integrable: "set_integrable lborel (einterval a b) f"
+  assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
+  assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
+  shows "(LBINT x=a..b. f x) = B - A"
+proof -
+  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx = this
+  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
+    by (rule order_less_le_trans, rule approx, force)
+  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
+    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
+  have FTCi: "\<And>i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)"
+    using assms approx
+    by (auto simp add: less_imp_le min_def max_def
+             intro!: f continuous_at_imp_continuous_on interval_integral_FTC_finite
+             intro: has_vector_derivative_at_within)
+  have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A"
+    apply (subst FTCi)
+    apply (intro tendsto_intros)
+    using B approx unfolding tendsto_at_iff_sequentially comp_def
+    apply (elim allE[of _ "\<lambda>i. ereal (u i)"], auto)
+    using A approx unfolding tendsto_at_iff_sequentially comp_def
+    by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
+  moreover have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
+    by (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx f_integrable])
+  ultimately show ?thesis
+    by (elim LIMSEQ_unique)
+qed
+
+(*
+  The second Fundamental Theorem of Calculus and existence of antiderivatives on an
+  einterval.
+*)
+
+lemma interval_integral_FTC2:
+  fixes a b c :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
+  assumes "a \<le> c" "c \<le> b"
+  and contf: "continuous_on {a..b} f"
+  fixes x :: real
+  assumes "a \<le> x" and "x \<le> b"
+  shows "((\<lambda>u. LBINT y=c..u. f y) has_vector_derivative (f x)) (at x within {a..b})"
+proof -
+  let ?F = "(\<lambda>u. LBINT y=a..u. f y)"
+  have intf: "set_integrable lborel {a..b} f"
+    by (rule borel_integrable_atLeastAtMost', rule contf)
+  have "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
+    apply (intro integral_has_vector_derivative)
+    using \<open>a \<le> x\<close> \<open>x \<le> b\<close> by (intro continuous_on_subset [OF contf], auto)
+  then have "((\<lambda>u. integral {a..u} f) has_vector_derivative (f x)) (at x within {a..b})"
+    by simp
+  then have "(?F has_vector_derivative (f x)) (at x within {a..b})"
+    by (rule has_vector_derivative_weaken)
+       (auto intro!: assms interval_integral_eq_integral[symmetric] set_integrable_subset [OF intf])
+  then have "((\<lambda>x. (LBINT y=c..a. f y) + ?F x) has_vector_derivative (f x)) (at x within {a..b})"
+    by (auto intro!: derivative_eq_intros)
+  then show ?thesis
+  proof (rule has_vector_derivative_weaken)
+    fix u assume "u \<in> {a .. b}"
+    then show "(LBINT y=c..a. f y) + (LBINT y=a..u. f y) = (LBINT y=c..u. f y)"
+      using assms
+      apply (intro interval_integral_sum)
+      apply (auto simp add: interval_lebesgue_integrable_def simp del: real_scaleR_def)
+      by (rule set_integrable_subset [OF intf], auto simp add: min_def max_def)
+  qed (insert assms, auto)
+qed
+
+lemma einterval_antiderivative:
+  fixes a b :: ereal and f :: "real \<Rightarrow> 'a::euclidean_space"
+  assumes "a < b" and contf: "\<And>x :: real. a < x \<Longrightarrow> x < b \<Longrightarrow> isCont f x"
+  shows "\<exists>F. \<forall>x :: real. a < x \<longrightarrow> x < b \<longrightarrow> (F has_vector_derivative f x) (at x)"
+proof -
+  from einterval_nonempty [OF \<open>a < b\<close>] obtain c :: real where [simp]: "a < c" "c < b"
+    by (auto simp add: einterval_def)
+  let ?F = "(\<lambda>u. LBINT y=c..u. f y)"
+  show ?thesis
+  proof (rule exI, clarsimp)
+    fix x :: real
+    assume [simp]: "a < x" "x < b"
+    have 1: "a < min c x" by simp
+    from einterval_nonempty [OF 1] obtain d :: real where [simp]: "a < d" "d < c" "d < x"
+      by (auto simp add: einterval_def)
+    have 2: "max c x < b" by simp
+    from einterval_nonempty [OF 2] obtain e :: real where [simp]: "c < e" "x < e" "e < b"
+      by (auto simp add: einterval_def)
+    show "(?F has_vector_derivative f x) (at x)"
+      (* TODO: factor out the next three lines to has_field_derivative_within_open *)
+      unfolding has_vector_derivative_def
+      apply (subst has_derivative_within_open [of _ "{d<..<e}", symmetric], auto)
+      apply (subst has_vector_derivative_def [symmetric])
+      apply (rule has_vector_derivative_within_subset [of _ _ _ "{d..e}"])
+      apply (rule interval_integral_FTC2, auto simp add: less_imp_le)
+      apply (rule continuous_at_imp_continuous_on)
+      apply (auto intro!: contf)
+      apply (rule order_less_le_trans, rule \<open>a < d\<close>, auto)
+      apply (rule order_le_less_trans) prefer 2
+      by (rule \<open>e < b\<close>, auto)
+  qed
+qed
+
+(*
+    The substitution theorem
+
+    Once again, three versions: first, for finite intervals, and then two versions for
+    arbitrary intervals.
+*)
+
+lemma interval_integral_substitution_finite:
+  fixes a b :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
+  assumes "a \<le> b"
+  and derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_real_derivative (g' x)) (at x within {a..b})"
+  and contf : "continuous_on (g ` {a..b}) f"
+  and contg': "continuous_on {a..b} g'"
+  shows "LBINT x=a..b. g' x *\<^sub>R f (g x) = LBINT y=g a..g b. f y"
+proof-
+  have v_derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_vector_derivative (g' x)) (at x within {a..b})"
+    using derivg unfolding has_field_derivative_iff_has_vector_derivative .
+  then have contg [simp]: "continuous_on {a..b} g"
+    by (rule continuous_on_vector_derivative) auto
+  have 1: "\<And>u. min (g a) (g b) \<le> u \<Longrightarrow> u \<le> max (g a) (g b) \<Longrightarrow>
+      \<exists>x\<in>{a..b}. u = g x"
+    apply (case_tac "g a \<le> g b")
+    apply (auto simp add: min_def max_def less_imp_le)
+    apply (frule (1) IVT' [of g], auto simp add: assms)
+    by (frule (1) IVT2' [of g], auto simp add: assms)
+  from contg \<open>a \<le> b\<close> have "\<exists>c d. g ` {a..b} = {c..d} \<and> c \<le> d"
+    by (elim continuous_image_closed_interval)
+  then obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \<le> d" by auto
+  have "\<exists>F. \<forall>x\<in>{a..b}. (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))"
+    apply (rule exI, auto, subst g_im)
+    apply (rule interval_integral_FTC2 [of c c d])
+    using \<open>c \<le> d\<close> apply auto
+    apply (rule continuous_on_subset [OF contf])
+    using g_im by auto
+  then guess F ..
+  then have derivF: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow>
+    (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))" by auto
+  have contf2: "continuous_on {min (g a) (g b)..max (g a) (g b)} f"
+    apply (rule continuous_on_subset [OF contf])
+    apply (auto simp add: image_def)
+    by (erule 1)
+  have contfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
+    by (blast intro: continuous_on_compose2 contf contg)
+  have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)"
+    apply (subst interval_integral_Icc, simp add: assms)
+    apply (rule integral_FTC_atLeastAtMost[of a b "\<lambda>x. F (g x)", OF \<open>a \<le> b\<close>])
+    apply (rule vector_diff_chain_within[OF v_derivg derivF, unfolded comp_def])
+    apply (auto intro!: continuous_on_scaleR contg' contfg)
+    done
+  moreover have "LBINT y=(g a)..(g b). f y = F (g b) - F (g a)"
+    apply (rule interval_integral_FTC_finite)
+    apply (rule contf2)
+    apply (frule (1) 1, auto)
+    apply (rule has_vector_derivative_within_subset [OF derivF])
+    apply (auto simp add: image_def)
+    by (rule 1, auto)
+  ultimately show ?thesis by simp
+qed
+
+(* TODO: is it possible to lift the assumption here that g' is nonnegative? *)
+
+lemma interval_integral_substitution_integrable:
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and a b u v :: ereal
+  assumes "a < b"
+  and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
+  and contf: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f (g x)"
+  and contg': "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont g' x"
+  and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x"
+  and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
+  and B: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
+  and integrable: "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
+  and integrable2: "set_integrable lborel (einterval A B) (\<lambda>x. f x)"
+  shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
+proof -
+  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx [simp] = this
+  note less_imp_le [simp]
+  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
+    by (rule order_less_le_trans, rule approx, force)
+  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
+    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
+  have [simp]: "\<And>i. l i < b"
+    apply (rule order_less_trans) prefer 2
+    by (rule approx, auto, rule approx)
+  have [simp]: "\<And>i. a < u i"
+    by (rule order_less_trans, rule approx, auto, rule approx)
+  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> l j \<le> l i" by (rule decseqD, rule approx)
+  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> u i \<le> u j" by (rule incseqD, rule approx)
+  have g_nondec [simp]: "\<And>x y. a < x \<Longrightarrow> x \<le> y \<Longrightarrow> y < b \<Longrightarrow> g x \<le> g y"
+    apply (erule DERIV_nonneg_imp_nondecreasing, auto)
+    apply (rule exI, rule conjI, rule deriv_g)
+    apply (erule order_less_le_trans, auto)
+    apply (rule order_le_less_trans, subst ereal_less_eq(3), assumption, auto)
+    apply (rule g'_nonneg)
+    apply (rule less_imp_le, erule order_less_le_trans, auto)
+    by (rule less_imp_le, rule le_less_trans, subst ereal_less_eq(3), assumption, auto)
+  have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
+  proof -
+    have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A"
+      using A apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
+      by (drule_tac x = "\<lambda>i. ereal (l i)" in spec, auto)
+    hence A3: "\<And>i. g (l i) \<ge> A"
+      by (intro decseq_le, auto simp add: decseq_def)
+    have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B"
+      using B apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
+      by (drule_tac x = "\<lambda>i. ereal (u i)" in spec, auto)
+    hence B3: "\<And>i. g (u i) \<le> B"
+      by (intro incseq_le, auto simp add: incseq_def)
+    show "A \<le> B"
+      apply (rule order_trans [OF A3 [of 0]])
+      apply (rule order_trans [OF _ B3 [of 0]])
+      by auto
+    { fix x :: real
+      assume "A < x" and "x < B"
+      then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially"
+        apply (intro eventually_conj order_tendstoD)
+        by (rule A2, assumption, rule B2, assumption)
+      hence "\<exists>i. g (l i) < x \<and> x < g (u i)"
+        by (simp add: eventually_sequentially, auto)
+    } note AB = this
+    show "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
+      apply (auto simp add: einterval_def)
+      apply (erule (1) AB)
+      apply (rule order_le_less_trans, rule A3, simp)
+      apply (rule order_less_le_trans) prefer 2
+      by (rule B3, simp)
+  qed
+  (* finally, the main argument *)
+  {
+     fix i
+     have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)"
+        apply (rule interval_integral_substitution_finite, auto)
+        apply (rule DERIV_subset)
+        unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
+        apply (rule deriv_g)
+        apply (auto intro!: continuous_at_imp_continuous_on contf contg')
+        done
+  } note eq1 = this
+  have "(\<lambda>i. LBINT x=l i..u i. g' x *\<^sub>R f (g x)) \<longlonglongrightarrow> (LBINT x=a..b. g' x *\<^sub>R f (g x))"
+    apply (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx])
+    by (rule assms)
+  hence 2: "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x=a..b. g' x *\<^sub>R f (g x))"
+    by (simp add: eq1)
+  have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
+    apply (auto simp add: incseq_def)
+    apply (rule order_le_less_trans)
+    prefer 2 apply (assumption, auto)
+    by (erule order_less_le_trans, rule g_nondec, auto)
+  have "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x = A..B. f x)"
+    apply (subst interval_lebesgue_integral_le_eq, auto simp del: real_scaleR_def)
+    apply (subst interval_lebesgue_integral_le_eq, rule \<open>A \<le> B\<close>)
+    apply (subst un, rule set_integral_cont_up, auto simp del: real_scaleR_def)
+    apply (rule incseq)
+    apply (subst un [symmetric])
+    by (rule integrable2)
+  thus ?thesis by (intro LIMSEQ_unique [OF _ 2])
+qed
+
+(* TODO: the last two proofs are only slightly different. Factor out common part?
+   An alternative: make the second one the main one, and then have another lemma
+   that says that if f is nonnegative and all the other hypotheses hold, then it is integrable. *)
+
+lemma interval_integral_substitution_nonneg:
+  fixes f g g':: "real \<Rightarrow> real" and a b u v :: ereal
+  assumes "a < b"
+  and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
+  and contf: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f (g x)"
+  and contg': "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont g' x"
+  and f_nonneg: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> 0 \<le> f (g x)" (* TODO: make this AE? *)
+  and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x"
+  and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
+  and B: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
+  and integrable_fg: "set_integrable lborel (einterval a b) (\<lambda>x. f (g x) * g' x)"
+  shows
+    "set_integrable lborel (einterval A B) f"
+    "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
+proof -
+  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx [simp] = this
+  note less_imp_le [simp]
+  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
+    by (rule order_less_le_trans, rule approx, force)
+  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
+    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
+  have [simp]: "\<And>i. l i < b"
+    apply (rule order_less_trans) prefer 2
+    by (rule approx, auto, rule approx)
+  have [simp]: "\<And>i. a < u i"
+    by (rule order_less_trans, rule approx, auto, rule approx)
+  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> l j \<le> l i" by (rule decseqD, rule approx)
+  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> u i \<le> u j" by (rule incseqD, rule approx)
+  have g_nondec [simp]: "\<And>x y. a < x \<Longrightarrow> x \<le> y \<Longrightarrow> y < b \<Longrightarrow> g x \<le> g y"
+    apply (erule DERIV_nonneg_imp_nondecreasing, auto)
+    apply (rule exI, rule conjI, rule deriv_g)
+    apply (erule order_less_le_trans, auto)
+    apply (rule order_le_less_trans, subst ereal_less_eq(3), assumption, auto)
+    apply (rule g'_nonneg)
+    apply (rule less_imp_le, erule order_less_le_trans, auto)
+    by (rule less_imp_le, rule le_less_trans, subst ereal_less_eq(3), assumption, auto)
+  have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
+  proof -
+    have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A"
+      using A apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
+      by (drule_tac x = "\<lambda>i. ereal (l i)" in spec, auto)
+    hence A3: "\<And>i. g (l i) \<ge> A"
+      by (intro decseq_le, auto simp add: decseq_def)
+    have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B"
+      using B apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
+      by (drule_tac x = "\<lambda>i. ereal (u i)" in spec, auto)
+    hence B3: "\<And>i. g (u i) \<le> B"
+      by (intro incseq_le, auto simp add: incseq_def)
+    show "A \<le> B"
+      apply (rule order_trans [OF A3 [of 0]])
+      apply (rule order_trans [OF _ B3 [of 0]])
+      by auto
+    { fix x :: real
+      assume "A < x" and "x < B"
+      then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially"
+        apply (intro eventually_conj order_tendstoD)
+        by (rule A2, assumption, rule B2, assumption)
+      hence "\<exists>i. g (l i) < x \<and> x < g (u i)"
+        by (simp add: eventually_sequentially, auto)
+    } note AB = this
+    show "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
+      apply (auto simp add: einterval_def)
+      apply (erule (1) AB)
+      apply (rule order_le_less_trans, rule A3, simp)
+      apply (rule order_less_le_trans) prefer 2
+      by (rule B3, simp)
+  qed
+  (* finally, the main argument *)
+  {
+     fix i
+     have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)"
+        apply (rule interval_integral_substitution_finite, auto)
+        apply (rule DERIV_subset, rule deriv_g, auto)
+        apply (rule continuous_at_imp_continuous_on, auto, rule contf, auto)
+        by (rule continuous_at_imp_continuous_on, auto, rule contg', auto)
+     then have "(LBINT x=l i.. u i. (f (g x) * g' x)) = (LBINT y=g (l i)..g (u i). f y)"
+       by (simp add: ac_simps)
+  } note eq1 = this
+  have "(\<lambda>i. LBINT x=l i..u i. f (g x) * g' x)
+      \<longlonglongrightarrow> (LBINT x=a..b. f (g x) * g' x)"
+    apply (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx])
+    by (rule assms)
+  hence 2: "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x=a..b. f (g x) * g' x)"
+    by (simp add: eq1)
+  have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
+    apply (auto simp add: incseq_def)
+    apply (rule order_le_less_trans)
+    prefer 2 apply assumption
+    apply (rule g_nondec, auto)
+    by (erule order_less_le_trans, rule g_nondec, auto)
+  have img: "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> \<exists>c \<ge> l i. c \<le> u i \<and> x = g c"
+    apply (frule (1) IVT' [of g], auto)
+    apply (rule continuous_at_imp_continuous_on, auto)
+    by (rule DERIV_isCont, rule deriv_g, auto)
+  have nonneg_f2: "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> 0 \<le> f x"
+    by (frule (1) img, auto, rule f_nonneg, auto)
+  have contf_2: "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> isCont f x"
+    by (frule (1) img, auto, rule contf, auto)
+  have integrable: "set_integrable lborel (\<Union>i. {g (l i)<..<g (u i)}) f"
+    apply (rule pos_integrable_to_top, auto simp del: real_scaleR_def)
+    apply (rule incseq)
+    apply (rule nonneg_f2, erule less_imp_le, erule less_imp_le)
+    apply (rule set_integrable_subset)
+    apply (rule borel_integrable_atLeastAtMost')
+    apply (rule continuous_at_imp_continuous_on)
+    apply (clarsimp, erule (1) contf_2, auto)
+    apply (erule less_imp_le)+
+    using 2 unfolding interval_lebesgue_integral_def
+    by auto
+  thus "set_integrable lborel (einterval A B) f"
+    by (simp add: un)
+
+  have "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
+  proof (rule interval_integral_substitution_integrable)
+    show "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
+      using integrable_fg by (simp add: ac_simps)
+  qed fact+
+  then show "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
+    by (simp add: ac_simps)
+qed
+
+
+syntax
+"_complex_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> complex"
+("(2CLBINT _. _)" [0,60] 60)
+
+translations
+"CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\<lambda>x. f)"
+
+syntax
+"_complex_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> complex"
+("(3CLBINT _:_. _)" [0,60,61] 60)
+
+translations
+"CLBINT x:A. f" == "CONST complex_set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
+
+abbreviation complex_interval_lebesgue_integral ::
+    "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> complex" where
+  "complex_interval_lebesgue_integral M a b f \<equiv> interval_lebesgue_integral M a b f"
+
+abbreviation complex_interval_lebesgue_integrable ::
+  "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool" where
+  "complex_interval_lebesgue_integrable M a b f \<equiv> interval_lebesgue_integrable M a b f"
+
+syntax
+  "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> real \<Rightarrow> complex"
+  ("(4CLBINT _=_.._. _)" [0,60,60,61] 60)
+
+translations
+  "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
+
+lemma interval_integral_norm:
+  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  shows "interval_lebesgue_integrable lborel a b f \<Longrightarrow> a \<le> b \<Longrightarrow>
+    norm (LBINT t=a..b. f t) \<le> LBINT t=a..b. norm (f t)"
+  using integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
+  by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def)
+
+lemma interval_integral_norm2:
+  "interval_lebesgue_integrable lborel a b f \<Longrightarrow>
+    norm (LBINT t=a..b. f t) \<le> \<bar>LBINT t=a..b. norm (f t)\<bar>"
+proof (induct a b rule: linorder_wlog)
+  case (sym a b) then show ?case
+    by (simp add: interval_integral_endpoints_reverse[of a b] interval_integrable_endpoints_reverse[of a b])
+next
+  case (le a b)
+  then have "\<bar>LBINT t=a..b. norm (f t)\<bar> = LBINT t=a..b. norm (f t)"
+    using integrable_norm[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
+    by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def
+             intro!: integral_nonneg_AE abs_of_nonneg)
+  then show ?case
+    using le by (simp add: interval_integral_norm)
+qed
+
+(* TODO: should we have a library of facts like these? *)
+lemma integral_cos: "t \<noteq> 0 \<Longrightarrow> LBINT x=a..b. cos (t * x) = sin (t * b) / t - sin (t * a) / t"
+  apply (intro interval_integral_FTC_finite continuous_intros)
+  by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative[symmetric])
+
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Lebesgue_Integral_Substitution.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,400 @@
+(*  Title:      HOL/Probability/Lebesgue_Integral_Substitution.thy
+    Author:     Manuel Eberl
+
+    Provides lemmas for integration by substitution for the basic integral types.
+    Note that the substitution function must have a nonnegative derivative.
+    This could probably be weakened somehow.
+*)
+
+section \<open>Integration by Substition\<close>
+
+theory Lebesgue_Integral_Substitution
+imports Interval_Integral
+begin
+
+lemma nn_integral_substitution_aux:
+  fixes f :: "real \<Rightarrow> ennreal"
+  assumes Mf: "f \<in> borel_measurable borel"
+  assumes nonnegf: "\<And>x. f x \<ge> 0"
+  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
+  assumes contg': "continuous_on {a..b} g'"
+  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
+  assumes "a < b"
+  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
+             (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
+proof-
+  from \<open>a < b\<close> have [simp]: "a \<le> b" by simp
+  from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
+  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and
+                             Mg': "set_borel_measurable borel {a..b} g'"
+      by (simp_all only: set_measurable_continuous_on_ivl)
+  from derivg have derivg': "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
+    by (simp only: has_field_derivative_iff_has_vector_derivative)
+
+  have real_ind[simp]: "\<And>A x. enn2real (indicator A x) = indicator A x"
+      by (auto split: split_indicator)
+  have ennreal_ind[simp]: "\<And>A x. ennreal (indicator A x) = indicator A x"
+      by (auto split: split_indicator)
+  have [simp]: "\<And>x A. indicator A (g x) = indicator (g -` A) x"
+      by (auto split: split_indicator)
+
+  from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
+    by (rule deriv_nonneg_imp_mono) simp_all
+  with monog have [simp]: "g a \<le> g b" by (auto intro: mono_onD)
+
+  show ?thesis
+  proof (induction rule: borel_measurable_induct[OF Mf, case_names cong set mult add sup])
+    case (cong f1 f2)
+    from cong.hyps(3) have "f1 = f2" by auto
+    with cong show ?case by simp
+  next
+    case (set A)
+    from set.hyps show ?case
+    proof (induction rule: borel_set_induct)
+      case empty
+      thus ?case by simp
+    next
+      case (interval c d)
+      {
+        fix u v :: real assume asm: "{u..v} \<subseteq> {g a..g b}" "u \<le> v"
+
+        obtain u' v' where u'v': "{a..b} \<inter> g-`{u..v} = {u'..v'}" "u' \<le> v'" "g u' = u" "g v' = v"
+             using asm by (rule_tac continuous_interval_vimage_Int[OF contg monog, of u v]) simp_all
+        hence "{u'..v'} \<subseteq> {a..b}" "{u'..v'} \<subseteq> g -` {u..v}" by blast+
+        with u'v'(2) have "u' \<in> g -` {u..v}" "v' \<in> g -` {u..v}" by auto
+        from u'v'(1) have [simp]: "{a..b} \<inter> g -` {u..v} \<in> sets borel" by simp
+
+        have A: "continuous_on {min u' v'..max u' v'} g'"
+            by (simp only: u'v' max_absorb2 min_absorb1)
+               (intro continuous_on_subset[OF contg'], insert u'v', auto)
+        have "\<And>x. x \<in> {u'..v'} \<Longrightarrow> (g has_real_derivative g' x) (at x within {u'..v'})"
+           using asm by (intro has_field_derivative_subset[OF derivg] set_mp[OF \<open>{u'..v'} \<subseteq> {a..b}\<close>]) auto
+        hence B: "\<And>x. min u' v' \<le> x \<Longrightarrow> x \<le> max u' v' \<Longrightarrow>
+                      (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})"
+            by (simp only: u'v' max_absorb2 min_absorb1)
+               (auto simp: has_field_derivative_iff_has_vector_derivative)
+        have "integrable lborel (\<lambda>x. indicator ({a..b} \<inter> g -` {u..v}) x *\<^sub>R g' x)"
+          by (rule set_integrable_subset[OF borel_integrable_atLeastAtMost'[OF contg']]) simp_all
+        hence "(\<integral>\<^sup>+x. ennreal (g' x) * indicator ({a..b} \<inter> g-` {u..v}) x \<partial>lborel) =
+                   LBINT x:{a..b} \<inter> g-`{u..v}. g' x"
+          by (subst nn_integral_eq_integral[symmetric])
+             (auto intro!: derivg_nonneg nn_integral_cong split: split_indicator)
+        also from interval_integral_FTC_finite[OF A B]
+            have "LBINT x:{a..b} \<inter> g-`{u..v}. g' x = v - u"
+                by (simp add: u'v' interval_integral_Icc \<open>u \<le> v\<close>)
+        finally have "(\<integral>\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \<inter> g -` {u..v}) x \<partial>lborel) =
+                           ennreal (v - u)" .
+      } note A = this
+
+      have "(\<integral>\<^sup>+x. indicator {c..d} (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) =
+               (\<integral>\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \<inter> g -` {c..d}) x \<partial>lborel)"
+        by (intro nn_integral_cong) (simp split: split_indicator)
+      also have "{a..b} \<inter> g-`{c..d} = {a..b} \<inter> g-`{max (g a) c..min (g b) d}"
+        using \<open>a \<le> b\<close> \<open>c \<le> d\<close>
+        by (auto intro!: monog intro: order.trans)
+      also have "(\<integral>\<^sup>+ x. ennreal (g' x) * indicator ... x \<partial>lborel) =
+        (if max (g a) c \<le> min (g b) d then min (g b) d - max (g a) c else 0)"
+         using \<open>c \<le> d\<close> by (simp add: A)
+      also have "... = (\<integral>\<^sup>+ x. indicator ({g a..g b} \<inter> {c..d}) x \<partial>lborel)"
+        by (subst nn_integral_indicator) (auto intro!: measurable_sets Mg simp:)
+      also have "... = (\<integral>\<^sup>+ x. indicator {c..d} x * indicator {g a..g b} x \<partial>lborel)"
+        by (intro nn_integral_cong) (auto split: split_indicator)
+      finally show ?case ..
+
+      next
+
+      case (compl A)
+      note \<open>A \<in> sets borel\<close>[measurable]
+      from emeasure_mono[of "A \<inter> {g a..g b}" "{g a..g b}" lborel]
+      have [simp]: "emeasure lborel (A \<inter> {g a..g b}) \<noteq> top" by (auto simp: top_unique)
+      have [simp]: "g -` A \<inter> {a..b} \<in> sets borel"
+        by (rule set_borel_measurable_sets[OF Mg]) auto
+      have [simp]: "g -` (-A) \<inter> {a..b} \<in> sets borel"
+        by (rule set_borel_measurable_sets[OF Mg]) auto
+
+      have "(\<integral>\<^sup>+x. indicator (-A) x * indicator {g a..g b} x \<partial>lborel) =
+                (\<integral>\<^sup>+x. indicator (-A \<inter> {g a..g b}) x \<partial>lborel)"
+        by (rule nn_integral_cong) (simp split: split_indicator)
+      also from compl have "... = emeasure lborel ({g a..g b} - A)" using derivg_nonneg
+        by (simp add: vimage_Compl diff_eq Int_commute[of "-A"])
+      also have "{g a..g b} - A = {g a..g b} - A \<inter> {g a..g b}" by blast
+      also have "emeasure lborel ... = g b - g a - emeasure lborel (A \<inter> {g a..g b})"
+             using \<open>A \<in> sets borel\<close> by (subst emeasure_Diff) (auto simp: )
+     also have "emeasure lborel (A \<inter> {g a..g b}) =
+                    \<integral>\<^sup>+x. indicator A x * indicator {g a..g b} x \<partial>lborel"
+       using \<open>A \<in> sets borel\<close>
+       by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong)
+          (simp split: split_indicator)
+      also have "... = \<integral>\<^sup>+ x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel" (is "_ = ?I")
+        by (subst compl.IH, intro nn_integral_cong) (simp split: split_indicator)
+      also have "g b - g a = LBINT x:{a..b}. g' x" using derivg'
+        by (intro integral_FTC_atLeastAtMost[symmetric])
+           (auto intro: continuous_on_subset[OF contg'] has_field_derivative_subset[OF derivg]
+                 has_vector_derivative_at_within)
+      also have "ennreal ... = \<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel"
+        using borel_integrable_atLeastAtMost'[OF contg']
+        by (subst nn_integral_eq_integral)
+           (simp_all add: mult.commute derivg_nonneg split: split_indicator)
+      also have Mg'': "(\<lambda>x. indicator (g -` A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x))
+                            \<in> borel_measurable borel" using Mg'
+        by (intro borel_measurable_times_ennreal borel_measurable_indicator)
+           (simp_all add: mult.commute)
+      have le: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel) \<le>
+                        (\<integral>\<^sup>+x. ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
+         by (intro nn_integral_mono) (simp split: split_indicator add: derivg_nonneg)
+      note integrable = borel_integrable_atLeastAtMost'[OF contg']
+      with le have notinf: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel) \<noteq> top"
+          by (auto simp: real_integrable_def nn_integral_set_ennreal mult.commute top_unique)
+      have "(\<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel) - ?I =
+                  \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) -
+                        indicator (g -` A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel"
+        apply (intro nn_integral_diff[symmetric])
+        apply (insert Mg', simp add: mult.commute) []
+        apply (insert Mg'', simp) []
+        apply (simp split: split_indicator add: derivg_nonneg)
+        apply (rule notinf)
+        apply (simp split: split_indicator add: derivg_nonneg)
+        done
+      also have "... = \<integral>\<^sup>+ x. indicator (-A) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+        by (intro nn_integral_cong) (simp split: split_indicator)
+      finally show ?case .
+
+    next
+      case (union f)
+      then have [simp]: "\<And>i. {a..b} \<inter> g -` f i \<in> sets borel"
+        by (subst Int_commute, intro set_borel_measurable_sets[OF Mg]) auto
+      have "g -` (\<Union>i. f i) \<inter> {a..b} = (\<Union>i. {a..b} \<inter> g -` f i)" by auto
+      hence "g -` (\<Union>i. f i) \<inter> {a..b} \<in> sets borel" by (auto simp del: UN_simps)
+
+      have "(\<integral>\<^sup>+x. indicator (\<Union>i. f i) x * indicator {g a..g b} x \<partial>lborel) =
+                \<integral>\<^sup>+x. indicator (\<Union>i. {g a..g b} \<inter> f i) x \<partial>lborel"
+          by (intro nn_integral_cong) (simp split: split_indicator)
+      also from union have "... = emeasure lborel (\<Union>i. {g a..g b} \<inter> f i)" by simp
+      also from union have "... = (\<Sum>i. emeasure lborel ({g a..g b} \<inter> f i))"
+        by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def)
+      also from union have "... = (\<Sum>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel)" by simp
+      also have "(\<lambda>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel) =
+                           (\<lambda>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel)"
+        by (intro ext nn_integral_cong) (simp split: split_indicator)
+      also from union.IH have "(\<Sum>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel) =
+          (\<Sum>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)" by simp
+      also have "(\<lambda>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) =
+                         (\<lambda>i. \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x \<partial>lborel)"
+        by (intro ext nn_integral_cong) (simp split: split_indicator)
+      also have "(\<Sum>i. ... i) = \<integral>\<^sup>+ x. (\<Sum>i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) \<partial>lborel"
+        using Mg'
+        apply (intro nn_integral_suminf[symmetric])
+        apply (rule borel_measurable_times_ennreal, simp add: mult.commute)
+        apply (rule borel_measurable_indicator, subst sets_lborel)
+        apply (simp_all split: split_indicator add: derivg_nonneg)
+        done
+      also have "(\<lambda>x i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) =
+                      (\<lambda>x i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x)"
+        by (intro ext) (simp split: split_indicator)
+      also have "(\<integral>\<^sup>+ x. (\<Sum>i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x) \<partial>lborel) =
+                     \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * (\<Sum>i. indicator (g -` f i) x) \<partial>lborel"
+        by (intro nn_integral_cong) (auto split: split_indicator simp: derivg_nonneg)
+      also from union have "(\<lambda>x. \<Sum>i. indicator (g -` f i) x :: ennreal) = (\<lambda>x. indicator (\<Union>i. g -` f i) x)"
+        by (intro ext suminf_indicator) (auto simp: disjoint_family_on_def)
+      also have "(\<integral>\<^sup>+x. ennreal (g' x * indicator {a..b} x) * ... x \<partial>lborel) =
+                    (\<integral>\<^sup>+x. indicator (\<Union>i. f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
+       by (intro nn_integral_cong) (simp split: split_indicator)
+      finally show ?case .
+  qed
+
+next
+  case (mult f c)
+    note Mf[measurable] = \<open>f \<in> borel_measurable borel\<close>
+    let ?I = "indicator {a..b}"
+    have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
+      by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf])
+         (simp_all add: mult.commute)
+    also have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. f (g x) * ennreal (g' x) * ?I x)"
+      by (intro ext) (simp split: split_indicator)
+    finally have Mf': "(\<lambda>x. f (g x) * ennreal (g' x) * ?I x) \<in> borel_measurable borel" .
+    with mult show ?case
+      by (subst (1 2 3) mult_ac, subst (1 2) nn_integral_cmult) (simp_all add: mult_ac)
+
+next
+  case (add f2 f1)
+    let ?I = "indicator {a..b}"
+    {
+      fix f :: "real \<Rightarrow> ennreal" assume Mf: "f \<in> borel_measurable borel"
+      have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
+        by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf])
+           (simp_all add:  mult.commute)
+      also have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. f (g x) * ennreal (g' x) * ?I x)"
+        by (intro ext) (simp split: split_indicator)
+      finally have "(\<lambda>x. f (g x) * ennreal (g' x) * ?I x) \<in> borel_measurable borel" .
+    } note Mf' = this[OF \<open>f1 \<in> borel_measurable borel\<close>] this[OF \<open>f2 \<in> borel_measurable borel\<close>]
+
+    have "(\<integral>\<^sup>+ x. (f1 x + f2 x) * indicator {g a..g b} x \<partial>lborel) =
+             (\<integral>\<^sup>+ x. f1 x * indicator {g a..g b} x + f2 x * indicator {g a..g b} x \<partial>lborel)"
+      by (intro nn_integral_cong) (simp split: split_indicator)
+    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) +
+                                (\<integral>\<^sup>+ x. f2 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
+      by (simp_all add: nn_integral_add)
+    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x +
+                                      f2 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
+      by (intro nn_integral_add[symmetric])
+         (auto simp add: Mf' derivg_nonneg split: split_indicator)
+    also have "... = \<integral>\<^sup>+ x. (f1 (g x) + f2 (g x)) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+      by (intro nn_integral_cong) (simp split: split_indicator add: distrib_right)
+    finally show ?case .
+
+next
+  case (sup F)
+  {
+    fix i
+    let ?I = "indicator {a..b}"
+    have "(\<lambda>x. F i (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
+      by (rule_tac borel_measurable_times_ennreal, rule_tac measurable_compose[OF _ sup.hyps(1)])
+         (simp_all add: mult.commute)
+    also have "(\<lambda>x. F i (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. F i (g x) * ennreal (g' x) * ?I x)"
+      by (intro ext) (simp split: split_indicator)
+     finally have "... \<in> borel_measurable borel" .
+  } note Mf' = this
+
+    have "(\<integral>\<^sup>+x. (SUP i. F i x) * indicator {g a..g b} x \<partial>lborel) =
+               \<integral>\<^sup>+x. (SUP i. F i x* indicator {g a..g b} x) \<partial>lborel"
+      by (intro nn_integral_cong) (simp split: split_indicator)
+    also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i x* indicator {g a..g b} x \<partial>lborel)"
+      by (intro nn_integral_monotone_convergence_SUP)
+         (auto simp: incseq_def le_fun_def split: split_indicator)
+    also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
+      by simp
+    also from sup have "... =  \<integral>\<^sup>+x. (SUP i. F i (g x) * ennreal (g' x) * indicator {a..b} x) \<partial>lborel"
+      by (intro nn_integral_monotone_convergence_SUP[symmetric])
+         (auto simp: incseq_def le_fun_def derivg_nonneg Mf' split: split_indicator
+               intro!: mult_right_mono)
+    also from sup have "... = \<integral>\<^sup>+x. (SUP i. F i (g x)) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+      by (subst mult.assoc, subst mult.commute, subst SUP_mult_left_ennreal)
+         (auto split: split_indicator simp: derivg_nonneg mult_ac)
+    finally show ?case by simp
+  qed
+qed
+
+lemma nn_integral_substitution:
+  fixes f :: "real \<Rightarrow> real"
+  assumes Mf[measurable]: "set_borel_measurable borel {g a..g b} f"
+  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
+  assumes contg': "continuous_on {a..b} g'"
+  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
+  assumes "a \<le> b"
+  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
+             (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
+proof (cases "a = b")
+  assume "a \<noteq> b"
+  with \<open>a \<le> b\<close> have "a < b" by auto
+  let ?f' = "\<lambda>x. f x * indicator {g a..g b} x"
+
+  from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
+    by (rule deriv_nonneg_imp_mono) simp_all
+  have bounds: "\<And>x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g x \<ge> g a" "\<And>x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g x \<le> g b"
+    by (auto intro: monog)
+
+  from derivg_nonneg have nonneg:
+    "\<And>f x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g' x \<noteq> 0 \<Longrightarrow> f x * ennreal (g' x) \<ge> 0 \<Longrightarrow> f x \<ge> 0"
+    by (force simp: field_simps)
+  have nonneg': "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<not> 0 \<le> f (g x) \<Longrightarrow> 0 \<le> f (g x) * g' x \<Longrightarrow> g' x = 0"
+    by (metis atLeastAtMost_iff derivg_nonneg eq_iff mult_eq_0_iff mult_le_0_iff)
+
+  have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
+            (\<integral>\<^sup>+x. ennreal (?f' x) * indicator {g a..g b} x \<partial>lborel)"
+    by (intro nn_integral_cong)
+       (auto split: split_indicator split_max simp: zero_ennreal.rep_eq ennreal_neg)
+  also have "... = \<integral>\<^sup>+ x. ?f' (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel" using Mf
+    by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \<open>a < b\<close>])
+       (auto simp add: mult.commute)
+  also have "... = \<integral>\<^sup>+ x. f (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
+    by (intro nn_integral_cong) (auto split: split_indicator simp: max_def dest: bounds)
+  also have "... = \<integral>\<^sup>+x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel"
+    by (intro nn_integral_cong) (auto simp: mult.commute derivg_nonneg ennreal_mult' split: split_indicator)
+  finally show ?thesis .
+qed auto
+
+lemma integral_substitution:
+  assumes integrable: "set_integrable lborel {g a..g b} f"
+  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
+  assumes contg': "continuous_on {a..b} g'"
+  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
+  assumes "a \<le> b"
+  shows "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
+    and "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)"
+proof-
+  from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
+  with contg' have Mg: "set_borel_measurable borel {a..b} g"
+    and Mg': "set_borel_measurable borel {a..b} g'"
+    by (simp_all only: set_measurable_continuous_on_ivl)
+  from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
+    by (rule deriv_nonneg_imp_mono) simp_all
+
+  have "(\<lambda>x. ennreal (f x) * indicator {g a..g b} x) =
+           (\<lambda>x. ennreal (f x * indicator {g a..g b} x))"
+    by (intro ext) (simp split: split_indicator)
+  with integrable have M1: "(\<lambda>x. f x * indicator {g a..g b} x) \<in> borel_measurable borel"
+    unfolding real_integrable_def by (force simp: mult.commute)
+  from integrable have M2: "(\<lambda>x. -f x * indicator {g a..g b} x) \<in> borel_measurable borel"
+    unfolding real_integrable_def by (force simp: mult.commute)
+
+  have "LBINT x. (f x :: real) * indicator {g a..g b} x =
+          enn2real (\<integral>\<^sup>+ x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) -
+          enn2real (\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel)" using integrable
+    by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ennreal mult.commute)
+  also have *: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) =
+      (\<integral>\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel)"
+    by (intro nn_integral_cong) (simp split: split_indicator)
+  also from M1 * have A: "(\<integral>\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel) =
+                            (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel)"
+    by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
+       (auto simp: nn_integral_set_ennreal mult.commute)
+  also have **: "(\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
+      (\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
+    by (intro nn_integral_cong) (simp split: split_indicator)
+  also from M2 ** have B: "(\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
+        (\<integral>\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
+    by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
+       (auto simp: nn_integral_set_ennreal mult.commute)
+
+  also {
+    from integrable have Mf: "set_borel_measurable borel {g a..g b} f"
+      unfolding real_integrable_def by simp
+    from borel_measurable_times[OF measurable_compose[OF Mg Mf] Mg']
+      have "(\<lambda>x. f (g x * indicator {a..b} x) * indicator {g a..g b} (g x * indicator {a..b} x) *
+                     (g' x * indicator {a..b} x)) \<in> borel_measurable borel"  (is "?f \<in> _")
+      by (simp add: mult.commute)
+    also have "?f = (\<lambda>x. f (g x) * g' x * indicator {a..b} x)"
+      using monog by (intro ext) (auto split: split_indicator)
+    finally show "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
+      using A B integrable unfolding real_integrable_def
+      by (simp_all add: nn_integral_set_ennreal mult.commute)
+  } note integrable' = this
+
+  have "enn2real (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel) -
+                  enn2real (\<integral>\<^sup>+ x. ennreal (-f (g x) * g' x * indicator {a..b} x) \<partial>lborel) =
+                (LBINT x. f (g x) * g' x * indicator {a..b} x)" using integrable'
+    by (subst real_lebesgue_integral_def) (simp_all add: field_simps)
+  finally show "(LBINT x. f x * indicator {g a..g b} x) =
+                     (LBINT x. f (g x) * g' x * indicator {a..b} x)" .
+qed
+
+lemma interval_integral_substitution:
+  assumes integrable: "set_integrable lborel {g a..g b} f"
+  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
+  assumes contg': "continuous_on {a..b} g'"
+  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
+  assumes "a \<le> b"
+  shows "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
+    and "(LBINT x=g a..g b. f x) = (LBINT x=a..b. f (g x) * g' x)"
+  apply (rule integral_substitution[OF assms], simp, simp)
+  apply (subst (1 2) interval_integral_Icc, fact)
+  apply (rule deriv_nonneg_imp_mono[OF derivg derivg_nonneg], simp, simp, fact)
+  using integral_substitution(2)[OF assms]
+  apply (simp add: mult.commute)
+  done
+
+lemma set_borel_integrable_singleton[simp]:
+  "set_integrable lborel {x} (f :: real \<Rightarrow> real)"
+  by (subst integrable_discrete_difference[where X="{x}" and g="\<lambda>_. 0"]) auto
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Lebesgue_Measure.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,1381 @@
+(*  Title:      HOL/Probability/Lebesgue_Measure.thy
+    Author:     Johannes Hölzl, TU München
+    Author:     Robert Himmelmann, TU München
+    Author:     Jeremy Avigad
+    Author:     Luke Serafin
+*)
+
+section \<open>Lebesgue measure\<close>
+
+theory Lebesgue_Measure
+  imports Finite_Product_Measure Bochner_Integration Caratheodory
+begin
+
+subsection \<open>Every right continuous and nondecreasing function gives rise to a measure\<close>
+
+definition interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
+  "interval_measure F = extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a <.. b}) (\<lambda>(a, b). ennreal (F b - F a))"
+
+lemma emeasure_interval_measure_Ioc:
+  assumes "a \<le> b"
+  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
+  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
+  shows "emeasure (interval_measure F) {a <.. b} = F b - F a"
+proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \<open>a \<le> b\<close>])
+  show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \<le> b}"
+  proof (unfold_locales, safe)
+    fix a b c d :: real assume *: "a \<le> b" "c \<le> d"
+    then show "\<exists>C\<subseteq>{{a<..b} |a b. a \<le> b}. finite C \<and> disjoint C \<and> {a<..b} - {c<..d} = \<Union>C"
+    proof cases
+      let ?C = "{{a<..b}}"
+      assume "b < c \<or> d \<le> a \<or> d \<le> c"
+      with * have "?C \<subseteq> {{a<..b} |a b. a \<le> b} \<and> finite ?C \<and> disjoint ?C \<and> {a<..b} - {c<..d} = \<Union>?C"
+        by (auto simp add: disjoint_def)
+      thus ?thesis ..
+    next
+      let ?C = "{{a<..c}, {d<..b}}"
+      assume "\<not> (b < c \<or> d \<le> a \<or> d \<le> c)"
+      with * have "?C \<subseteq> {{a<..b} |a b. a \<le> b} \<and> finite ?C \<and> disjoint ?C \<and> {a<..b} - {c<..d} = \<Union>?C"
+        by (auto simp add: disjoint_def Ioc_inj) (metis linear)+
+      thus ?thesis ..
+    qed
+  qed (auto simp: Ioc_inj, metis linear)
+next
+  fix l r :: "nat \<Rightarrow> real" and a b :: real
+  assume l_r[simp]: "\<And>n. l n \<le> r n" and "a \<le> b" and disj: "disjoint_family (\<lambda>n. {l n<..r n})"
+  assume lr_eq_ab: "(\<Union>i. {l i<..r i}) = {a<..b}"
+
+  have [intro, simp]: "\<And>a b. a \<le> b \<Longrightarrow> F a \<le> F b"
+    by (auto intro!: l_r mono_F)
+
+  { fix S :: "nat set" assume "finite S"
+    moreover note \<open>a \<le> b\<close>
+    moreover have "\<And>i. i \<in> S \<Longrightarrow> {l i <.. r i} \<subseteq> {a <.. b}"
+      unfolding lr_eq_ab[symmetric] by auto
+    ultimately have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<le> F b - F a"
+    proof (induction S arbitrary: a rule: finite_psubset_induct)
+      case (psubset S)
+      show ?case
+      proof cases
+        assume "\<exists>i\<in>S. l i < r i"
+        with \<open>finite S\<close> have "Min (l ` {i\<in>S. l i < r i}) \<in> l ` {i\<in>S. l i < r i}"
+          by (intro Min_in) auto
+        then obtain m where m: "m \<in> S" "l m < r m" "l m = Min (l ` {i\<in>S. l i < r i})"
+          by fastforce
+
+        have "(\<Sum>i\<in>S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (\<Sum>i\<in>S - {m}. F (r i) - F (l i))"
+          using m psubset by (intro setsum.remove) auto
+        also have "(\<Sum>i\<in>S - {m}. F (r i) - F (l i)) \<le> F b - F (r m)"
+        proof (intro psubset.IH)
+          show "S - {m} \<subset> S"
+            using \<open>m\<in>S\<close> by auto
+          show "r m \<le> b"
+            using psubset.prems(2)[OF \<open>m\<in>S\<close>] \<open>l m < r m\<close> by auto
+        next
+          fix i assume "i \<in> S - {m}"
+          then have i: "i \<in> S" "i \<noteq> m" by auto
+          { assume i': "l i < r i" "l i < r m"
+            with \<open>finite S\<close> i m have "l m \<le> l i"
+              by auto
+            with i' have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
+              by auto
+            then have False
+              using disjoint_family_onD[OF disj, of i m] i by auto }
+          then have "l i \<noteq> r i \<Longrightarrow> r m \<le> l i"
+            unfolding not_less[symmetric] using l_r[of i] by auto
+          then show "{l i <.. r i} \<subseteq> {r m <.. b}"
+            using psubset.prems(2)[OF \<open>i\<in>S\<close>] by auto
+        qed
+        also have "F (r m) - F (l m) \<le> F (r m) - F a"
+          using psubset.prems(2)[OF \<open>m \<in> S\<close>] \<open>l m < r m\<close>
+          by (auto simp add: Ioc_subset_iff intro!: mono_F)
+        finally show ?case
+          by (auto intro: add_mono)
+      qed (auto simp add: \<open>a \<le> b\<close> less_le)
+    qed }
+  note claim1 = this
+
+  (* second key induction: a lower bound on the measures of any finite collection of Ai's
+     that cover an interval {u..v} *)
+
+  { fix S u v and l r :: "nat \<Rightarrow> real"
+    assume "finite S" "\<And>i. i\<in>S \<Longrightarrow> l i < r i" "{u..v} \<subseteq> (\<Union>i\<in>S. {l i<..< r i})"
+    then have "F v - F u \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
+    proof (induction arbitrary: v u rule: finite_psubset_induct)
+      case (psubset S)
+      show ?case
+      proof cases
+        assume "S = {}" then show ?case
+          using psubset by (simp add: mono_F)
+      next
+        assume "S \<noteq> {}"
+        then obtain j where "j \<in> S"
+          by auto
+
+        let ?R = "r j < u \<or> l j > v \<or> (\<exists>i\<in>S-{j}. l i \<le> l j \<and> r j \<le> r i)"
+        show ?case
+        proof cases
+          assume "?R"
+          with \<open>j \<in> S\<close> psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>S-{j}. {l i<..< r i})"
+            apply (auto simp: subset_eq Ball_def)
+            apply (metis Diff_iff less_le_trans leD linear singletonD)
+            apply (metis Diff_iff less_le_trans leD linear singletonD)
+            apply (metis order_trans less_le_not_le linear)
+            done
+          with \<open>j \<in> S\<close> have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l i))"
+            by (intro psubset) auto
+          also have "\<dots> \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
+            using psubset.prems
+            by (intro setsum_mono2 psubset) (auto intro: less_imp_le)
+          finally show ?thesis .
+        next
+          assume "\<not> ?R"
+          then have j: "u \<le> r j" "l j \<le> v" "\<And>i. i \<in> S - {j} \<Longrightarrow> r i < r j \<or> l i > l j"
+            by (auto simp: not_less)
+          let ?S1 = "{i \<in> S. l i < l j}"
+          let ?S2 = "{i \<in> S. r i > r j}"
+
+          have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<ge> (\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i))"
+            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
+            by (intro setsum_mono2) (auto intro: less_imp_le)
+          also have "(\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i)) =
+            (\<Sum>i\<in>?S1. F (r i) - F (l i)) + (\<Sum>i\<in>?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))"
+            using psubset(1) psubset.prems(1) j
+            apply (subst setsum.union_disjoint)
+            apply simp_all
+            apply (subst setsum.union_disjoint)
+            apply auto
+            apply (metis less_le_not_le)
+            done
+          also (xtrans) have "(\<Sum>i\<in>?S1. F (r i) - F (l i)) \<ge> F (l j) - F u"
+            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
+            apply (intro psubset.IH psubset)
+            apply (auto simp: subset_eq Ball_def)
+            apply (metis less_le_trans not_le)
+            done
+          also (xtrans) have "(\<Sum>i\<in>?S2. F (r i) - F (l i)) \<ge> F v - F (r j)"
+            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
+            apply (intro psubset.IH psubset)
+            apply (auto simp: subset_eq Ball_def)
+            apply (metis le_less_trans not_le)
+            done
+          finally (xtrans) show ?case
+            by (auto simp: add_mono)
+        qed
+      qed
+    qed }
+  note claim2 = this
+
+  (* now prove the inequality going the other way *)
+  have "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i)))"
+  proof (rule ennreal_le_epsilon)
+    fix epsilon :: real assume egt0: "epsilon > 0"
+    have "\<forall>i. \<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
+    proof
+      fix i
+      note right_cont_F [of "r i"]
+      thus "\<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
+        apply -
+        apply (subst (asm) continuous_at_right_real_increasing)
+        apply (rule mono_F, assumption)
+        apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec)
+        apply (erule impE)
+        using egt0 by (auto simp add: field_simps)
+    qed
+    then obtain delta where
+        deltai_gt0: "\<And>i. delta i > 0" and
+        deltai_prop: "\<And>i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)"
+      by metis
+    have "\<exists>a' > a. F a' - F a < epsilon / 2"
+      apply (insert right_cont_F [of a])
+      apply (subst (asm) continuous_at_right_real_increasing)
+      using mono_F apply force
+      apply (drule_tac x = "epsilon / 2" in spec)
+      using egt0 unfolding mult.commute [of 2] by force
+    then obtain a' where a'lea [arith]: "a' > a" and
+      a_prop: "F a' - F a < epsilon / 2"
+      by auto
+    define S' where "S' = {i. l i < r i}"
+    obtain S :: "nat set" where
+      "S \<subseteq> S'" and finS: "finite S" and
+      Sprop: "{a'..b} \<subseteq> (\<Union>i \<in> S. {l i<..<r i + delta i})"
+    proof (rule compactE_image)
+      show "compact {a'..b}"
+        by (rule compact_Icc)
+      show "\<forall>i \<in> S'. open ({l i<..<r i + delta i})" by auto
+      have "{a'..b} \<subseteq> {a <.. b}"
+        by auto
+      also have "{a <.. b} = (\<Union>i\<in>S'. {l i<..r i})"
+        unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans)
+      also have "\<dots> \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})"
+        apply (intro UN_mono)
+        apply (auto simp: S'_def)
+        apply (cut_tac i=i in deltai_gt0)
+        apply simp
+        done
+      finally show "{a'..b} \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})" .
+    qed
+    with S'_def have Sprop2: "\<And>i. i \<in> S \<Longrightarrow> l i < r i" by auto
+    from finS have "\<exists>n. \<forall>i \<in> S. i \<le> n"
+      by (subst finite_nat_set_iff_bounded_le [symmetric])
+    then obtain n where Sbound [rule_format]: "\<forall>i \<in> S. i \<le> n" ..
+    have "F b - F a' \<le> (\<Sum>i\<in>S. F (r i + delta i) - F (l i))"
+      apply (rule claim2 [rule_format])
+      using finS Sprop apply auto
+      apply (frule Sprop2)
+      apply (subgoal_tac "delta i > 0")
+      apply arith
+      by (rule deltai_gt0)
+    also have "... \<le> (\<Sum>i \<in> S. F(r i) - F(l i) + epsilon / 2^(i+2))"
+      apply (rule setsum_mono)
+      apply simp
+      apply (rule order_trans)
+      apply (rule less_imp_le)
+      apply (rule deltai_prop)
+      by auto
+    also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
+        (epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
+      by (subst setsum.distrib) (simp add: field_simps setsum_right_distrib)
+    also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
+      apply (rule add_left_mono)
+      apply (rule mult_left_mono)
+      apply (rule setsum_mono2)
+      using egt0 apply auto
+      by (frule Sbound, auto)
+    also have "... \<le> ?t + (epsilon / 2)"
+      apply (rule add_left_mono)
+      apply (subst geometric_sum)
+      apply auto
+      apply (rule mult_left_mono)
+      using egt0 apply auto
+      done
+    finally have aux2: "F b - F a' \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2"
+      by simp
+
+    have "F b - F a = (F b - F a') + (F a' - F a)"
+      by auto
+    also have "... \<le> (F b - F a') + epsilon / 2"
+      using a_prop by (intro add_left_mono) simp
+    also have "... \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2"
+      apply (intro add_right_mono)
+      apply (rule aux2)
+      done
+    also have "... = (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon"
+      by auto
+    also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
+      using finS Sbound Sprop by (auto intro!: add_right_mono setsum_mono3)
+    finally have "ennreal (F b - F a) \<le> (\<Sum>i\<le>n. ennreal (F (r i) - F (l i))) + epsilon"
+      using egt0 by (simp add: ennreal_plus[symmetric] setsum_nonneg del: ennreal_plus)
+    then show "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
+      by (rule order_trans) (auto intro!: add_mono setsum_le_suminf simp del: setsum_ennreal)
+  qed
+  moreover have "(\<Sum>i. ennreal (F (r i) - F (l i))) \<le> ennreal (F b - F a)"
+    using \<open>a \<le> b\<close> by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1)
+  ultimately show "(\<Sum>n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)"
+    by (rule antisym[rotated])
+qed (auto simp: Ioc_inj mono_F)
+
+lemma measure_interval_measure_Ioc:
+  assumes "a \<le> b"
+  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
+  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
+  shows "measure (interval_measure F) {a <.. b} = F b - F a"
+  unfolding measure_def
+  apply (subst emeasure_interval_measure_Ioc)
+  apply fact+
+  apply (simp add: assms)
+  done
+
+lemma emeasure_interval_measure_Ioc_eq:
+  "(\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y) \<Longrightarrow> (\<And>a. continuous (at_right a) F) \<Longrightarrow>
+    emeasure (interval_measure F) {a <.. b} = (if a \<le> b then F b - F a else 0)"
+  using emeasure_interval_measure_Ioc[of a b F] by auto
+
+lemma sets_interval_measure [simp, measurable_cong]: "sets (interval_measure F) = sets borel"
+  apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc)
+  apply (rule sigma_sets_eqI)
+  apply auto
+  apply (case_tac "a \<le> ba")
+  apply (auto intro: sigma_sets.Empty)
+  done
+
+lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV"
+  by (simp add: interval_measure_def space_extend_measure)
+
+lemma emeasure_interval_measure_Icc:
+  assumes "a \<le> b"
+  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
+  assumes cont_F : "continuous_on UNIV F"
+  shows "emeasure (interval_measure F) {a .. b} = F b - F a"
+proof (rule tendsto_unique)
+  { fix a b :: real assume "a \<le> b" then have "emeasure (interval_measure F) {a <.. b} = F b - F a"
+      using cont_F
+      by (subst emeasure_interval_measure_Ioc)
+         (auto intro: mono_F continuous_within_subset simp: continuous_on_eq_continuous_within) }
+  note * = this
+
+  let ?F = "interval_measure F"
+  show "((\<lambda>a. F b - F a) \<longlongrightarrow> emeasure ?F {a..b}) (at_left a)"
+  proof (rule tendsto_at_left_sequentially)
+    show "a - 1 < a" by simp
+    fix X assume "\<And>n. X n < a" "incseq X" "X \<longlonglongrightarrow> a"
+    with \<open>a \<le> b\<close> have "(\<lambda>n. emeasure ?F {X n<..b}) \<longlonglongrightarrow> emeasure ?F (\<Inter>n. {X n <..b})"
+      apply (intro Lim_emeasure_decseq)
+      apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *)
+      apply force
+      apply (subst (asm ) *)
+      apply (auto intro: less_le_trans less_imp_le)
+      done
+    also have "(\<Inter>n. {X n <..b}) = {a..b}"
+      using \<open>\<And>n. X n < a\<close>
+      apply auto
+      apply (rule LIMSEQ_le_const2[OF \<open>X \<longlonglongrightarrow> a\<close>])
+      apply (auto intro: less_imp_le)
+      apply (auto intro: less_le_trans)
+      done
+    also have "(\<lambda>n. emeasure ?F {X n<..b}) = (\<lambda>n. F b - F (X n))"
+      using \<open>\<And>n. X n < a\<close> \<open>a \<le> b\<close> by (subst *) (auto intro: less_imp_le less_le_trans)
+    finally show "(\<lambda>n. F b - F (X n)) \<longlonglongrightarrow> emeasure ?F {a..b}" .
+  qed
+  show "((\<lambda>a. ennreal (F b - F a)) \<longlongrightarrow> F b - F a) (at_left a)"
+    by (rule continuous_on_tendsto_compose[where g="\<lambda>x. x" and s=UNIV])
+       (auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const)
+qed (rule trivial_limit_at_left_real)
+
+lemma sigma_finite_interval_measure:
+  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
+  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
+  shows "sigma_finite_measure (interval_measure F)"
+  apply unfold_locales
+  apply (intro exI[of _ "(\<lambda>(a, b). {a <.. b}) ` (\<rat> \<times> \<rat>)"])
+  apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms])
+  done
+
+subsection \<open>Lebesgue-Borel measure\<close>
+
+definition lborel :: "('a :: euclidean_space) measure" where
+  "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. interval_measure (\<lambda>x. x)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
+
+lemma
+  shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel"
+    and space_lborel[simp]: "space lborel = space borel"
+    and measurable_lborel1[simp]: "measurable M lborel = measurable M borel"
+    and measurable_lborel2[simp]: "measurable lborel M = measurable borel M"
+  by (simp_all add: lborel_def)
+
+context
+begin
+
+interpretation sigma_finite_measure "interval_measure (\<lambda>x. x)"
+  by (rule sigma_finite_interval_measure) auto
+interpretation finite_product_sigma_finite "\<lambda>_. interval_measure (\<lambda>x. x)" Basis
+  proof qed simp
+
+lemma lborel_eq_real: "lborel = interval_measure (\<lambda>x. x)"
+  unfolding lborel_def Basis_real_def
+  using distr_id[of "interval_measure (\<lambda>x. x)"]
+  by (subst distr_component[symmetric])
+     (simp_all add: distr_distr comp_def del: distr_id cong: distr_cong)
+
+lemma lborel_eq: "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. lborel) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
+  by (subst lborel_def) (simp add: lborel_eq_real)
+
+lemma nn_integral_lborel_setprod:
+  assumes [measurable]: "\<And>b. b \<in> Basis \<Longrightarrow> f b \<in> borel_measurable borel"
+  assumes nn[simp]: "\<And>b x. b \<in> Basis \<Longrightarrow> 0 \<le> f b x"
+  shows "(\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. f b (x \<bullet> b)) \<partial>lborel) = (\<Prod>b\<in>Basis. (\<integral>\<^sup>+x. f b x \<partial>lborel))"
+  by (simp add: lborel_def nn_integral_distr product_nn_integral_setprod
+                product_nn_integral_singleton)
+
+lemma emeasure_lborel_Icc[simp]:
+  fixes l u :: real
+  assumes [simp]: "l \<le> u"
+  shows "emeasure lborel {l .. u} = u - l"
+proof -
+  have "((\<lambda>f. f 1) -` {l..u} \<inter> space (Pi\<^sub>M {1} (\<lambda>b. interval_measure (\<lambda>x. x)))) = {1::real} \<rightarrow>\<^sub>E {l..u}"
+    by (auto simp: space_PiM)
+  then show ?thesis
+    by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc continuous_on_id)
+qed
+
+lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \<le> u then u - l else 0)"
+  by simp
+
+lemma emeasure_lborel_cbox[simp]:
+  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
+  shows "emeasure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+proof -
+  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (cbox l u)"
+    by (auto simp: fun_eq_iff cbox_def split: split_indicator)
+  then have "emeasure lborel (cbox l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
+    by simp
+  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
+  finally show ?thesis .
+qed
+
+lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \<noteq> c"
+  using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c]
+  by (auto simp add: cbox_sing setprod_constant power_0_left)
+
+lemma emeasure_lborel_Ioo[simp]:
+  assumes [simp]: "l \<le> u"
+  shows "emeasure lborel {l <..< u} = ennreal (u - l)"
+proof -
+  have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}"
+    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
+  then show ?thesis
+    by simp
+qed
+
+lemma emeasure_lborel_Ioc[simp]:
+  assumes [simp]: "l \<le> u"
+  shows "emeasure lborel {l <.. u} = ennreal (u - l)"
+proof -
+  have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}"
+    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
+  then show ?thesis
+    by simp
+qed
+
+lemma emeasure_lborel_Ico[simp]:
+  assumes [simp]: "l \<le> u"
+  shows "emeasure lborel {l ..< u} = ennreal (u - l)"
+proof -
+  have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}"
+    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
+  then show ?thesis
+    by simp
+qed
+
+lemma emeasure_lborel_box[simp]:
+  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
+  shows "emeasure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+proof -
+  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (box l u)"
+    by (auto simp: fun_eq_iff box_def split: split_indicator)
+  then have "emeasure lborel (box l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
+    by simp
+  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
+  finally show ?thesis .
+qed
+
+lemma emeasure_lborel_cbox_eq:
+  "emeasure lborel (cbox l u) = (if \<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b then \<Prod>b\<in>Basis. (u - l) \<bullet> b else 0)"
+  using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le)
+
+lemma emeasure_lborel_box_eq:
+  "emeasure lborel (box l u) = (if \<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b then \<Prod>b\<in>Basis. (u - l) \<bullet> b else 0)"
+  using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force
+
+lemma
+  fixes l u :: real
+  assumes [simp]: "l \<le> u"
+  shows measure_lborel_Icc[simp]: "measure lborel {l .. u} = u - l"
+    and measure_lborel_Ico[simp]: "measure lborel {l ..< u} = u - l"
+    and measure_lborel_Ioc[simp]: "measure lborel {l <.. u} = u - l"
+    and measure_lborel_Ioo[simp]: "measure lborel {l <..< u} = u - l"
+  by (simp_all add: measure_def)
+
+lemma
+  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
+  shows measure_lborel_box[simp]: "measure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+    and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+  by (simp_all add: measure_def inner_diff_left setprod_nonneg)
+
+lemma sigma_finite_lborel: "sigma_finite_measure lborel"
+proof
+  show "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets lborel \<and> \<Union>A = space lborel \<and> (\<forall>a\<in>A. emeasure lborel a \<noteq> \<infinity>)"
+    by (intro exI[of _ "range (\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One))"])
+       (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV)
+qed
+
+end
+
+lemma emeasure_lborel_UNIV: "emeasure lborel (UNIV::'a::euclidean_space set) = \<infinity>"
+proof -
+  { fix n::nat
+    let ?Ba = "Basis :: 'a set"
+    have "real n \<le> (2::real) ^ card ?Ba * real n"
+      by (simp add: mult_le_cancel_right1)
+    also
+    have "... \<le> (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba"
+      apply (rule mult_left_mono)
+      apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc)
+      apply (simp add: DIM_positive)
+      done
+    finally have "real n \<le> (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" .
+  } note [intro!] = this
+  show ?thesis
+    unfolding UN_box_eq_UNIV[symmetric]
+    apply (subst SUP_emeasure_incseq[symmetric])
+    apply (auto simp: incseq_def subset_box inner_add_left setprod_constant
+      simp del: Sup_eq_top_iff SUP_eq_top_iff
+      intro!: ennreal_SUP_eq_top)
+    done
+qed
+
+lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0"
+  using emeasure_lborel_cbox[of x x] nonempty_Basis
+  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: cbox_sing setprod_constant)
+
+lemma emeasure_lborel_countable:
+  fixes A :: "'a::euclidean_space set"
+  assumes "countable A"
+  shows "emeasure lborel A = 0"
+proof -
+  have "A \<subseteq> (\<Union>i. {from_nat_into A i})" using from_nat_into_surj assms by force
+  then have "emeasure lborel A \<le> emeasure lborel (\<Union>i. {from_nat_into A i})"
+    by (intro emeasure_mono) auto
+  also have "emeasure lborel (\<Union>i. {from_nat_into A i}) = 0"
+    by (rule emeasure_UN_eq_0) auto
+  finally show ?thesis
+    by (auto simp add: )
+qed
+
+lemma countable_imp_null_set_lborel: "countable A \<Longrightarrow> A \<in> null_sets lborel"
+  by (simp add: null_sets_def emeasure_lborel_countable sets.countable)
+
+lemma finite_imp_null_set_lborel: "finite A \<Longrightarrow> A \<in> null_sets lborel"
+  by (intro countable_imp_null_set_lborel countable_finite)
+
+lemma lborel_neq_count_space[simp]: "lborel \<noteq> count_space (A::('a::ordered_euclidean_space) set)"
+proof
+  assume asm: "lborel = count_space A"
+  have "space lborel = UNIV" by simp
+  hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space)
+  have "emeasure lborel {undefined::'a} = 1"
+      by (subst asm, subst emeasure_count_space_finite) auto
+  moreover have "emeasure lborel {undefined} \<noteq> 1" by simp
+  ultimately show False by contradiction
+qed
+
+subsection \<open>Affine transformation on the Lebesgue-Borel\<close>
+
+lemma lborel_eqI:
+  fixes M :: "'a::euclidean_space measure"
+  assumes emeasure_eq: "\<And>l u. (\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b) \<Longrightarrow> emeasure M (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
+  assumes sets_eq: "sets M = sets borel"
+  shows "lborel = M"
+proof (rule measure_eqI_generator_eq)
+  let ?E = "range (\<lambda>(a, b). box a b::'a set)"
+  show "Int_stable ?E"
+    by (auto simp: Int_stable_def box_Int_box)
+
+  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
+    by (simp_all add: borel_eq_box sets_eq)
+
+  let ?A = "\<lambda>n::nat. box (- (real n *\<^sub>R One)) (real n *\<^sub>R One) :: 'a set"
+  show "range ?A \<subseteq> ?E" "(\<Union>i. ?A i) = UNIV"
+    unfolding UN_box_eq_UNIV by auto
+
+  { fix i show "emeasure lborel (?A i) \<noteq> \<infinity>" by auto }
+  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
+      apply (auto simp: emeasure_eq emeasure_lborel_box_eq )
+      apply (subst box_eq_empty(1)[THEN iffD2])
+      apply (auto intro: less_imp_le simp: not_le)
+      done }
+qed
+
+lemma lborel_affine:
+  fixes t :: "'a::euclidean_space" assumes "c \<noteq> 0"
+  shows "lborel = density (distr lborel borel (\<lambda>x. t + c *\<^sub>R x)) (\<lambda>_. \<bar>c\<bar>^DIM('a))" (is "_ = ?D")
+proof (rule lborel_eqI)
+  let ?B = "Basis :: 'a set"
+  fix l u assume le: "\<And>b. b \<in> ?B \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
+  show "emeasure ?D (box l u) = (\<Prod>b\<in>?B. (u - l) \<bullet> b)"
+  proof cases
+    assume "0 < c"
+    then have "(\<lambda>x. t + c *\<^sub>R x) -` box l u = box ((l - t) /\<^sub>R c) ((u - t) /\<^sub>R c)"
+      by (auto simp: field_simps box_def inner_simps)
+    with \<open>0 < c\<close> show ?thesis
+      using le
+      by (auto simp: field_simps inner_simps setprod_dividef setprod_constant setprod_nonneg
+                     ennreal_mult[symmetric] emeasure_density nn_integral_distr emeasure_distr
+                     nn_integral_cmult emeasure_lborel_box_eq borel_measurable_indicator')
+  next
+    assume "\<not> 0 < c" with \<open>c \<noteq> 0\<close> have "c < 0" by auto
+    then have "box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c) = (\<lambda>x. t + c *\<^sub>R x) -` box l u"
+      by (auto simp: field_simps box_def inner_simps)
+    then have *: "\<And>x. indicator (box l u) (t + c *\<^sub>R x) = (indicator (box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c)) x :: ennreal)"
+      by (auto split: split_indicator)
+    have **: "(\<Prod>x\<in>Basis. (l \<bullet> x - u \<bullet> x) / c) = (\<Prod>x\<in>Basis. u \<bullet> x - l \<bullet> x) / (-c) ^ card (Basis::'a set)"
+      using \<open>c < 0\<close>
+      by (auto simp add: field_simps setprod_dividef[symmetric] setprod_constant[symmetric]
+               intro!: setprod.cong)
+    show ?thesis
+      using \<open>c < 0\<close> le
+      by (auto simp: * ** field_simps emeasure_density nn_integral_distr nn_integral_cmult
+                     emeasure_lborel_box_eq inner_simps setprod_nonneg ennreal_mult[symmetric]
+                     borel_measurable_indicator')
+  qed
+qed simp
+
+lemma lborel_real_affine:
+  "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. ennreal (abs c))"
+  using lborel_affine[of c t] by simp
+
+lemma AE_borel_affine:
+  fixes P :: "real \<Rightarrow> bool"
+  shows "c \<noteq> 0 \<Longrightarrow> Measurable.pred borel P \<Longrightarrow> AE x in lborel. P x \<Longrightarrow> AE x in lborel. P (t + c * x)"
+  by (subst lborel_real_affine[where t="- t / c" and c="1 / c"])
+     (simp_all add: AE_density AE_distr_iff field_simps)
+
+lemma nn_integral_real_affine:
+  fixes c :: real assumes [measurable]: "f \<in> borel_measurable borel" and c: "c \<noteq> 0"
+  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = \<bar>c\<bar> * (\<integral>\<^sup>+x. f (t + c * x) \<partial>lborel)"
+  by (subst lborel_real_affine[OF c, of t])
+     (simp add: nn_integral_density nn_integral_distr nn_integral_cmult)
+
+lemma lborel_integrable_real_affine:
+  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  assumes f: "integrable lborel f"
+  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x))"
+  using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
+  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top)
+
+lemma lborel_integrable_real_affine_iff:
+  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x)) \<longleftrightarrow> integrable lborel f"
+  using
+    lborel_integrable_real_affine[of f c t]
+    lborel_integrable_real_affine[of "\<lambda>x. f (t + c * x)" "1/c" "-t/c"]
+  by (auto simp add: field_simps)
+
+lemma lborel_integral_real_affine:
+  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" and c :: real
+  assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lborel) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f (t + c * x) \<partial>lborel)"
+proof cases
+  assume f[measurable]: "integrable lborel f" then show ?thesis
+    using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t]
+    by (subst lborel_real_affine[OF c, of t])
+       (simp add: integral_density integral_distr)
+next
+  assume "\<not> integrable lborel f" with c show ?thesis
+    by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq)
+qed
+
+lemma divideR_right:
+  fixes x y :: "'a::real_normed_vector"
+  shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x"
+  using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp
+
+lemma lborel_has_bochner_integral_real_affine_iff:
+  fixes x :: "'a :: {banach, second_countable_topology}"
+  shows "c \<noteq> 0 \<Longrightarrow>
+    has_bochner_integral lborel f x \<longleftrightarrow>
+    has_bochner_integral lborel (\<lambda>x. f (t + c * x)) (x /\<^sub>R \<bar>c\<bar>)"
+  unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff
+  by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong)
+
+lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)"
+  by (subst lborel_real_affine[of "-1" 0])
+     (auto simp: density_1 one_ennreal_def[symmetric])
+
+lemma lborel_distr_mult:
+  assumes "(c::real) \<noteq> 0"
+  shows "distr lborel borel (op * c) = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
+proof-
+  have "distr lborel borel (op * c) = distr lborel lborel (op * c)" by (simp cong: distr_cong)
+  also from assms have "... = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
+    by (subst lborel_real_affine[of "inverse c" 0]) (auto simp: o_def distr_density_distr)
+  finally show ?thesis .
+qed
+
+lemma lborel_distr_mult':
+  assumes "(c::real) \<noteq> 0"
+  shows "lborel = density (distr lborel borel (op * c)) (\<lambda>_. \<bar>c\<bar>)"
+proof-
+  have "lborel = density lborel (\<lambda>_. 1)" by (rule density_1[symmetric])
+  also from assms have "(\<lambda>_. 1 :: ennreal) = (\<lambda>_. inverse \<bar>c\<bar> * \<bar>c\<bar>)" by (intro ext) simp
+  also have "density lborel ... = density (density lborel (\<lambda>_. inverse \<bar>c\<bar>)) (\<lambda>_. \<bar>c\<bar>)"
+    by (subst density_density_eq) (auto simp: ennreal_mult)
+  also from assms have "density lborel (\<lambda>_. inverse \<bar>c\<bar>) = distr lborel borel (op * c)"
+    by (rule lborel_distr_mult[symmetric])
+  finally show ?thesis .
+qed
+
+lemma lborel_distr_plus: "distr lborel borel (op + c) = (lborel :: real measure)"
+  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ennreal_def[symmetric])
+
+interpretation lborel: sigma_finite_measure lborel
+  by (rule sigma_finite_lborel)
+
+interpretation lborel_pair: pair_sigma_finite lborel lborel ..
+
+lemma lborel_prod:
+  "lborel \<Otimes>\<^sub>M lborel = (lborel :: ('a::euclidean_space \<times> 'b::euclidean_space) measure)"
+proof (rule lborel_eqI[symmetric], clarify)
+  fix la ua :: 'a and lb ub :: 'b
+  assume lu: "\<And>a b. (a, b) \<in> Basis \<Longrightarrow> (la, lb) \<bullet> (a, b) \<le> (ua, ub) \<bullet> (a, b)"
+  have [simp]:
+    "\<And>b. b \<in> Basis \<Longrightarrow> la \<bullet> b \<le> ua \<bullet> b"
+    "\<And>b. b \<in> Basis \<Longrightarrow> lb \<bullet> b \<le> ub \<bullet> b"
+    "inj_on (\<lambda>u. (u, 0)) Basis" "inj_on (\<lambda>u. (0, u)) Basis"
+    "(\<lambda>u. (u, 0)) ` Basis \<inter> (\<lambda>u. (0, u)) ` Basis = {}"
+    "box (la, lb) (ua, ub) = box la ua \<times> box lb ub"
+    using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def)
+  show "emeasure (lborel \<Otimes>\<^sub>M lborel) (box (la, lb) (ua, ub)) =
+      ennreal (setprod (op \<bullet> ((ua, ub) - (la, lb))) Basis)"
+    by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def setprod.union_disjoint
+                  setprod.reindex ennreal_mult inner_diff_left setprod_nonneg)
+qed (simp add: borel_prod[symmetric])
+
+(* FIXME: conversion in measurable prover *)
+lemma lborelD_Collect[measurable (raw)]: "{x\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" by simp
+lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
+
+subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
+
+lemma has_integral_measure_lborel:
+  fixes A :: "'a::euclidean_space set"
+  assumes A[measurable]: "A \<in> sets borel" and finite: "emeasure lborel A < \<infinity>"
+  shows "((\<lambda>x. 1) has_integral measure lborel A) A"
+proof -
+  { fix l u :: 'a
+    have "((\<lambda>x. 1) has_integral measure lborel (box l u)) (box l u)"
+    proof cases
+      assume "\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b"
+      then show ?thesis
+        apply simp
+        apply (subst has_integral_restrict[symmetric, OF box_subset_cbox])
+        apply (subst has_integral_spike_interior_eq[where g="\<lambda>_. 1"])
+        using has_integral_const[of "1::real" l u]
+        apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases)
+        done
+    next
+      assume "\<not> (\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b)"
+      then have "box l u = {}"
+        unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le)
+      then show ?thesis
+        by simp
+    qed }
+  note has_integral_box = this
+
+  { fix a b :: 'a let ?M = "\<lambda>A. measure lborel (A \<inter> box a b)"
+    have "Int_stable  (range (\<lambda>(a, b). box a b))"
+      by (auto simp: Int_stable_def box_Int_box)
+    moreover have "(range (\<lambda>(a, b). box a b)) \<subseteq> Pow UNIV"
+      by auto
+    moreover have "A \<in> sigma_sets UNIV (range (\<lambda>(a, b). box a b))"
+       using A unfolding borel_eq_box by simp
+    ultimately have "((\<lambda>x. 1) has_integral ?M A) (A \<inter> box a b)"
+    proof (induction rule: sigma_sets_induct_disjoint)
+      case (basic A) then show ?case
+        by (auto simp: box_Int_box has_integral_box)
+    next
+      case empty then show ?case
+        by simp
+    next
+      case (compl A)
+      then have [measurable]: "A \<in> sets borel"
+        by (simp add: borel_eq_box)
+
+      have "((\<lambda>x. 1) has_integral ?M (box a b)) (box a b)"
+        by (simp add: has_integral_box)
+      moreover have "((\<lambda>x. if x \<in> A \<inter> box a b then 1 else 0) has_integral ?M A) (box a b)"
+        by (subst has_integral_restrict) (auto intro: compl)
+      ultimately have "((\<lambda>x. 1 - (if x \<in> A \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
+        by (rule has_integral_sub)
+      then have "((\<lambda>x. (if x \<in> (UNIV - A) \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
+        by (rule has_integral_cong[THEN iffD1, rotated 1]) auto
+      then have "((\<lambda>x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \<inter> box a b)"
+        by (subst (asm) has_integral_restrict) auto
+      also have "?M (box a b) - ?M A = ?M (UNIV - A)"
+        by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2)
+      finally show ?case .
+    next
+      case (union F)
+      then have [measurable]: "\<And>i. F i \<in> sets borel"
+        by (simp add: borel_eq_box subset_eq)
+      have "((\<lambda>x. if x \<in> UNION UNIV F \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)"
+      proof (rule has_integral_monotone_convergence_increasing)
+        let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real"
+        show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)"
+          using union.IH by (auto intro!: has_integral_setsum simp del: Int_iff)
+        show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
+          by (intro setsum_mono2) auto
+        from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i"
+          by (auto simp add: disjoint_family_on_def)
+        show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)"
+          apply (auto simp: * setsum.If_cases Iio_Int_singleton)
+          apply (rule_tac k="Suc xa" in LIMSEQ_offset)
+          apply simp
+          done
+        have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)"
+          by (intro emeasure_mono) auto
+
+        with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
+          unfolding sums_def[symmetric] UN_extend_simps
+          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique)
+      qed
+      then show ?case
+        by (subst (asm) has_integral_restrict) auto
+    qed }
+  note * = this
+
+  show ?thesis
+  proof (rule has_integral_monotone_convergence_increasing)
+    let ?B = "\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set"
+    let ?f = "\<lambda>n::nat. \<lambda>x. if x \<in> A \<inter> ?B n then 1 else 0 :: real"
+    let ?M = "\<lambda>n. measure lborel (A \<inter> ?B n)"
+
+    show "\<And>n::nat. (?f n has_integral ?M n) A"
+      using * by (subst has_integral_restrict) simp_all
+    show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
+      by (auto simp: box_def)
+    { fix x assume "x \<in> A"
+      moreover have "(\<lambda>k. indicator (A \<inter> ?B k) x :: real) \<longlonglongrightarrow> indicator (\<Union>k::nat. A \<inter> ?B k) x"
+        by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def)
+      ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) \<longlonglongrightarrow> 1"
+        by (simp add: indicator_def UN_box_eq_UNIV) }
+
+    have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)"
+      by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
+    also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))"
+    proof (intro ext emeasure_eq_ennreal_measure)
+      fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)"
+        by (intro emeasure_mono) auto
+      then show "emeasure lborel (A \<inter> ?B n) \<noteq> top"
+        by (auto simp: top_unique)
+    qed
+    finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> measure lborel A"
+      using emeasure_eq_ennreal_measure[of lborel A] finite
+      by (simp add: UN_box_eq_UNIV less_top)
+  qed
+qed
+
+lemma nn_integral_has_integral:
+  fixes f::"'a::euclidean_space \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
+  shows "(f has_integral r) UNIV"
+using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
+  case (set A)
+  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
+    by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
+  with set show ?case
+    by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
+next
+  case (mult g c)
+  then have "ennreal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal r"
+    by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult)
+  with \<open>0 \<le> r\<close> \<open>0 \<le> c\<close>
+  obtain r' where "(c = 0 \<and> r = 0) \<or> (0 \<le> r' \<and> (\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel) = ennreal r' \<and> r = c * r')"
+    by (cases "\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel" rule: ennreal_cases)
+       (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric])
+  with mult show ?case
+    by (auto intro!: has_integral_cmult_real)
+next
+  case (add g h)
+  then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
+    by (simp add: nn_integral_add)
+  with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
+    by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
+       (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
+  with add show ?case
+    by (auto intro!: has_integral_add)
+next
+  case (seq U)
+  note seq(1)[measurable] and f[measurable]
+
+  { fix i x
+    have "U i x \<le> f x"
+      using seq(5)
+      apply (rule LIMSEQ_le_const)
+      using seq(4)
+      apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
+      done }
+  note U_le_f = this
+
+  { fix i
+    have "(\<integral>\<^sup>+x. U i x \<partial>lborel) \<le> (\<integral>\<^sup>+x. f x \<partial>lborel)"
+      using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp
+    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p" "p \<le> r" "0 \<le> p"
+      using seq(6) \<open>0\<le>r\<close> by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel" rule: ennreal_cases) (auto simp: top_unique)
+    moreover note seq
+    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
+      by auto }
+  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ennreal (U i x) \<partial>lborel) = ennreal (p i)"
+    and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
+    and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
+
+  have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)
+
+  have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) \<longlonglongrightarrow> integral UNIV f"
+  proof (rule monotone_convergence_increasing)
+    show "\<forall>k. U k integrable_on UNIV" using U_int by auto
+    show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using \<open>incseq U\<close> by (auto simp: incseq_def le_fun_def)
+    then show "bounded {integral UNIV (U k) |k. True}"
+      using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
+    show "\<forall>x\<in>UNIV. (\<lambda>k. U k x) \<longlonglongrightarrow> f x"
+      using seq by auto
+  qed
+  moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. f x \<partial>lborel)"
+    using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
+  ultimately have "integral UNIV f = r"
+    by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique)
+  with * show ?case
+    by (simp add: has_integral_integral)
+qed
+
+lemma nn_integral_lborel_eq_integral:
+  fixes f::"'a::euclidean_space \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
+  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f"
+proof -
+  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
+    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
+  then show ?thesis
+    using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
+qed
+
+lemma nn_integral_integrable_on:
+  fixes f::"'a::euclidean_space \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
+  shows "f integrable_on UNIV"
+proof -
+  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
+    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
+  then show ?thesis
+    by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
+qed
+
+lemma nn_integral_has_integral_lborel:
+  fixes f :: "'a::euclidean_space \<Rightarrow> real"
+  assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"
+  assumes I: "(f has_integral I) UNIV"
+  shows "integral\<^sup>N lborel f = I"
+proof -
+  from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto
+  from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
+  let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set"
+
+  note F(1)[THEN borel_measurable_simple_function, measurable]
+
+  have "0 \<le> I"
+    using I by (rule has_integral_nonneg) (simp add: nonneg)
+
+  have F_le_f: "enn2real (F i x) \<le> f x" for i x
+    using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\<lambda>i. F i x"]
+    by (cases "F i x" rule: ennreal_cases) auto
+  let ?F = "\<lambda>i x. F i x * indicator (?B i) x"
+  have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
+  proof (subst nn_integral_monotone_convergence_SUP[symmetric])
+    { fix x
+      obtain j where j: "x \<in> ?B j"
+        using UN_box_eq_UNIV by auto
+
+      have "ennreal (f x) = (SUP i. F i x)"
+        using F(4)[of x] nonneg[of x] by (simp add: max_def)
+      also have "\<dots> = (SUP i. ?F i x)"
+      proof (rule SUP_eq)
+        fix i show "\<exists>j\<in>UNIV. F i x \<le> ?F j x"
+          using j F(2)
+          by (intro bexI[of _ "max i j"])
+             (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
+      qed (auto intro!: F split: split_indicator)
+      finally have "ennreal (f x) =  (SUP i. ?F i x)" . }
+    then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
+      by simp
+  qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
+  also have "\<dots> \<le> ennreal I"
+  proof (rule SUP_least)
+    fix i :: nat
+    have finite_F: "(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
+    proof (rule nn_integral_bound_simple_function)
+      have "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
+        emeasure lborel (?B i)"
+        by (intro emeasure_mono)  (auto split: split_indicator)
+      then show "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
+        by (auto simp: less_top[symmetric] top_unique)
+    qed (auto split: split_indicator
+              intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal)
+
+    have int_F: "(\<lambda>x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV"
+      using F(4) finite_F
+      by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg)
+
+    have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) =
+      (\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel)"
+      using F(3,4)
+      by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator)
+    also have "\<dots> = ennreal (integral UNIV (\<lambda>x. enn2real (F i x) * indicator (?B i) x))"
+      using F
+      by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
+         (auto split: split_indicator intro: enn2real_nonneg)
+    also have "\<dots> \<le> ennreal I"
+      by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
+               simp: \<open>0 \<le> I\<close> split: split_indicator )
+    finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ennreal I" .
+  qed
+  finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) < \<infinity>"
+    by (auto simp: less_top[symmetric] top_unique)
+  from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
+    by (simp add: integral_unique)
+qed
+
+lemma has_integral_iff_emeasure_lborel:
+  fixes A :: "'a::euclidean_space set"
+  assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r"
+  shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r"
+proof (cases "emeasure lborel A = \<infinity>")
+  case emeasure_A: True
+  have "\<not> (\<lambda>x. 1::real) integrable_on A"
+  proof
+    assume int: "(\<lambda>x. 1::real) integrable_on A"
+    then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV"
+      unfolding indicator_def[abs_def] integrable_restrict_univ .
+    then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV"
+      by auto
+    from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
+      by (simp add: ennreal_indicator)
+  qed
+  with emeasure_A show ?thesis
+    by auto
+next
+  case False
+  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
+    by (simp add: has_integral_measure_lborel less_top)
+  with False show ?thesis
+    by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
+qed
+
+lemma has_integral_integral_real:
+  fixes f::"'a::euclidean_space \<Rightarrow> real"
+  assumes f: "integrable lborel f"
+  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
+using f proof induct
+  case (base A c) then show ?case
+    by (auto intro!: has_integral_mult_left simp: )
+       (simp add: emeasure_eq_ennreal_measure indicator_def has_integral_measure_lborel)
+next
+  case (add f g) then show ?case
+    by (auto intro!: has_integral_add)
+next
+  case (lim f s)
+  show ?case
+  proof (rule has_integral_dominated_convergence)
+    show "\<And>i. (s i has_integral integral\<^sup>L lborel (s i)) UNIV" by fact
+    show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV"
+      using \<open>integrable lborel f\<close>
+      by (intro nn_integral_integrable_on)
+         (auto simp: integrable_iff_bounded abs_mult  nn_integral_cmult ennreal_mult ennreal_mult_less_top)
+    show "\<And>k. \<forall>x\<in>UNIV. norm (s k x) \<le> norm (2 * f x)"
+      using lim by (auto simp add: abs_mult)
+    show "\<forall>x\<in>UNIV. (\<lambda>k. s k x) \<longlonglongrightarrow> f x"
+      using lim by auto
+    show "(\<lambda>k. integral\<^sup>L lborel (s k)) \<longlonglongrightarrow> integral\<^sup>L lborel f"
+      using lim lim(1)[THEN borel_measurable_integrable]
+      by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto
+  qed
+qed
+
+context
+  fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
+begin
+
+lemma has_integral_integral_lborel:
+  assumes f: "integrable lborel f"
+  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
+proof -
+  have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
+    using f by (intro has_integral_setsum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
+  also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
+    by (simp add: fun_eq_iff euclidean_representation)
+  also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f"
+    using f by (subst (2) eq_f[symmetric]) simp
+  finally show ?thesis .
+qed
+
+lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV"
+  using has_integral_integral_lborel by auto
+
+lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)"
+  using has_integral_integral_lborel by auto
+
+end
+
+subsection \<open>Fundamental Theorem of Calculus for the Lebesgue integral\<close>
+
+lemma emeasure_bounded_finite:
+  assumes "bounded A" shows "emeasure lborel A < \<infinity>"
+proof -
+  from bounded_subset_cbox[OF \<open>bounded A\<close>] obtain a b where "A \<subseteq> cbox a b"
+    by auto
+  then have "emeasure lborel A \<le> emeasure lborel (cbox a b)"
+    by (intro emeasure_mono) auto
+  then show ?thesis
+    by (auto simp: emeasure_lborel_cbox_eq setprod_nonneg less_top[symmetric] top_unique split: if_split_asm)
+qed
+
+lemma emeasure_compact_finite: "compact A \<Longrightarrow> emeasure lborel A < \<infinity>"
+  using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded)
+
+lemma borel_integrable_compact:
+  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes "compact S" "continuous_on S f"
+  shows "integrable lborel (\<lambda>x. indicator S x *\<^sub>R f x)"
+proof cases
+  assume "S \<noteq> {}"
+  have "continuous_on S (\<lambda>x. norm (f x))"
+    using assms by (intro continuous_intros)
+  from continuous_attains_sup[OF \<open>compact S\<close> \<open>S \<noteq> {}\<close> this]
+  obtain M where M: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> M"
+    by auto
+
+  show ?thesis
+  proof (rule integrable_bound)
+    show "integrable lborel (\<lambda>x. indicator S x * M)"
+      using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left)
+    show "(\<lambda>x. indicator S x *\<^sub>R f x) \<in> borel_measurable lborel"
+      using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact)
+    show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \<le> norm (indicator S x * M)"
+      by (auto split: split_indicator simp: abs_real_def dest!: M)
+  qed
+qed simp
+
+lemma borel_integrable_atLeastAtMost:
+  fixes f :: "real \<Rightarrow> real"
+  assumes f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
+  shows "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is "integrable _ ?f")
+proof -
+  have "integrable lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x)"
+  proof (rule borel_integrable_compact)
+    from f show "continuous_on {a..b} f"
+      by (auto intro: continuous_at_imp_continuous_on)
+  qed simp
+  then show ?thesis
+    by (auto simp: mult.commute)
+qed
+
+text \<open>
+
+For the positive integral we replace continuity with Borel-measurability.
+
+\<close>
+
+lemma
+  fixes f :: "real \<Rightarrow> real"
+  assumes [measurable]: "f \<in> borel_measurable borel"
+  assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
+  shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
+    and has_bochner_integral_FTC_Icc_nonneg:
+      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
+    and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
+    and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
+proof -
+  have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
+    using f(2) by (auto split: split_indicator)
+
+  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b\<Longrightarrow> F x \<le> F y" for x y
+    using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
+
+  have "(f has_integral F b - F a) {a..b}"
+    by (intro fundamental_theorem_of_calculus)
+       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
+             intro: has_field_derivative_subset[OF f(1)] \<open>a \<le> b\<close>)
+  then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
+    unfolding indicator_def if_distrib[where f="\<lambda>x. a * x" for a]
+    by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
+  then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
+    by (rule nn_integral_has_integral_lborel[OF *])
+  then show ?has
+    by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \<open>a \<le> b\<close>)
+  then show ?eq ?int
+    unfolding has_bochner_integral_iff by auto
+  show ?nn
+    by (subst nn[symmetric])
+       (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator)
+qed
+
+lemma
+  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
+  assumes "a \<le> b"
+  assumes "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
+  assumes cont: "continuous_on {a .. b} f"
+  shows has_bochner_integral_FTC_Icc:
+      "has_bochner_integral lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has)
+    and integral_FTC_Icc: "(\<integral>x. indicator {a .. b} x *\<^sub>R f x \<partial>lborel) = F b - F a" (is ?eq)
+proof -
+  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
+  have int: "integrable lborel ?f"
+    using borel_integrable_compact[OF _ cont] by auto
+  have "(f has_integral F b - F a) {a..b}"
+    using assms(1,2) by (intro fundamental_theorem_of_calculus) auto
+  moreover
+  have "(f has_integral integral\<^sup>L lborel ?f) {a..b}"
+    using has_integral_integral_lborel[OF int]
+    unfolding indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a]
+    by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
+  ultimately show ?eq
+    by (auto dest: has_integral_unique)
+  then show ?has
+    using int by (auto simp: has_bochner_integral_iff)
+qed
+
+lemma
+  fixes f :: "real \<Rightarrow> real"
+  assumes "a \<le> b"
+  assumes deriv: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
+  assumes cont: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
+  shows has_bochner_integral_FTC_Icc_real:
+      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
+    and integral_FTC_Icc_real: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
+proof -
+  have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
+    unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
+    using deriv by (auto intro: DERIV_subset)
+  have 2: "continuous_on {a .. b} f"
+    using cont by (intro continuous_at_imp_continuous_on) auto
+  show ?has ?eq
+    using has_bochner_integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2]
+    by (auto simp: mult.commute)
+qed
+
+lemma nn_integral_FTC_atLeast:
+  fixes f :: "real \<Rightarrow> real"
+  assumes f_borel: "f \<in> borel_measurable borel"
+  assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x"
+  assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
+  assumes lim: "(F \<longlongrightarrow> T) at_top"
+  shows "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
+proof -
+  let ?f = "\<lambda>(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
+  let ?fR = "\<lambda>x. ennreal (f x) * indicator {a ..} x"
+
+  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> F x \<le> F y" for x y
+    using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
+  then have F_le_T: "a \<le> x \<Longrightarrow> F x \<le> T" for x
+    by (intro tendsto_le_const[OF _ lim])
+       (auto simp: trivial_limit_at_top_linorder eventually_at_top_linorder)
+
+  have "(SUP i::nat. ?f i x) = ?fR x" for x
+  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
+    from reals_Archimedean2[of "x - a"] guess n ..
+    then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
+      by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
+    then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x"
+      by (rule Lim_eventually)
+  qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator)
+  then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
+    by simp
+  also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
+  proof (rule nn_integral_monotone_convergence_SUP)
+    show "incseq ?f"
+      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
+    show "\<And>i. (?f i) \<in> borel_measurable lborel"
+      using f_borel by auto
+  qed
+  also have "\<dots> = (SUP i::nat. ennreal (F (a + real i) - F a))"
+    by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
+  also have "\<dots> = T - F a"
+  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
+    have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T"
+      apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
+      apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
+      apply (rule filterlim_real_sequentially)
+      done
+    then show "(\<lambda>n. ennreal (F (a + real n) - F a)) \<longlonglongrightarrow> ennreal (T - F a)"
+      by (simp add: F_mono F_le_T tendsto_diff)
+  qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono)
+  finally show ?thesis .
+qed
+
+lemma integral_power:
+  "a \<le> b \<Longrightarrow> (\<integral>x. x^k * indicator {a..b} x \<partial>lborel) = (b^Suc k - a^Suc k) / Suc k"
+proof (subst integral_FTC_Icc_real)
+  fix x show "DERIV (\<lambda>x. x^Suc k / Suc k) x :> x^k"
+    by (intro derivative_eq_intros) auto
+qed (auto simp: field_simps simp del: of_nat_Suc)
+
+subsection \<open>Integration by parts\<close>
+
+lemma integral_by_parts_integrable:
+  fixes f g F G::"real \<Rightarrow> real"
+  assumes "a \<le> b"
+  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
+  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
+  assumes [intro]: "!!x. DERIV F x :> f x"
+  assumes [intro]: "!!x. DERIV G x :> g x"
+  shows  "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
+  by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont)
+
+lemma integral_by_parts:
+  fixes f g F G::"real \<Rightarrow> real"
+  assumes [arith]: "a \<le> b"
+  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
+  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
+  assumes [intro]: "!!x. DERIV F x :> f x"
+  assumes [intro]: "!!x. DERIV G x :> g x"
+  shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel)
+            =  F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
+proof-
+  have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
+    by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros)
+      (auto intro!: DERIV_isCont)
+
+  have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
+    (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
+    apply (subst integral_add[symmetric])
+    apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros)
+    by (auto intro!: DERIV_isCont integral_cong split:split_indicator)
+
+  thus ?thesis using 0 by auto
+qed
+
+lemma integral_by_parts':
+  fixes f g F G::"real \<Rightarrow> real"
+  assumes "a \<le> b"
+  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
+  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
+  assumes "!!x. DERIV F x :> f x"
+  assumes "!!x. DERIV G x :> g x"
+  shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel)
+            =  F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel"
+  using integral_by_parts[OF assms] by (simp add: ac_simps)
+
+lemma has_bochner_integral_even_function:
+  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
+  assumes even: "\<And>x. f (- x) = f x"
+  shows "has_bochner_integral lborel f (2 *\<^sub>R x)"
+proof -
+  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
+    by (auto split: split_indicator)
+  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x"
+    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
+       (auto simp: indicator even f)
+  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)"
+    by (rule has_bochner_integral_add)
+  then have "has_bochner_integral lborel f (x + x)"
+    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
+       (auto split: split_indicator)
+  then show ?thesis
+    by (simp add: scaleR_2)
+qed
+
+lemma has_bochner_integral_odd_function:
+  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
+  assumes odd: "\<And>x. f (- x) = - f x"
+  shows "has_bochner_integral lborel f 0"
+proof -
+  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
+    by (auto split: split_indicator)
+  have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x"
+    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
+       (auto simp: indicator odd f)
+  from has_bochner_integral_minus[OF this]
+  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)"
+    by simp
+  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)"
+    by (rule has_bochner_integral_add)
+  then have "has_bochner_integral lborel f (x + - x)"
+    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
+       (auto split: split_indicator)
+  then show ?thesis
+    by simp
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Measurable.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,624 @@
+(*  Title:      HOL/Probability/Measurable.thy
+    Author:     Johannes Hölzl <hoelzl@in.tum.de>
+*)
+theory Measurable
+  imports
+    Sigma_Algebra
+    "~~/src/HOL/Library/Order_Continuity"
+begin
+
+subsection \<open>Measurability prover\<close>
+
+lemma (in algebra) sets_Collect_finite_All:
+  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
+  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (if S = {} then \<Omega> else \<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
+    by auto
+  with assms show ?thesis by (auto intro!: sets_Collect_finite_All')
+qed
+
+abbreviation "pred M P \<equiv> P \<in> measurable M (count_space (UNIV::bool set))"
+
+lemma pred_def: "pred M P \<longleftrightarrow> {x\<in>space M. P x} \<in> sets M"
+proof
+  assume "pred M P"
+  then have "P -` {True} \<inter> space M \<in> sets M"
+    by (auto simp: measurable_count_space_eq2)
+  also have "P -` {True} \<inter> space M = {x\<in>space M. P x}" by auto
+  finally show "{x\<in>space M. P x} \<in> sets M" .
+next
+  assume P: "{x\<in>space M. P x} \<in> sets M"
+  moreover
+  { fix X
+    have "X \<in> Pow (UNIV :: bool set)" by simp
+    then have "P -` X \<inter> space M = {x\<in>space M. ((X = {True} \<longrightarrow> P x) \<and> (X = {False} \<longrightarrow> \<not> P x) \<and> X \<noteq> {})}"
+      unfolding UNIV_bool Pow_insert Pow_empty by auto
+    then have "P -` X \<inter> space M \<in> sets M"
+      by (auto intro!: sets.sets_Collect_neg sets.sets_Collect_imp sets.sets_Collect_conj sets.sets_Collect_const P) }
+  then show "pred M P"
+    by (auto simp: measurable_def)
+qed
+
+lemma pred_sets1: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> f \<in> measurable N M \<Longrightarrow> pred N (\<lambda>x. P (f x))"
+  by (rule measurable_compose[where f=f and N=M]) (auto simp: pred_def)
+
+lemma pred_sets2: "A \<in> sets N \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A)"
+  by (rule measurable_compose[where f=f and N=N]) (auto simp: pred_def Int_def[symmetric])
+
+ML_file "measurable.ML"
+
+attribute_setup measurable = \<open>
+  Scan.lift (
+    (Args.add >> K true || Args.del >> K false || Scan.succeed true) --
+    Scan.optional (Args.parens (
+      Scan.optional (Args.$$$ "raw" >> K true) false --
+      Scan.optional (Args.$$$ "generic" >> K Measurable.Generic) Measurable.Concrete))
+    (false, Measurable.Concrete) >>
+    Measurable.measurable_thm_attr)
+\<close> "declaration of measurability theorems"
+
+attribute_setup measurable_dest = Measurable.dest_thm_attr
+  "add dest rule to measurability prover"
+
+attribute_setup measurable_cong = Measurable.cong_thm_attr
+  "add congurence rules to measurability prover"
+
+method_setup measurable = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close>
+  "measurability prover"
+
+simproc_setup measurable ("A \<in> sets M" | "f \<in> measurable M N") = \<open>K Measurable.simproc\<close>
+
+setup \<open>
+  Global_Theory.add_thms_dynamic (@{binding measurable}, Measurable.get_all)
+\<close>
+
+declare
+  pred_sets1[measurable_dest]
+  pred_sets2[measurable_dest]
+  sets.sets_into_space[measurable_dest]
+
+declare
+  sets.top[measurable]
+  sets.empty_sets[measurable (raw)]
+  sets.Un[measurable (raw)]
+  sets.Diff[measurable (raw)]
+
+declare
+  measurable_count_space[measurable (raw)]
+  measurable_ident[measurable (raw)]
+  measurable_id[measurable (raw)]
+  measurable_const[measurable (raw)]
+  measurable_If[measurable (raw)]
+  measurable_comp[measurable (raw)]
+  measurable_sets[measurable (raw)]
+
+declare measurable_cong_sets[measurable_cong]
+declare sets_restrict_space_cong[measurable_cong]
+declare sets_restrict_UNIV[measurable_cong]
+
+lemma predE[measurable (raw)]:
+  "pred M P \<Longrightarrow> {x\<in>space M. P x} \<in> sets M"
+  unfolding pred_def .
+
+lemma pred_intros_imp'[measurable (raw)]:
+  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<longrightarrow> P x)"
+  by (cases K) auto
+
+lemma pred_intros_conj1'[measurable (raw)]:
+  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<and> P x)"
+  by (cases K) auto
+
+lemma pred_intros_conj2'[measurable (raw)]:
+  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<and> K)"
+  by (cases K) auto
+
+lemma pred_intros_disj1'[measurable (raw)]:
+  "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<or> P x)"
+  by (cases K) auto
+
+lemma pred_intros_disj2'[measurable (raw)]:
+  "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<or> K)"
+  by (cases K) auto
+
+lemma pred_intros_logic[measurable (raw)]:
+  "pred M (\<lambda>x. x \<in> space M)"
+  "pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. \<not> P x)"
+  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<and> P x)"
+  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<longrightarrow> P x)"
+  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<or> P x)"
+  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x = P x)"
+  "pred M (\<lambda>x. f x \<in> UNIV)"
+  "pred M (\<lambda>x. f x \<in> {})"
+  "pred M (\<lambda>x. P' (f x) x) \<Longrightarrow> pred M (\<lambda>x. f x \<in> {y. P' y x})"
+  "pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> - (B x))"
+  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) - (B x))"
+  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<inter> (B x))"
+  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<union> (B x))"
+  "pred M (\<lambda>x. g x (f x) \<in> (X x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (g x) -` (X x))"
+  by (auto simp: iff_conv_conj_imp pred_def)
+
+lemma pred_intros_countable[measurable (raw)]:
+  fixes P :: "'a \<Rightarrow> 'i :: countable \<Rightarrow> bool"
+  shows
+    "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i. P x i)"
+    "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i. P x i)"
+  by (auto intro!: sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex simp: pred_def)
+
+lemma pred_intros_countable_bounded[measurable (raw)]:
+  fixes X :: "'i :: countable set"
+  shows
+    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>X. N x i))"
+    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>X. N x i))"
+    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
+    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
+  by simp_all (auto simp: Bex_def Ball_def)
+
+lemma pred_intros_finite[measurable (raw)]:
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>I. N x i))"
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>I. N x i))"
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>I. P x i)"
+  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>I. P x i)"
+  by (auto intro!: sets.sets_Collect_finite_Ex sets.sets_Collect_finite_All simp: iff_conv_conj_imp pred_def)
+
+lemma countable_Un_Int[measurable (raw)]:
+  "(\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Union>i\<in>I. N i) \<in> sets M"
+  "I \<noteq> {} \<Longrightarrow> (\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Inter>i\<in>I. N i) \<in> sets M"
+  by auto
+
+declare
+  finite_UN[measurable (raw)]
+  finite_INT[measurable (raw)]
+
+lemma sets_Int_pred[measurable (raw)]:
+  assumes space: "A \<inter> B \<subseteq> space M" and [measurable]: "pred M (\<lambda>x. x \<in> A)" "pred M (\<lambda>x. x \<in> B)"
+  shows "A \<inter> B \<in> sets M"
+proof -
+  have "{x\<in>space M. x \<in> A \<inter> B} \<in> sets M" by auto
+  also have "{x\<in>space M. x \<in> A \<inter> B} = A \<inter> B"
+    using space by auto
+  finally show ?thesis .
+qed
+
+lemma [measurable (raw generic)]:
+  assumes f: "f \<in> measurable M N" and c: "c \<in> space N \<Longrightarrow> {c} \<in> sets N"
+  shows pred_eq_const1: "pred M (\<lambda>x. f x = c)"
+    and pred_eq_const2: "pred M (\<lambda>x. c = f x)"
+proof -
+  show "pred M (\<lambda>x. f x = c)"
+  proof cases
+    assume "c \<in> space N"
+    with measurable_sets[OF f c] show ?thesis
+      by (auto simp: Int_def conj_commute pred_def)
+  next
+    assume "c \<notin> space N"
+    with f[THEN measurable_space] have "{x \<in> space M. f x = c} = {}" by auto
+    then show ?thesis by (auto simp: pred_def cong: conj_cong)
+  qed
+  then show "pred M (\<lambda>x. c = f x)"
+    by (simp add: eq_commute)
+qed
+
+lemma pred_count_space_const1[measurable (raw)]:
+  "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. f x = c)"
+  by (intro pred_eq_const1[where N="count_space UNIV"]) (auto )
+
+lemma pred_count_space_const2[measurable (raw)]:
+  "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. c = f x)"
+  by (intro pred_eq_const2[where N="count_space UNIV"]) (auto )
+
+lemma pred_le_const[measurable (raw generic)]:
+  assumes f: "f \<in> measurable M N" and c: "{.. c} \<in> sets N" shows "pred M (\<lambda>x. f x \<le> c)"
+  using measurable_sets[OF f c]
+  by (auto simp: Int_def conj_commute eq_commute pred_def)
+
+lemma pred_const_le[measurable (raw generic)]:
+  assumes f: "f \<in> measurable M N" and c: "{c ..} \<in> sets N" shows "pred M (\<lambda>x. c \<le> f x)"
+  using measurable_sets[OF f c]
+  by (auto simp: Int_def conj_commute eq_commute pred_def)
+
+lemma pred_less_const[measurable (raw generic)]:
+  assumes f: "f \<in> measurable M N" and c: "{..< c} \<in> sets N" shows "pred M (\<lambda>x. f x < c)"
+  using measurable_sets[OF f c]
+  by (auto simp: Int_def conj_commute eq_commute pred_def)
+
+lemma pred_const_less[measurable (raw generic)]:
+  assumes f: "f \<in> measurable M N" and c: "{c <..} \<in> sets N" shows "pred M (\<lambda>x. c < f x)"
+  using measurable_sets[OF f c]
+  by (auto simp: Int_def conj_commute eq_commute pred_def)
+
+declare
+  sets.Int[measurable (raw)]
+
+lemma pred_in_If[measurable (raw)]:
+  "(P \<Longrightarrow> pred M (\<lambda>x. x \<in> A x)) \<Longrightarrow> (\<not> P \<Longrightarrow> pred M (\<lambda>x. x \<in> B x)) \<Longrightarrow>
+    pred M (\<lambda>x. x \<in> (if P then A x else B x))"
+  by auto
+
+lemma sets_range[measurable_dest]:
+  "A ` I \<subseteq> sets M \<Longrightarrow> i \<in> I \<Longrightarrow> A i \<in> sets M"
+  by auto
+
+lemma pred_sets_range[measurable_dest]:
+  "A ` I \<subseteq> sets N \<Longrightarrow> i \<in> I \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
+  using pred_sets2[OF sets_range] by auto
+
+lemma sets_All[measurable_dest]:
+  "\<forall>i. A i \<in> sets (M i) \<Longrightarrow> A i \<in> sets (M i)"
+  by auto
+
+lemma pred_sets_All[measurable_dest]:
+  "\<forall>i. A i \<in> sets (N i) \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
+  using pred_sets2[OF sets_All, of A N f] by auto
+
+lemma sets_Ball[measurable_dest]:
+  "\<forall>i\<in>I. A i \<in> sets (M i) \<Longrightarrow> i\<in>I \<Longrightarrow> A i \<in> sets (M i)"
+  by auto
+
+lemma pred_sets_Ball[measurable_dest]:
+  "\<forall>i\<in>I. A i \<in> sets (N i) \<Longrightarrow> i\<in>I \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
+  using pred_sets2[OF sets_Ball, of _ _ _ f] by auto
+
+lemma measurable_finite[measurable (raw)]:
+  fixes S :: "'a \<Rightarrow> nat set"
+  assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
+  shows "pred M (\<lambda>x. finite (S x))"
+  unfolding finite_nat_set_iff_bounded by (simp add: Ball_def)
+
+lemma measurable_Least[measurable]:
+  assumes [measurable]: "(\<And>i::nat. (\<lambda>x. P i x) \<in> measurable M (count_space UNIV))"
+  shows "(\<lambda>x. LEAST i. P i x) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_def by (safe intro!: sets_Least) simp_all
+
+lemma measurable_Max_nat[measurable (raw)]:
+  fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
+  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
+  shows "(\<lambda>x. Max {i. P i x}) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_count_space_eq2_countable
+proof safe
+  fix n
+
+  { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
+    then have "infinite {i. P i x}"
+      unfolding infinite_nat_iff_unbounded_le by auto
+    then have "Max {i. P i x} = the None"
+      by (rule Max.infinite) }
+  note 1 = this
+
+  { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
+    then have "finite {i. P i x}"
+      by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded)
+    with \<open>P i x\<close> have "P (Max {i. P i x}) x" "i \<le> Max {i. P i x}" "finite {i. P i x}"
+      using Max_in[of "{i. P i x}"] by auto }
+  note 2 = this
+
+  have "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Max {i. P i x} = n}"
+    by auto
+  also have "\<dots> =
+    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
+      if (\<exists>i. P i x) then P n x \<and> (\<forall>i>n. \<not> P i x)
+      else Max {} = n}"
+    by (intro arg_cong[where f=Collect] ext conj_cong)
+       (auto simp add: 1 2 not_le[symmetric] intro!: Max_eqI)
+  also have "\<dots> \<in> sets M"
+    by measurable
+  finally show "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
+qed simp
+
+lemma measurable_Min_nat[measurable (raw)]:
+  fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
+  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
+  shows "(\<lambda>x. Min {i. P i x}) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_count_space_eq2_countable
+proof safe
+  fix n
+
+  { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
+    then have "infinite {i. P i x}"
+      unfolding infinite_nat_iff_unbounded_le by auto
+    then have "Min {i. P i x} = the None"
+      by (rule Min.infinite) }
+  note 1 = this
+
+  { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
+    then have "finite {i. P i x}"
+      by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded)
+    with \<open>P i x\<close> have "P (Min {i. P i x}) x" "Min {i. P i x} \<le> i" "finite {i. P i x}"
+      using Min_in[of "{i. P i x}"] by auto }
+  note 2 = this
+
+  have "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Min {i. P i x} = n}"
+    by auto
+  also have "\<dots> =
+    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
+      if (\<exists>i. P i x) then P n x \<and> (\<forall>i<n. \<not> P i x)
+      else Min {} = n}"
+    by (intro arg_cong[where f=Collect] ext conj_cong)
+       (auto simp add: 1 2 not_le[symmetric] intro!: Min_eqI)
+  also have "\<dots> \<in> sets M"
+    by measurable
+  finally show "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
+qed simp
+
+lemma measurable_count_space_insert[measurable (raw)]:
+  "s \<in> S \<Longrightarrow> A \<in> sets (count_space S) \<Longrightarrow> insert s A \<in> sets (count_space S)"
+  by simp
+
+lemma sets_UNIV [measurable (raw)]: "A \<in> sets (count_space UNIV)"
+  by simp
+
+lemma measurable_card[measurable]:
+  fixes S :: "'a \<Rightarrow> nat set"
+  assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
+  shows "(\<lambda>x. card (S x)) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_count_space_eq2_countable
+proof safe
+  fix n show "(\<lambda>x. card (S x)) -` {n} \<inter> space M \<in> sets M"
+  proof (cases n)
+    case 0
+    then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M = {x\<in>space M. infinite (S x) \<or> (\<forall>i. i \<notin> S x)}"
+      by auto
+    also have "\<dots> \<in> sets M"
+      by measurable
+    finally show ?thesis .
+  next
+    case (Suc i)
+    then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M =
+      (\<Union>F\<in>{A\<in>{A. finite A}. card A = n}. {x\<in>space M. (\<forall>i. i \<in> S x \<longleftrightarrow> i \<in> F)})"
+      unfolding set_eq_iff[symmetric] Collect_bex_eq[symmetric] by (auto intro: card_ge_0_finite)
+    also have "\<dots> \<in> sets M"
+      by (intro sets.countable_UN' countable_Collect countable_Collect_finite) auto
+    finally show ?thesis .
+  qed
+qed rule
+
+lemma measurable_pred_countable[measurable (raw)]:
+  assumes "countable X"
+  shows
+    "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
+    "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
+  unfolding pred_def
+  by (auto intro!: sets.sets_Collect_countable_All' sets.sets_Collect_countable_Ex' assms)
+
+subsection \<open>Measurability for (co)inductive predicates\<close>
+
+lemma measurable_bot[measurable]: "bot \<in> measurable M (count_space UNIV)"
+  by (simp add: bot_fun_def)
+
+lemma measurable_top[measurable]: "top \<in> measurable M (count_space UNIV)"
+  by (simp add: top_fun_def)
+
+lemma measurable_SUP[measurable]:
+  fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
+  assumes [simp]: "countable I"
+  assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)"
+  shows "(\<lambda>x. SUP i:I. F i x) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_count_space_eq2_countable
+proof (safe intro!: UNIV_I)
+  fix a
+  have "(\<lambda>x. SUP i:I. F i x) -` {a} \<inter> space M =
+    {x\<in>space M. (\<forall>i\<in>I. F i x \<le> a) \<and> (\<forall>b. (\<forall>i\<in>I. F i x \<le> b) \<longrightarrow> a \<le> b)}"
+    unfolding SUP_le_iff[symmetric] by auto
+  also have "\<dots> \<in> sets M"
+    by measurable
+  finally show "(\<lambda>x. SUP i:I. F i x) -` {a} \<inter> space M \<in> sets M" .
+qed
+
+lemma measurable_INF[measurable]:
+  fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
+  assumes [simp]: "countable I"
+  assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)"
+  shows "(\<lambda>x. INF i:I. F i x) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_count_space_eq2_countable
+proof (safe intro!: UNIV_I)
+  fix a
+  have "(\<lambda>x. INF i:I. F i x) -` {a} \<inter> space M =
+    {x\<in>space M. (\<forall>i\<in>I. a \<le> F i x) \<and> (\<forall>b. (\<forall>i\<in>I. b \<le> F i x) \<longrightarrow> b \<le> a)}"
+    unfolding le_INF_iff[symmetric] by auto
+  also have "\<dots> \<in> sets M"
+    by measurable
+  finally show "(\<lambda>x. INF i:I. F i x) -` {a} \<inter> space M \<in> sets M" .
+qed
+
+lemma measurable_lfp_coinduct[consumes 1, case_names continuity step]:
+  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
+  assumes "P M"
+  assumes F: "sup_continuous F"
+  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
+  shows "lfp F \<in> measurable M (count_space UNIV)"
+proof -
+  { fix i from \<open>P M\<close> have "((F ^^ i) bot) \<in> measurable M (count_space UNIV)"
+      by (induct i arbitrary: M) (auto intro!: *) }
+  then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> measurable M (count_space UNIV)"
+    by measurable
+  also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = lfp F"
+    by (subst sup_continuous_lfp) (auto intro: F)
+  finally show ?thesis .
+qed
+
+lemma measurable_lfp:
+  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
+  assumes F: "sup_continuous F"
+  assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
+  shows "lfp F \<in> measurable M (count_space UNIV)"
+  by (coinduction rule: measurable_lfp_coinduct[OF _ F]) (blast intro: *)
+
+lemma measurable_gfp_coinduct[consumes 1, case_names continuity step]:
+  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
+  assumes "P M"
+  assumes F: "inf_continuous F"
+  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
+  shows "gfp F \<in> measurable M (count_space UNIV)"
+proof -
+  { fix i from \<open>P M\<close> have "((F ^^ i) top) \<in> measurable M (count_space UNIV)"
+      by (induct i arbitrary: M) (auto intro!: *) }
+  then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> measurable M (count_space UNIV)"
+    by measurable
+  also have "(\<lambda>x. INF i. (F ^^ i) top x) = gfp F"
+    by (subst inf_continuous_gfp) (auto intro: F)
+  finally show ?thesis .
+qed
+
+lemma measurable_gfp:
+  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
+  assumes F: "inf_continuous F"
+  assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
+  shows "gfp F \<in> measurable M (count_space UNIV)"
+  by (coinduction rule: measurable_gfp_coinduct[OF _ F]) (blast intro: *)
+
+lemma measurable_lfp2_coinduct[consumes 1, case_names continuity step]:
+  fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
+  assumes "P M s"
+  assumes F: "sup_continuous F"
+  assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
+  shows "lfp F s \<in> measurable M (count_space UNIV)"
+proof -
+  { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
+      by (induct i arbitrary: M s) (auto intro!: *) }
+  then have "(\<lambda>x. SUP i. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
+    by measurable
+  also have "(\<lambda>x. SUP i. (F ^^ i) bot s x) = lfp F s"
+    by (subst sup_continuous_lfp) (auto simp: F)
+  finally show ?thesis .
+qed
+
+lemma measurable_gfp2_coinduct[consumes 1, case_names continuity step]:
+  fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
+  assumes "P M s"
+  assumes F: "inf_continuous F"
+  assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
+  shows "gfp F s \<in> measurable M (count_space UNIV)"
+proof -
+  { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
+      by (induct i arbitrary: M s) (auto intro!: *) }
+  then have "(\<lambda>x. INF i. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
+    by measurable
+  also have "(\<lambda>x. INF i. (F ^^ i) top s x) = gfp F s"
+    by (subst inf_continuous_gfp) (auto simp: F)
+  finally show ?thesis .
+qed
+
+lemma measurable_enat_coinduct:
+  fixes f :: "'a \<Rightarrow> enat"
+  assumes "R f"
+  assumes *: "\<And>f. R f \<Longrightarrow> \<exists>g h i P. R g \<and> f = (\<lambda>x. if P x then h x else eSuc (g (i x))) \<and>
+    Measurable.pred M P \<and>
+    i \<in> measurable M M \<and>
+    h \<in> measurable M (count_space UNIV)"
+  shows "f \<in> measurable M (count_space UNIV)"
+proof (simp add: measurable_count_space_eq2_countable, rule )
+  fix a :: enat
+  have "f -` {a} \<inter> space M = {x\<in>space M. f x = a}"
+    by auto
+  { fix i :: nat
+    from \<open>R f\<close> have "Measurable.pred M (\<lambda>x. f x = enat i)"
+    proof (induction i arbitrary: f)
+      case 0
+      from *[OF this] obtain g h i P
+        where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and
+          [measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)"
+        by auto
+      have "Measurable.pred M (\<lambda>x. P x \<and> h x = 0)"
+        by measurable
+      also have "(\<lambda>x. P x \<and> h x = 0) = (\<lambda>x. f x = enat 0)"
+        by (auto simp: f zero_enat_def[symmetric])
+      finally show ?case .
+    next
+      case (Suc n)
+      from *[OF Suc.prems] obtain g h i P
+        where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and "R g" and
+          M[measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)"
+        by auto
+      have "(\<lambda>x. f x = enat (Suc n)) =
+        (\<lambda>x. (P x \<longrightarrow> h x = enat (Suc n)) \<and> (\<not> P x \<longrightarrow> g (i x) = enat n))"
+        by (auto simp: f zero_enat_def[symmetric] eSuc_enat[symmetric])
+      also have "Measurable.pred M \<dots>"
+        by (intro pred_intros_logic measurable_compose[OF M(2)] Suc \<open>R g\<close>) measurable
+      finally show ?case .
+    qed
+    then have "f -` {enat i} \<inter> space M \<in> sets M"
+      by (simp add: pred_def Int_def conj_commute) }
+  note fin = this
+  show "f -` {a} \<inter> space M \<in> sets M"
+  proof (cases a)
+    case infinity
+    then have "f -` {a} \<inter> space M = space M - (\<Union>n. f -` {enat n} \<inter> space M)"
+      by auto
+    also have "\<dots> \<in> sets M"
+      by (intro sets.Diff sets.top sets.Un sets.countable_UN) (auto intro!: fin)
+    finally show ?thesis .
+  qed (simp add: fin)
+qed
+
+lemma measurable_THE:
+  fixes P :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
+  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
+  assumes I[simp]: "countable I" "\<And>i x. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> i \<in> I"
+  assumes unique: "\<And>x i j. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> P j x \<Longrightarrow> i = j"
+  shows "(\<lambda>x. THE i. P i x) \<in> measurable M (count_space UNIV)"
+  unfolding measurable_def
+proof safe
+  fix X
+  define f where "f x = (THE i. P i x)" for x
+  define undef where "undef = (THE i::'a. False)"
+  { fix i x assume "x \<in> space M" "P i x" then have "f x = i"
+      unfolding f_def using unique by auto }
+  note f_eq = this
+  { fix x assume "x \<in> space M" "\<forall>i\<in>I. \<not> P i x"
+    then have "\<And>i. \<not> P i x"
+      using I(2)[of x] by auto
+    then have "f x = undef"
+      by (auto simp: undef_def f_def) }
+  then have "f -` X \<inter> space M = (\<Union>i\<in>I \<inter> X. {x\<in>space M. P i x}) \<union>
+     (if undef \<in> X then space M - (\<Union>i\<in>I. {x\<in>space M. P i x}) else {})"
+    by (auto dest: f_eq)
+  also have "\<dots> \<in> sets M"
+    by (auto intro!: sets.Diff sets.countable_UN')
+  finally show "f -` X \<inter> space M \<in> sets M" .
+qed simp
+
+lemma measurable_Ex1[measurable (raw)]:
+  assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> Measurable.pred M (P i)"
+  shows "Measurable.pred M (\<lambda>x. \<exists>!i\<in>I. P i x)"
+  unfolding bex1_def by measurable
+
+lemma measurable_Sup_nat[measurable (raw)]:
+  fixes F :: "'a \<Rightarrow> nat set"
+  assumes [measurable]: "\<And>i. Measurable.pred M (\<lambda>x. i \<in> F x)"
+  shows "(\<lambda>x. Sup (F x)) \<in> M \<rightarrow>\<^sub>M count_space UNIV"
+proof (clarsimp simp add: measurable_count_space_eq2_countable)
+  fix a
+  have F_empty_iff: "F x = {} \<longleftrightarrow> (\<forall>i. i \<notin> F x)" for x
+    by auto
+  have "Measurable.pred M (\<lambda>x. if finite (F x) then if F x = {} then a = Max {}
+    else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None)"
+    unfolding finite_nat_set_iff_bounded Ball_def F_empty_iff by measurable
+  moreover have "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M =
+    {x\<in>space M. if finite (F x) then if F x = {} then a = Max {}
+      else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None}"
+    by (intro set_eqI)
+       (auto simp: Sup_nat_def Max.infinite intro!: Max_in Max_eqI)
+  ultimately show "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M \<in> sets M"
+    by auto
+qed
+
+lemma measurable_if_split[measurable (raw)]:
+  "(c \<Longrightarrow> Measurable.pred M f) \<Longrightarrow> (\<not> c \<Longrightarrow> Measurable.pred M g) \<Longrightarrow>
+   Measurable.pred M (if c then f else g)"
+  by simp
+
+lemma pred_restrict_space:
+  assumes "S \<in> sets M"
+  shows "Measurable.pred (restrict_space M S) P \<longleftrightarrow> Measurable.pred M (\<lambda>x. x \<in> S \<and> P x)"
+  unfolding pred_def sets_Collect_restrict_space_iff[OF assms] ..
+
+lemma measurable_predpow[measurable]:
+  assumes "Measurable.pred M T"
+  assumes "\<And>Q. Measurable.pred M Q \<Longrightarrow> Measurable.pred M (R Q)"
+  shows "Measurable.pred M ((R ^^ n) T)"
+  by (induct n) (auto intro: assms)
+
+hide_const (open) pred
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Measure_Space.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,3326 @@
+(*  Title:      HOL/Probability/Measure_Space.thy
+    Author:     Lawrence C Paulson
+    Author:     Johannes Hölzl, TU München
+    Author:     Armin Heller, TU München
+*)
+
+section \<open>Measure spaces and their properties\<close>
+
+theory Measure_Space
+imports
+  Measurable "~~/src/HOL/Library/Extended_Nonnegative_Real"
+begin
+
+subsection "Relate extended reals and the indicator function"
+
+lemma suminf_cmult_indicator:
+  fixes f :: "nat \<Rightarrow> ennreal"
+  assumes "disjoint_family A" "x \<in> A i"
+  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
+proof -
+  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)"
+    using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto
+  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)"
+    by (auto simp: setsum.If_cases)
+  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)"
+  proof (rule SUP_eqI)
+    fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+    from this[of "Suc i"] show "f i \<le> y" by auto
+  qed (insert assms, simp)
+  ultimately show ?thesis using assms
+    by (subst suminf_eq_SUP) (auto simp: indicator_def)
+qed
+
+lemma suminf_indicator:
+  assumes "disjoint_family A"
+  shows "(\<Sum>n. indicator (A n) x :: ennreal) = indicator (\<Union>i. A i) x"
+proof cases
+  assume *: "x \<in> (\<Union>i. A i)"
+  then obtain i where "x \<in> A i" by auto
+  from suminf_cmult_indicator[OF assms(1), OF \<open>x \<in> A i\<close>, of "\<lambda>k. 1"]
+  show ?thesis using * by simp
+qed simp
+
+lemma setsum_indicator_disjoint_family:
+  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
+  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
+  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
+proof -
+  have "P \<inter> {i. x \<in> A i} = {j}"
+    using d \<open>x \<in> A j\<close> \<open>j \<in> P\<close> unfolding disjoint_family_on_def
+    by auto
+  thus ?thesis
+    unfolding indicator_def
+    by (simp add: if_distrib setsum.If_cases[OF \<open>finite P\<close>])
+qed
+
+text \<open>
+  The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
+  represent sigma algebras (with an arbitrary emeasure).
+\<close>
+
+subsection "Extend binary sets"
+
+lemma LIMSEQ_binaryset:
+  assumes f: "f {} = 0"
+  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
+proof -
+  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
+    proof
+      fix n
+      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
+        by (induct n)  (auto simp add: binaryset_def f)
+    qed
+  moreover
+  have "... \<longlonglongrightarrow> f A + f B" by (rule tendsto_const)
+  ultimately
+  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
+    by metis
+  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
+    by simp
+  thus ?thesis by (rule LIMSEQ_offset [where k=2])
+qed
+
+lemma binaryset_sums:
+  assumes f: "f {} = 0"
+  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
+    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
+
+lemma suminf_binaryset_eq:
+  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
+  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
+  by (metis binaryset_sums sums_unique)
+
+subsection \<open>Properties of a premeasure @{term \<mu>}\<close>
+
+text \<open>
+  The definitions for @{const positive} and @{const countably_additive} should be here, by they are
+  necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
+\<close>
+
+definition subadditive where
+  "subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
+
+lemma subadditiveD: "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
+  by (auto simp add: subadditive_def)
+
+definition countably_subadditive where
+  "countably_subadditive M f \<longleftrightarrow>
+    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
+
+lemma (in ring_of_sets) countably_subadditive_subadditive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" and cs: "countably_subadditive M f"
+  shows  "subadditive M f"
+proof (auto simp add: subadditive_def)
+  fix x y
+  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
+  hence "disjoint_family (binaryset x y)"
+    by (auto simp add: disjoint_family_on_def binaryset_def)
+  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
+         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
+         f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
+    using cs by (auto simp add: countably_subadditive_def)
+  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
+         f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
+    by (simp add: range_binaryset_eq UN_binaryset_eq)
+  thus "f (x \<union> y) \<le>  f x + f y" using f x y
+    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
+qed
+
+definition additive where
+  "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
+
+definition increasing where
+  "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
+
+lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
+
+lemma positiveD_empty:
+  "positive M f \<Longrightarrow> f {} = 0"
+  by (auto simp add: positive_def)
+
+lemma additiveD:
+  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
+  by (auto simp add: additive_def)
+
+lemma increasingD:
+  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
+  by (auto simp add: increasing_def)
+
+lemma countably_additiveI[case_names countably]:
+  "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
+  \<Longrightarrow> countably_additive M f"
+  by (simp add: countably_additive_def)
+
+lemma (in ring_of_sets) disjointed_additive:
+  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
+  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
+proof (induct n)
+  case (Suc n)
+  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
+    by simp
+  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
+    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono)
+  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
+    using \<open>incseq A\<close> by (auto dest: incseq_SucD simp: disjointed_mono)
+  finally show ?case .
+qed simp
+
+lemma (in ring_of_sets) additive_sum:
+  fixes A:: "'i \<Rightarrow> 'a set"
+  assumes f: "positive M f" and ad: "additive M f" and "finite S"
+      and A: "A`S \<subseteq> M"
+      and disj: "disjoint_family_on A S"
+  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
+  using \<open>finite S\<close> disj A
+proof induct
+  case empty show ?case using f by (simp add: positive_def)
+next
+  case (insert s S)
+  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
+    by (auto simp add: disjoint_family_on_def neq_iff)
+  moreover
+  have "A s \<in> M" using insert by blast
+  moreover have "(\<Union>i\<in>S. A i) \<in> M"
+    using insert \<open>finite S\<close> by auto
+  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
+    using ad UNION_in_sets A by (auto simp add: additive_def)
+  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
+    by (auto simp add: additive_def subset_insertI)
+qed
+
+lemma (in ring_of_sets) additive_increasing:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes posf: "positive M f" and addf: "additive M f"
+  shows "increasing M f"
+proof (auto simp add: increasing_def)
+  fix x y
+  assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
+  then have "y - x \<in> M" by auto
+  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono zero_le)
+  also have "... = f (x \<union> (y-x))" using addf
+    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
+  also have "... = f y"
+    by (metis Un_Diff_cancel Un_absorb1 xy(3))
+  finally show "f x \<le> f y" by simp
+qed
+
+lemma (in ring_of_sets) subadditive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" "additive M f" and A: "A`S \<subseteq> M" and S: "finite S"
+  shows "f (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. f (A i))"
+using S A
+proof (induct S)
+  case empty thus ?case using f by (auto simp: positive_def)
+next
+  case (insert x F)
+  hence in_M: "A x \<in> M" "(\<Union>i\<in>F. A i) \<in> M" "(\<Union>i\<in>F. A i) - A x \<in> M" using A by force+
+  have subs: "(\<Union>i\<in>F. A i) - A x \<subseteq> (\<Union>i\<in>F. A i)" by auto
+  have "(\<Union>i\<in>(insert x F). A i) = A x \<union> ((\<Union>i\<in>F. A i) - A x)" by auto
+  hence "f (\<Union>i\<in>(insert x F). A i) = f (A x \<union> ((\<Union>i\<in>F. A i) - A x))"
+    by simp
+  also have "\<dots> = f (A x) + f ((\<Union>i\<in>F. A i) - A x)"
+    using f(2) by (rule additiveD) (insert in_M, auto)
+  also have "\<dots> \<le> f (A x) + f (\<Union>i\<in>F. A i)"
+    using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono)
+  also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" using insert by (auto intro: add_left_mono)
+  finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))" using insert by simp
+qed
+
+lemma (in ring_of_sets) countably_additive_additive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes posf: "positive M f" and ca: "countably_additive M f"
+  shows "additive M f"
+proof (auto simp add: additive_def)
+  fix x y
+  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
+  hence "disjoint_family (binaryset x y)"
+    by (auto simp add: disjoint_family_on_def binaryset_def)
+  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
+         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
+         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
+    using ca
+    by (simp add: countably_additive_def)
+  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
+         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
+    by (simp add: range_binaryset_eq UN_binaryset_eq)
+  thus "f (x \<union> y) = f x + f y" using posf x y
+    by (auto simp add: Un suminf_binaryset_eq positive_def)
+qed
+
+lemma (in algebra) increasing_additive_bound:
+  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" and ad: "additive M f"
+      and inc: "increasing M f"
+      and A: "range A \<subseteq> M"
+      and disj: "disjoint_family A"
+  shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
+proof (safe intro!: suminf_le_const)
+  fix N
+  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
+  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
+    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
+  also have "... \<le> f \<Omega>" using space_closed A
+    by (intro increasingD[OF inc] finite_UN) auto
+  finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
+qed (insert f A, auto simp: positive_def)
+
+lemma (in ring_of_sets) countably_additiveI_finite:
+  fixes \<mu> :: "'a set \<Rightarrow> ennreal"
+  assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
+  shows "countably_additive M \<mu>"
+proof (rule countably_additiveI)
+  fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
+
+  have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
+  from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
+
+  have inj_f: "inj_on f {i. F i \<noteq> {}}"
+  proof (rule inj_onI, simp)
+    fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
+    then have "f i \<in> F i" "f j \<in> F j" using f by force+
+    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
+  qed
+  have "finite (\<Union>i. F i)"
+    by (metis F(2) assms(1) infinite_super sets_into_space)
+
+  have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
+    by (auto simp: positiveD_empty[OF \<open>positive M \<mu>\<close>])
+  moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
+  proof (rule finite_imageD)
+    from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
+    then show "finite (f`{i. F i \<noteq> {}})"
+      by (rule finite_subset) fact
+  qed fact
+  ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
+    by (rule finite_subset)
+
+  have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
+    using disj by (auto simp: disjoint_family_on_def)
+
+  from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
+    by (rule suminf_finite) auto
+  also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
+    using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
+  also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
+    using \<open>positive M \<mu>\<close> \<open>additive M \<mu>\<close> fin_not_empty disj_not_empty F by (intro additive_sum) auto
+  also have "\<dots> = \<mu> (\<Union>i. F i)"
+    by (rule arg_cong[where f=\<mu>]) auto
+  finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
+qed
+
+lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" "additive M f"
+  shows "countably_additive M f \<longleftrightarrow>
+    (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i))"
+  unfolding countably_additive_def
+proof safe
+  assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
+  then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets)
+  with count_sum[THEN spec, of "disjointed A"] A(3)
+  have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
+    by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
+  moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
+    using f(1)[unfolded positive_def] dA
+    by (auto intro!: summable_LIMSEQ)
+  from LIMSEQ_Suc[OF this]
+  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
+    unfolding lessThan_Suc_atMost .
+  moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
+    using disjointed_additive[OF f A(1,2)] .
+  ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" by simp
+next
+  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M"
+  have *: "(\<Union>n. (\<Union>i<n. A i)) = (\<Union>i. A i)" by auto
+  have "(\<lambda>n. f (\<Union>i<n. A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+  proof (unfold *[symmetric], intro cont[rule_format])
+    show "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M"
+      using A * by auto
+  qed (force intro!: incseq_SucI)
+  moreover have "\<And>n. f (\<Union>i<n. A i) = (\<Sum>i<n. f (A i))"
+    using A
+    by (intro additive_sum[OF f, of _ A, symmetric])
+       (auto intro: disjoint_family_on_mono[where B=UNIV])
+  ultimately
+  have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
+    unfolding sums_def by simp
+  from sums_unique[OF this]
+  show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
+qed
+
+lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" "additive M f"
+  shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))
+     \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0)"
+proof safe
+  assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
+  with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
+    using \<open>positive M f\<close>[unfolded positive_def] by auto
+next
+  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
+
+  have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
+    using additive_increasing[OF f] unfolding increasing_def by simp
+
+  have decseq_fA: "decseq (\<lambda>i. f (A i))"
+    using A by (auto simp: decseq_def intro!: f_mono)
+  have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))"
+    using A by (auto simp: decseq_def)
+  then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))"
+    using A unfolding decseq_def by (auto intro!: f_mono Diff)
+  have "f (\<Inter>x. A x) \<le> f (A 0)"
+    using A by (auto intro!: f_mono)
+  then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
+    using A by (auto simp: top_unique)
+  { fix i
+    have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
+    then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
+      using A by (auto simp: top_unique) }
+  note f_fin = this
+  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) \<longlonglongrightarrow> 0"
+  proof (intro cont[rule_format, OF _ decseq _ f_fin])
+    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
+      using A by auto
+  qed
+  from INF_Lim_ereal[OF decseq_f this]
+  have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
+  moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
+    by auto
+  ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
+    using A(4) f_fin f_Int_fin
+    by (subst INF_ennreal_add_const) (auto simp: decseq_f)
+  moreover {
+    fix n
+    have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
+      using A by (subst f(2)[THEN additiveD]) auto
+    also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n"
+      by auto
+    finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
+  ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
+    by simp
+  with LIMSEQ_INF[OF decseq_fA]
+  show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i)" by simp
+qed
+
+lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>"
+  assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
+  assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
+  shows "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+proof -
+  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) \<longlonglongrightarrow> 0"
+    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
+  moreover
+  { fix i
+    have "f ((\<Union>i. A i) - A i \<union> A i) = f ((\<Union>i. A i) - A i) + f (A i)"
+      using A by (intro f(2)[THEN additiveD]) auto
+    also have "((\<Union>i. A i) - A i) \<union> A i = (\<Union>i. A i)"
+      by auto
+    finally have "f ((\<Union>i. A i) - A i) = f (\<Union>i. A i) - f (A i)"
+      using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) }
+  moreover have "\<forall>\<^sub>F i in sequentially. f (A i) \<le> f (\<Union>i. A i)"
+    using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\<Union>i. A i"] A
+    by (auto intro!: always_eventually simp: subset_eq)
+  ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
+    by (auto intro: ennreal_tendsto_const_minus)
+qed
+
+lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
+  fixes f :: "'a set \<Rightarrow> ennreal"
+  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. 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 "countably_additive M f"
+  using countably_additive_iff_continuous_from_below[OF f]
+  using empty_continuous_imp_continuous_from_below[OF f fin] cont
+  by blast
+
+subsection \<open>Properties of @{const emeasure}\<close>
+
+lemma emeasure_positive: "positive (sets M) (emeasure M)"
+  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
+
+lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
+  using emeasure_positive[of M] by (simp add: positive_def)
+
+lemma emeasure_single_in_space: "emeasure M {x} \<noteq> 0 \<Longrightarrow> x \<in> space M"
+  using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2])
+
+lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
+  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
+
+lemma suminf_emeasure:
+  "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
+  using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M]
+  by (simp add: countably_additive_def)
+
+lemma sums_emeasure:
+  "disjoint_family F \<Longrightarrow> (\<And>i. F i \<in> sets M) \<Longrightarrow> (\<lambda>i. emeasure M (F i)) sums emeasure M (\<Union>i. F i)"
+  unfolding sums_iff by (intro conjI suminf_emeasure) auto
+
+lemma emeasure_additive: "additive (sets M) (emeasure M)"
+  by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive)
+
+lemma plus_emeasure:
+  "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
+  using additiveD[OF emeasure_additive] ..
+
+lemma setsum_emeasure:
+  "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
+    (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
+  by (metis sets.additive_sum emeasure_positive emeasure_additive)
+
+lemma emeasure_mono:
+  "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
+  by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD)
+
+lemma emeasure_space:
+  "emeasure M A \<le> emeasure M (space M)"
+  by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le)
+
+lemma emeasure_Diff:
+  assumes finite: "emeasure M B \<noteq> \<infinity>"
+  and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
+  shows "emeasure M (A - B) = emeasure M A - emeasure M B"
+proof -
+  have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto
+  then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
+  also have "\<dots> = emeasure M (A - B) + emeasure M B"
+    by (subst plus_emeasure[symmetric]) auto
+  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
+    using finite by simp
+qed
+
+lemma emeasure_compl:
+  "s \<in> sets M \<Longrightarrow> emeasure M s \<noteq> \<infinity> \<Longrightarrow> emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
+  by (rule emeasure_Diff) (auto dest: sets.sets_into_space)
+
+lemma Lim_emeasure_incseq:
+  "range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) \<longlonglongrightarrow> emeasure M (\<Union>i. A i)"
+  using emeasure_countably_additive
+  by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive
+    emeasure_additive)
+
+lemma incseq_emeasure:
+  assumes "range B \<subseteq> sets M" "incseq B"
+  shows "incseq (\<lambda>i. emeasure M (B i))"
+  using assms by (auto simp: incseq_def intro!: emeasure_mono)
+
+lemma SUP_emeasure_incseq:
+  assumes A: "range A \<subseteq> sets M" "incseq A"
+  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
+  using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
+  by (simp add: LIMSEQ_unique)
+
+lemma decseq_emeasure:
+  assumes "range B \<subseteq> sets M" "decseq B"
+  shows "decseq (\<lambda>i. emeasure M (B i))"
+  using assms by (auto simp: decseq_def intro!: emeasure_mono)
+
+lemma INF_emeasure_decseq:
+  assumes A: "range A \<subseteq> sets M" and "decseq A"
+  and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
+proof -
+  have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
+    using A by (auto intro!: emeasure_mono)
+  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by (auto simp: top_unique)
+
+  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))"
+    by (simp add: ennreal_INF_const_minus)
+  also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
+    using A finite \<open>decseq A\<close>[unfolded decseq_def] by (subst emeasure_Diff) auto
+  also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
+  proof (rule SUP_emeasure_incseq)
+    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
+      using A by auto
+    show "incseq (\<lambda>n. A 0 - A n)"
+      using \<open>decseq A\<close> by (auto simp add: incseq_def decseq_def)
+  qed
+  also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
+    using A finite * by (simp, subst emeasure_Diff) auto
+  finally show ?thesis
+    by (rule ennreal_minus_cancel[rotated 3])
+       (insert finite A, auto intro: INF_lower emeasure_mono)
+qed
+
+lemma emeasure_INT_decseq_subset:
+  fixes F :: "nat \<Rightarrow> 'a set"
+  assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
+  assumes F_sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M"
+    and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (F i) \<noteq> \<infinity>"
+  shows "emeasure M (\<Inter>i\<in>I. F i) = (INF i:I. emeasure M (F i))"
+proof cases
+  assume "finite I"
+  have "(\<Inter>i\<in>I. F i) = F (Max I)"
+    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F) auto
+  moreover have "(INF i:I. emeasure M (F i)) = emeasure M (F (Max I))"
+    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F emeasure_mono) auto
+  ultimately show ?thesis
+    by simp
+next
+  assume "infinite I"
+  define L where "L n = (LEAST i. i \<in> I \<and> i \<ge> n)" for n
+  have L: "L n \<in> I \<and> n \<le> L n" for n
+    unfolding L_def
+  proof (rule LeastI_ex)
+    show "\<exists>x. x \<in> I \<and> n \<le> x"
+      using \<open>infinite I\<close> finite_subset[of I "{..< n}"]
+      by (rule_tac ccontr) (auto simp: not_le)
+  qed
+  have L_eq[simp]: "i \<in> I \<Longrightarrow> L i = i" for i
+    unfolding L_def by (intro Least_equality) auto
+  have L_mono: "i \<le> j \<Longrightarrow> L i \<le> L j" for i j
+    using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def)
+
+  have "emeasure M (\<Inter>i. F (L i)) = (INF i. emeasure M (F (L i)))"
+  proof (intro INF_emeasure_decseq[symmetric])
+    show "decseq (\<lambda>i. F (L i))"
+      using L by (intro antimonoI F L_mono) auto
+  qed (insert L fin, auto)
+  also have "\<dots> = (INF i:I. emeasure M (F i))"
+  proof (intro antisym INF_greatest)
+    show "i \<in> I \<Longrightarrow> (INF i. emeasure M (F (L i))) \<le> emeasure M (F i)" for i
+      by (intro INF_lower2[of i]) auto
+  qed (insert L, auto intro: INF_lower)
+  also have "(\<Inter>i. F (L i)) = (\<Inter>i\<in>I. F i)"
+  proof (intro antisym INF_greatest)
+    show "i \<in> I \<Longrightarrow> (\<Inter>i. F (L i)) \<subseteq> F i" for i
+      by (intro INF_lower2[of i]) auto
+  qed (insert L, auto)
+  finally show ?thesis .
+qed
+
+lemma Lim_emeasure_decseq:
+  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. emeasure M (A i)) \<longlonglongrightarrow> emeasure M (\<Inter>i. A i)"
+  using LIMSEQ_INF[OF decseq_emeasure, OF A]
+  using INF_emeasure_decseq[OF A fin] by simp
+
+lemma emeasure_lfp'[consumes 1, case_names cont measurable]:
+  assumes "P M"
+  assumes cont: "sup_continuous F"
+  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
+  shows "emeasure M {x\<in>space M. lfp F x} = (SUP i. emeasure M {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
+proof -
+  have "emeasure M {x\<in>space M. lfp F x} = emeasure M (\<Union>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
+    using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure])
+  moreover { fix i from \<open>P M\<close> have "{x\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
+    by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) }
+  moreover have "incseq (\<lambda>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
+  proof (rule incseq_SucI)
+    fix i
+    have "(F ^^ i) (\<lambda>x. False) \<le> (F ^^ (Suc i)) (\<lambda>x. False)"
+    proof (induct i)
+      case 0 show ?case by (simp add: le_fun_def)
+    next
+      case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto
+    qed
+    then show "{x \<in> space M. (F ^^ i) (\<lambda>x. False) x} \<subseteq> {x \<in> space M. (F ^^ Suc i) (\<lambda>x. False) x}"
+      by auto
+  qed
+  ultimately show ?thesis
+    by (subst SUP_emeasure_incseq) auto
+qed
+
+lemma emeasure_lfp:
+  assumes [simp]: "\<And>s. sets (M s) = sets N"
+  assumes cont: "sup_continuous F" "sup_continuous f"
+  assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
+  assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
+  shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
+proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
+  fix C assume "incseq C" "\<And>i. Measurable.pred N (C i)"
+  then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
+    unfolding SUP_apply[abs_def]
+    by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
+qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont)
+
+lemma emeasure_subadditive_finite:
+  "finite I \<Longrightarrow> A ` I \<subseteq> sets M \<Longrightarrow> emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
+  by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto
+
+lemma emeasure_subadditive:
+  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
+  using emeasure_subadditive_finite[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B" M] by simp
+
+lemma emeasure_subadditive_countably:
+  assumes "range f \<subseteq> sets M"
+  shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
+proof -
+  have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
+    unfolding UN_disjointed_eq ..
+  also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
+    using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
+    by (simp add:  disjoint_family_disjointed comp_def)
+  also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
+    using sets.range_disjointed_sets[OF assms] assms
+    by (auto intro!: suminf_le emeasure_mono disjointed_subset)
+  finally show ?thesis .
+qed
+
+lemma emeasure_insert:
+  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
+  shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
+proof -
+  have "{x} \<inter> A = {}" using \<open>x \<notin> A\<close> by auto
+  from plus_emeasure[OF sets this] show ?thesis by simp
+qed
+
+lemma emeasure_insert_ne:
+  "A \<noteq> {} \<Longrightarrow> {x} \<in> sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> x \<notin> A \<Longrightarrow> emeasure M (insert x A) = emeasure M {x} + emeasure M A"
+  by (rule emeasure_insert)
+
+lemma emeasure_eq_setsum_singleton:
+  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
+  using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
+  by (auto simp: disjoint_family_on_def subset_eq)
+
+lemma setsum_emeasure_cover:
+  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
+  assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
+  assumes disj: "disjoint_family_on B S"
+  shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
+proof -
+  have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
+  proof (rule setsum_emeasure)
+    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
+      using \<open>disjoint_family_on B S\<close>
+      unfolding disjoint_family_on_def by auto
+  qed (insert assms, auto)
+  also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
+    using A by auto
+  finally show ?thesis by simp
+qed
+
+lemma emeasure_eq_0:
+  "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
+  by (metis emeasure_mono order_eq_iff zero_le)
+
+lemma emeasure_UN_eq_0:
+  assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
+  shows "emeasure M (\<Union>i. N i) = 0"
+proof -
+  have "emeasure M (\<Union>i. N i) \<le> 0"
+    using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
+  then show ?thesis
+    by (auto intro: antisym zero_le)
+qed
+
+lemma measure_eqI_finite:
+  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
+  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
+  shows "M = N"
+proof (rule measure_eqI)
+  fix X assume "X \<in> sets M"
+  then have X: "X \<subseteq> A" by auto
+  then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
+    using \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+  also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
+    using X eq by (auto intro!: setsum.cong)
+  also have "\<dots> = emeasure N X"
+    using X \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+  finally show "emeasure M X = emeasure N X" .
+qed simp
+
+lemma measure_eqI_generator_eq:
+  fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
+  assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
+  and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
+  and M: "sets M = sigma_sets \<Omega> E"
+  and N: "sets N = sigma_sets \<Omega> E"
+  and A: "range A \<subseteq> E" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "M = N"
+proof -
+  let ?\<mu>  = "emeasure M" and ?\<nu> = "emeasure N"
+  interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
+  have "space M = \<Omega>"
+    using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \<open>sets M = sigma_sets \<Omega> E\<close>
+    by blast
+
+  { fix F D assume "F \<in> E" and "?\<mu> F \<noteq> \<infinity>"
+    then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
+    have "?\<nu> F \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> \<open>F \<in> E\<close> eq by simp
+    assume "D \<in> sets M"
+    with \<open>Int_stable E\<close> \<open>E \<subseteq> Pow \<Omega>\<close> have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
+      unfolding M
+    proof (induct rule: sigma_sets_induct_disjoint)
+      case (basic A)
+      then have "F \<inter> A \<in> E" using \<open>Int_stable E\<close> \<open>F \<in> E\<close> by (auto simp: Int_stable_def)
+      then show ?case using eq by auto
+    next
+      case empty then show ?case by simp
+    next
+      case (compl A)
+      then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
+        and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
+        using \<open>F \<in> E\<close> S.sets_into_space by (auto simp: M)
+      have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N)
+      then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
+      have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N)
+      then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
+      then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
+        using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> by (auto intro!: emeasure_Diff simp: M N)
+      also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq \<open>F \<in> E\<close> compl by simp
+      also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding **
+        using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> \<open>?\<nu> (F \<inter> A) \<noteq> \<infinity>\<close>
+        by (auto intro!: emeasure_Diff[symmetric] simp: M N)
+      finally show ?case
+        using \<open>space M = \<Omega>\<close> by auto
+    next
+      case (union A)
+      then have "?\<mu> (\<Union>x. F \<inter> A x) = ?\<nu> (\<Union>x. F \<inter> A x)"
+        by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N)
+      with A show ?case
+        by auto
+    qed }
+  note * = this
+  show "M = N"
+  proof (rule measure_eqI)
+    show "sets M = sets N"
+      using M N by simp
+    have [simp, intro]: "\<And>i. A i \<in> sets M"
+      using A(1) by (auto simp: subset_eq M)
+    fix F assume "F \<in> sets M"
+    let ?D = "disjointed (\<lambda>i. F \<inter> A i)"
+    from \<open>space M = \<Omega>\<close> have F_eq: "F = (\<Union>i. ?D i)"
+      using \<open>F \<in> sets M\<close>[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
+    have [simp, intro]: "\<And>i. ?D i \<in> sets M"
+      using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] \<open>F \<in> sets M\<close>
+      by (auto simp: subset_eq)
+    have "disjoint_family ?D"
+      by (auto simp: disjoint_family_disjointed)
+    moreover
+    have "(\<Sum>i. emeasure M (?D i)) = (\<Sum>i. emeasure N (?D i))"
+    proof (intro arg_cong[where f=suminf] ext)
+      fix i
+      have "A i \<inter> ?D i = ?D i"
+        by (auto simp: disjointed_def)
+      then show "emeasure M (?D i) = emeasure N (?D i)"
+        using *[of "A i" "?D i", OF _ A(3)] A(1) by auto
+    qed
+    ultimately show "emeasure M F = emeasure N F"
+      by (simp add: image_subset_iff \<open>sets M = sets N\<close>[symmetric] F_eq[symmetric] suminf_emeasure)
+  qed
+qed
+
+lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
+proof (intro measure_eqI emeasure_measure_of_sigma)
+  show "sigma_algebra (space M) (sets M)" ..
+  show "positive (sets M) (emeasure M)"
+    by (simp add: positive_def)
+  show "countably_additive (sets M) (emeasure M)"
+    by (simp add: emeasure_countably_additive)
+qed simp_all
+
+subsection \<open>\<open>\<mu>\<close>-null sets\<close>
+
+definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
+  "null_sets M = {N\<in>sets M. emeasure M N = 0}"
+
+lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
+  by (simp add: null_sets_def)
+
+lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
+  unfolding null_sets_def by simp
+
+lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
+  unfolding null_sets_def by simp
+
+interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
+proof (rule ring_of_setsI)
+  show "null_sets M \<subseteq> Pow (space M)"
+    using sets.sets_into_space by auto
+  show "{} \<in> null_sets M"
+    by auto
+  fix A B assume null_sets: "A \<in> null_sets M" "B \<in> null_sets M"
+  then have sets: "A \<in> sets M" "B \<in> sets M"
+    by auto
+  then have *: "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
+    "emeasure M (A - B) \<le> emeasure M A"
+    by (auto intro!: emeasure_subadditive emeasure_mono)
+  then have "emeasure M B = 0" "emeasure M A = 0"
+    using null_sets by auto
+  with sets * show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
+    by (auto intro!: antisym zero_le)
+qed
+
+lemma UN_from_nat_into:
+  assumes I: "countable I" "I \<noteq> {}"
+  shows "(\<Union>i\<in>I. N i) = (\<Union>i. N (from_nat_into I i))"
+proof -
+  have "(\<Union>i\<in>I. N i) = \<Union>(N ` range (from_nat_into I))"
+    using I by simp
+  also have "\<dots> = (\<Union>i. (N \<circ> from_nat_into I) i)"
+    by simp
+  finally show ?thesis by simp
+qed
+
+lemma null_sets_UN':
+  assumes "countable I"
+  assumes "\<And>i. i \<in> I \<Longrightarrow> N i \<in> null_sets M"
+  shows "(\<Union>i\<in>I. N i) \<in> null_sets M"
+proof cases
+  assume "I = {}" then show ?thesis by simp
+next
+  assume "I \<noteq> {}"
+  show ?thesis
+  proof (intro conjI CollectI null_setsI)
+    show "(\<Union>i\<in>I. N i) \<in> sets M"
+      using assms by (intro sets.countable_UN') auto
+    have "emeasure M (\<Union>i\<in>I. N i) \<le> (\<Sum>n. emeasure M (N (from_nat_into I n)))"
+      unfolding UN_from_nat_into[OF \<open>countable I\<close> \<open>I \<noteq> {}\<close>]
+      using assms \<open>I \<noteq> {}\<close> by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
+    also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)"
+      using assms \<open>I \<noteq> {}\<close> by (auto intro: from_nat_into)
+    finally show "emeasure M (\<Union>i\<in>I. N i) = 0"
+      by (intro antisym zero_le) simp
+  qed
+qed
+
+lemma null_sets_UN[intro]:
+  "(\<And>i::'i::countable. N i \<in> null_sets M) \<Longrightarrow> (\<Union>i. N i) \<in> null_sets M"
+  by (rule null_sets_UN') auto
+
+lemma null_set_Int1:
+  assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
+proof (intro CollectI conjI null_setsI)
+  show "emeasure M (A \<inter> B) = 0" using assms
+    by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
+qed (insert assms, auto)
+
+lemma null_set_Int2:
+  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
+  using assms by (subst Int_commute) (rule null_set_Int1)
+
+lemma emeasure_Diff_null_set:
+  assumes "B \<in> null_sets M" "A \<in> sets M"
+  shows "emeasure M (A - B) = emeasure M A"
+proof -
+  have *: "A - B = (A - (A \<inter> B))" by auto
+  have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
+  then show ?thesis
+    unfolding * using assms
+    by (subst emeasure_Diff) auto
+qed
+
+lemma null_set_Diff:
+  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
+proof (intro CollectI conjI null_setsI)
+  show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
+qed (insert assms, auto)
+
+lemma emeasure_Un_null_set:
+  assumes "A \<in> sets M" "B \<in> null_sets M"
+  shows "emeasure M (A \<union> B) = emeasure M A"
+proof -
+  have *: "A \<union> B = A \<union> (B - A)" by auto
+  have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
+  then show ?thesis
+    unfolding * using assms
+    by (subst plus_emeasure[symmetric]) auto
+qed
+
+subsection \<open>The almost everywhere filter (i.e.\ quantifier)\<close>
+
+definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
+  "ae_filter M = (INF N:null_sets M. principal (space M - N))"
+
+abbreviation almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
+  "almost_everywhere M P \<equiv> eventually P (ae_filter M)"
+
+syntax
+  "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
+
+translations
+  "AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)"
+
+lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
+  unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)
+
+lemma AE_I':
+  "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
+  unfolding eventually_ae_filter by auto
+
+lemma AE_iff_null:
+  assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
+  shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
+proof
+  assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
+    unfolding eventually_ae_filter by auto
+  have "emeasure M ?P \<le> emeasure M N"
+    using assms N(1,2) by (auto intro: emeasure_mono)
+  then have "emeasure M ?P = 0"
+    unfolding \<open>emeasure M N = 0\<close> by auto
+  then show "?P \<in> null_sets M" using assms by auto
+next
+  assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
+qed
+
+lemma AE_iff_null_sets:
+  "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
+  using Int_absorb1[OF sets.sets_into_space, of N M]
+  by (subst AE_iff_null) (auto simp: Int_def[symmetric])
+
+lemma AE_not_in:
+  "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
+  by (metis AE_iff_null_sets null_setsD2)
+
+lemma AE_iff_measurable:
+  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
+  using AE_iff_null[of _ P] by auto
+
+lemma AE_E[consumes 1]:
+  assumes "AE x in M. P x"
+  obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
+  using assms unfolding eventually_ae_filter by auto
+
+lemma AE_E2:
+  assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
+  shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
+proof -
+  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
+  with AE_iff_null[of M P] assms show ?thesis by auto
+qed
+
+lemma AE_I:
+  assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
+  shows "AE x in M. P x"
+  using assms unfolding eventually_ae_filter by auto
+
+lemma AE_mp[elim!]:
+  assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
+  shows "AE x in M. Q x"
+proof -
+  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
+    and A: "A \<in> sets M" "emeasure M A = 0"
+    by (auto elim!: AE_E)
+
+  from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
+    and B: "B \<in> sets M" "emeasure M B = 0"
+    by (auto elim!: AE_E)
+
+  show ?thesis
+  proof (intro AE_I)
+    have "emeasure M (A \<union> B) \<le> 0"
+      using emeasure_subadditive[of A M B] A B by auto
+    then show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0"
+      using A B by auto
+    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
+      using P imp by auto
+  qed
+qed
+
+(* depricated replace by laws about eventually *)
+lemma
+  shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
+    and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
+    and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
+    and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x"
+    and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)"
+  by auto
+
+lemma AE_impI:
+  "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
+  by (cases P) auto
+
+lemma AE_measure:
+  assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
+  shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
+proof -
+  from AE_E[OF AE] guess N . note N = this
+  with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
+    by (intro emeasure_mono) auto
+  also have "\<dots> \<le> emeasure M ?P + emeasure M N"
+    using sets N by (intro emeasure_subadditive) auto
+  also have "\<dots> = emeasure M ?P" using N by simp
+  finally show "emeasure M ?P = emeasure M (space M)"
+    using emeasure_space[of M "?P"] by auto
+qed
+
+lemma AE_space: "AE x in M. x \<in> space M"
+  by (rule AE_I[where N="{}"]) auto
+
+lemma AE_I2[simp, intro]:
+  "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
+  using AE_space by force
+
+lemma AE_Ball_mp:
+  "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
+  by auto
+
+lemma AE_cong[cong]:
+  "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)"
+  by auto
+
+lemma AE_all_countable:
+  "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
+proof
+  assume "\<forall>i. AE x in M. P i x"
+  from this[unfolded eventually_ae_filter Bex_def, THEN choice]
+  obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
+  have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
+  also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
+  finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
+  moreover from N have "(\<Union>i. N i) \<in> null_sets M"
+    by (intro null_sets_UN) auto
+  ultimately show "AE x in M. \<forall>i. P i x"
+    unfolding eventually_ae_filter by auto
+qed auto
+
+lemma AE_ball_countable:
+  assumes [intro]: "countable X"
+  shows "(AE x in M. \<forall>y\<in>X. P x y) \<longleftrightarrow> (\<forall>y\<in>X. AE x in M. P x y)"
+proof
+  assume "\<forall>y\<in>X. AE x in M. P x y"
+  from this[unfolded eventually_ae_filter Bex_def, THEN bchoice]
+  obtain N where N: "\<And>y. y \<in> X \<Longrightarrow> N y \<in> null_sets M" "\<And>y. y \<in> X \<Longrightarrow> {x\<in>space M. \<not> P x y} \<subseteq> N y"
+    by auto
+  have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. {x\<in>space M. \<not> P x y})"
+    by auto
+  also have "\<dots> \<subseteq> (\<Union>y\<in>X. N y)"
+    using N by auto
+  finally have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. N y)" .
+  moreover from N have "(\<Union>y\<in>X. N y) \<in> null_sets M"
+    by (intro null_sets_UN') auto
+  ultimately show "AE x in M. \<forall>y\<in>X. P x y"
+    unfolding eventually_ae_filter by auto
+qed auto
+
+lemma AE_discrete_difference:
+  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"
+  shows "AE x in M. x \<notin> X"
+proof -
+  have "(\<Union>x\<in>X. {x}) \<in> null_sets M"
+    using assms by (intro null_sets_UN') auto
+  from AE_not_in[OF this] show "AE x in M. x \<notin> X"
+    by auto
+qed
+
+lemma AE_finite_all:
+  assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)"
+  using f by induct auto
+
+lemma AE_finite_allI:
+  assumes "finite S"
+  shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
+  using AE_finite_all[OF \<open>finite S\<close>] by auto
+
+lemma emeasure_mono_AE:
+  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
+    and B: "B \<in> sets M"
+  shows "emeasure M A \<le> emeasure M B"
+proof cases
+  assume A: "A \<in> sets M"
+  from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
+    by (auto simp: eventually_ae_filter)
+  have "emeasure M A = emeasure M (A - N)"
+    using N A by (subst emeasure_Diff_null_set) auto
+  also have "emeasure M (A - N) \<le> emeasure M (B - N)"
+    using N A B sets.sets_into_space by (auto intro!: emeasure_mono)
+  also have "emeasure M (B - N) = emeasure M B"
+    using N B by (subst emeasure_Diff_null_set) auto
+  finally show ?thesis .
+qed (simp add: emeasure_notin_sets)
+
+lemma emeasure_eq_AE:
+  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "emeasure M A = emeasure M B"
+  using assms by (safe intro!: antisym emeasure_mono_AE) auto
+
+lemma emeasure_Collect_eq_AE:
+  "AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> Measurable.pred M Q \<Longrightarrow> Measurable.pred M P \<Longrightarrow>
+   emeasure M {x\<in>space M. P x} = emeasure M {x\<in>space M. Q x}"
+   by (intro emeasure_eq_AE) auto
+
+lemma emeasure_eq_0_AE: "AE x in M. \<not> P x \<Longrightarrow> emeasure M {x\<in>space M. P x} = 0"
+  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"]
+  by (cases "{x\<in>space M. P x} \<in> sets M") (simp_all add: emeasure_notin_sets)
+
+lemma emeasure_add_AE:
+  assumes [measurable]: "A \<in> sets M" "B \<in> sets M" "C \<in> sets M"
+  assumes 1: "AE x in M. x \<in> C \<longleftrightarrow> x \<in> A \<or> x \<in> B"
+  assumes 2: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)"
+  shows "emeasure M C = emeasure M A + emeasure M B"
+proof -
+  have "emeasure M C = emeasure M (A \<union> B)"
+    by (rule emeasure_eq_AE) (insert 1, auto)
+  also have "\<dots> = emeasure M A + emeasure M (B - A)"
+    by (subst plus_emeasure) auto
+  also have "emeasure M (B - A) = emeasure M B"
+    by (rule emeasure_eq_AE) (insert 2, auto)
+  finally show ?thesis .
+qed
+
+subsection \<open>\<open>\<sigma>\<close>-finite Measures\<close>
+
+locale sigma_finite_measure =
+  fixes M :: "'a measure"
+  assumes sigma_finite_countable:
+    "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets M \<and> (\<Union>A) = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
+
+lemma (in sigma_finite_measure) sigma_finite:
+  obtains A :: "nat \<Rightarrow> 'a set"
+  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+proof -
+  obtain A :: "'a set set" where
+    [simp]: "countable A" and
+    A: "A \<subseteq> sets M" "(\<Union>A) = space M" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
+    using sigma_finite_countable by metis
+  show thesis
+  proof cases
+    assume "A = {}" with \<open>(\<Union>A) = space M\<close> show thesis
+      by (intro that[of "\<lambda>_. {}"]) auto
+  next
+    assume "A \<noteq> {}"
+    show thesis
+    proof
+      show "range (from_nat_into A) \<subseteq> sets M"
+        using \<open>A \<noteq> {}\<close> A by auto
+      have "(\<Union>i. from_nat_into A i) = \<Union>A"
+        using range_from_nat_into[OF \<open>A \<noteq> {}\<close> \<open>countable A\<close>] by auto
+      with A show "(\<Union>i. from_nat_into A i) = space M"
+        by auto
+    qed (intro A from_nat_into \<open>A \<noteq> {}\<close>)
+  qed
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_disjoint:
+  obtains A :: "nat \<Rightarrow> 'a set"
+  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
+proof -
+  obtain A :: "nat \<Rightarrow> 'a set" where
+    range: "range A \<subseteq> sets M" and
+    space: "(\<Union>i. A i) = space M" and
+    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+    using sigma_finite by blast
+  show thesis
+  proof (rule that[of "disjointed A"])
+    show "range (disjointed A) \<subseteq> sets M"
+      by (rule sets.range_disjointed_sets[OF range])
+    show "(\<Union>i. disjointed A i) = space M"
+      and "disjoint_family (disjointed A)"
+      using disjoint_family_disjointed UN_disjointed_eq[of A] space range
+      by auto
+    show "emeasure M (disjointed A i) \<noteq> \<infinity>" for i
+    proof -
+      have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
+        using range disjointed_subset[of A i] by (auto intro!: emeasure_mono)
+      then show ?thesis using measure[of i] by (auto simp: top_unique)
+    qed
+  qed
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_incseq:
+  obtains A :: "nat \<Rightarrow> 'a set"
+  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
+proof -
+  obtain F :: "nat \<Rightarrow> 'a set" where
+    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
+    using sigma_finite by blast
+  show thesis
+  proof (rule that[of "\<lambda>n. \<Union>i\<le>n. F i"])
+    show "range (\<lambda>n. \<Union>i\<le>n. F i) \<subseteq> sets M"
+      using F by (force simp: incseq_def)
+    show "(\<Union>n. \<Union>i\<le>n. F i) = space M"
+    proof -
+      from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
+      with F show ?thesis by fastforce
+    qed
+    show "emeasure M (\<Union>i\<le>n. F i) \<noteq> \<infinity>" for n
+    proof -
+      have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))"
+        using F by (auto intro!: emeasure_subadditive_finite)
+      also have "\<dots> < \<infinity>"
+        using F by (auto simp: setsum_Pinfty less_top)
+      finally show ?thesis by simp
+    qed
+    show "incseq (\<lambda>n. \<Union>i\<le>n. F i)"
+      by (force simp: incseq_def)
+  qed
+qed
+
+subsection \<open>Measure space induced by distribution of @{const measurable}-functions\<close>
+
+definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
+  "distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
+
+lemma
+  shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N"
+    and space_distr[simp]: "space (distr M N f) = space N"
+  by (auto simp: distr_def)
+
+lemma
+  shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
+    and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
+  by (auto simp: measurable_def)
+
+lemma distr_cong:
+  "M = K \<Longrightarrow> sets N = sets L \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> distr M N f = distr K L g"
+  using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong)
+
+lemma emeasure_distr:
+  fixes f :: "'a \<Rightarrow> 'b"
+  assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
+  shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
+  unfolding distr_def
+proof (rule emeasure_measure_of_sigma)
+  show "positive (sets N) ?\<mu>"
+    by (auto simp: positive_def)
+
+  show "countably_additive (sets N) ?\<mu>"
+  proof (intro countably_additiveI)
+    fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
+    then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
+    then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
+      using f by (auto simp: measurable_def)
+    moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
+      using * by blast
+    moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
+      using \<open>disjoint_family A\<close> by (auto simp: disjoint_family_on_def)
+    ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
+      using suminf_emeasure[OF _ **] A f
+      by (auto simp: comp_def vimage_UN)
+  qed
+  show "sigma_algebra (space N) (sets N)" ..
+qed fact
+
+lemma emeasure_Collect_distr:
+  assumes X[measurable]: "X \<in> measurable M N" "Measurable.pred N P"
+  shows "emeasure (distr M N X) {x\<in>space N. P x} = emeasure M {x\<in>space M. P (X x)}"
+  by (subst emeasure_distr)
+     (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space])
+
+lemma emeasure_lfp2[consumes 1, case_names cont f measurable]:
+  assumes "P M"
+  assumes cont: "sup_continuous F"
+  assumes f: "\<And>M. P M \<Longrightarrow> f \<in> measurable M' M"
+  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
+  shows "emeasure M' {x\<in>space M'. lfp F (f x)} = (SUP i. emeasure M' {x\<in>space M'. (F ^^ i) (\<lambda>x. False) (f x)})"
+proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f])
+  show "f \<in> measurable M' M"  "f \<in> measurable M' M"
+    using f[OF \<open>P M\<close>] by auto
+  { fix i show "Measurable.pred M ((F ^^ i) (\<lambda>x. False))"
+    using \<open>P M\<close> by (induction i arbitrary: M) (auto intro!: *) }
+  show "Measurable.pred M (lfp F)"
+    using \<open>P M\<close> cont * by (rule measurable_lfp_coinduct[of P])
+
+  have "emeasure (distr M' M f) {x \<in> space (distr M' M f). lfp F x} =
+    (SUP i. emeasure (distr M' M f) {x \<in> space (distr M' M f). (F ^^ i) (\<lambda>x. False) x})"
+    using \<open>P M\<close>
+  proof (coinduction arbitrary: M rule: emeasure_lfp')
+    case (measurable A N) then have "\<And>N. P N \<Longrightarrow> Measurable.pred (distr M' N f) A"
+      by metis
+    then have "\<And>N. P N \<Longrightarrow> Measurable.pred N A"
+      by simp
+    with \<open>P N\<close>[THEN *] show ?case
+      by auto
+  qed fact
+  then show "emeasure (distr M' M f) {x \<in> space M. lfp F x} =
+    (SUP i. emeasure (distr M' M f) {x \<in> space M. (F ^^ i) (\<lambda>x. False) x})"
+   by simp
+qed
+
+lemma distr_id[simp]: "distr N N (\<lambda>x. x) = N"
+  by (rule measure_eqI) (auto simp: emeasure_distr)
+
+lemma measure_distr:
+  "f \<in> measurable M N \<Longrightarrow> S \<in> sets N \<Longrightarrow> measure (distr M N f) S = measure M (f -` S \<inter> space M)"
+  by (simp add: emeasure_distr measure_def)
+
+lemma distr_cong_AE:
+  assumes 1: "M = K" "sets N = sets L" and
+    2: "(AE x in M. f x = g x)" and "f \<in> measurable M N" and "g \<in> measurable K L"
+  shows "distr M N f = distr K L g"
+proof (rule measure_eqI)
+  fix A assume "A \<in> sets (distr M N f)"
+  with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A"
+    by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets)
+qed (insert 1, simp)
+
+lemma AE_distrD:
+  assumes f: "f \<in> measurable M M'"
+    and AE: "AE x in distr M M' f. P x"
+  shows "AE x in M. P (f x)"
+proof -
+  from AE[THEN AE_E] guess N .
+  with f show ?thesis
+    unfolding eventually_ae_filter
+    by (intro bexI[of _ "f -` N \<inter> space M"])
+       (auto simp: emeasure_distr measurable_def)
+qed
+
+lemma AE_distr_iff:
+  assumes f[measurable]: "f \<in> measurable M N" and P[measurable]: "{x \<in> space N. P x} \<in> sets N"
+  shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
+proof (subst (1 2) AE_iff_measurable[OF _ refl])
+  have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
+    using f[THEN measurable_space] by auto
+  then show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
+    (emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
+    by (simp add: emeasure_distr)
+qed auto
+
+lemma null_sets_distr_iff:
+  "f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N"
+  by (auto simp add: null_sets_def emeasure_distr)
+
+lemma distr_distr:
+  "g \<in> measurable N L \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> distr (distr M N f) L g = distr M L (g \<circ> f)"
+  by (auto simp add: emeasure_distr measurable_space
+           intro!: arg_cong[where f="emeasure M"] measure_eqI)
+
+subsection \<open>Real measure values\<close>
+
+lemma ring_of_finite_sets: "ring_of_sets (space M) {A\<in>sets M. emeasure M A \<noteq> top}"
+proof (rule ring_of_setsI)
+  show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
+    a \<union> b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
+    using emeasure_subadditive[of a M b] by (auto simp: top_unique)
+
+  show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
+    a - b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
+    using emeasure_mono[of "a - b" a M] by (auto simp: Diff_subset top_unique)
+qed (auto dest: sets.sets_into_space)
+
+lemma measure_nonneg[simp]: "0 \<le> measure M A"
+  unfolding measure_def by auto
+
+lemma zero_less_measure_iff: "0 < measure M A \<longleftrightarrow> measure M A \<noteq> 0"
+  using measure_nonneg[of M A] by (auto simp add: le_less)
+
+lemma measure_le_0_iff: "measure M X \<le> 0 \<longleftrightarrow> measure M X = 0"
+  using measure_nonneg[of M X] by linarith
+
+lemma measure_empty[simp]: "measure M {} = 0"
+  unfolding measure_def by (simp add: zero_ennreal.rep_eq)
+
+lemma emeasure_eq_ennreal_measure:
+  "emeasure M A \<noteq> top \<Longrightarrow> emeasure M A = ennreal (measure M A)"
+  by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def)
+
+lemma measure_zero_top: "emeasure M A = top \<Longrightarrow> measure M A = 0"
+  by (simp add: measure_def enn2ereal_top)
+
+lemma measure_eq_emeasure_eq_ennreal: "0 \<le> x \<Longrightarrow> emeasure M A = ennreal x \<Longrightarrow> measure M A = x"
+  using emeasure_eq_ennreal_measure[of M A]
+  by (cases "A \<in> M") (auto simp: measure_notin_sets emeasure_notin_sets)
+
+lemma enn2real_plus:"a < top \<Longrightarrow> b < top \<Longrightarrow> enn2real (a + b) = enn2real a + enn2real b"
+  by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top
+           del: real_of_ereal_enn2ereal)
+
+lemma measure_Union:
+  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M B \<noteq> \<infinity> \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow>
+    measure M (A \<union> B) = measure M A + measure M B"
+  by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top)
+
+lemma disjoint_family_on_insert:
+  "i \<notin> I \<Longrightarrow> disjoint_family_on A (insert i I) \<longleftrightarrow> A i \<inter> (\<Union>i\<in>I. A i) = {} \<and> disjoint_family_on A I"
+  by (fastforce simp: disjoint_family_on_def)
+
+lemma measure_finite_Union:
+  "finite S \<Longrightarrow> A`S \<subseteq> sets M \<Longrightarrow> disjoint_family_on A S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>) \<Longrightarrow>
+    measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
+  by (induction S rule: finite_induct)
+     (auto simp: disjoint_family_on_insert measure_Union setsum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])
+
+lemma measure_Diff:
+  assumes finite: "emeasure M A \<noteq> \<infinity>"
+  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
+  shows "measure M (A - B) = measure M A - measure M B"
+proof -
+  have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
+    using measurable by (auto intro!: emeasure_mono)
+  hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
+    using measurable finite by (rule_tac measure_Union) (auto simp: top_unique)
+  thus ?thesis using \<open>B \<subseteq> A\<close> by (auto simp: Un_absorb2)
+qed
+
+lemma measure_UNION:
+  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
+  assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
+proof -
+  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))"
+    unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums)
+  moreover
+  { fix i
+    have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
+      using measurable by (auto intro!: emeasure_mono)
+    then have "emeasure M (A i) = ennreal ((measure M (A i)))"
+      using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) }
+  ultimately show ?thesis using finite
+    by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all
+qed
+
+lemma measure_subadditive:
+  assumes measurable: "A \<in> sets M" "B \<in> sets M"
+  and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
+  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
+proof -
+  have "emeasure M (A \<union> B) \<noteq> \<infinity>"
+    using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique)
+  then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
+    using emeasure_subadditive[OF measurable] fin
+    apply simp
+    apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
+    apply (auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus)
+    done
+qed
+
+lemma measure_subadditive_finite:
+  assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
+  shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
+proof -
+  { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
+      using emeasure_subadditive_finite[OF A] .
+    also have "\<dots> < \<infinity>"
+      using fin by (simp add: less_top A)
+    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> top" by simp }
+  note * = this
+  show ?thesis
+    using emeasure_subadditive_finite[OF A] fin
+    unfolding emeasure_eq_ennreal_measure[OF *]
+    by (simp_all add: setsum_nonneg emeasure_eq_ennreal_measure)
+qed
+
+lemma measure_subadditive_countably:
+  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
+  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
+proof -
+  from fin have **: "\<And>i. emeasure M (A i) \<noteq> top"
+    using ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top)
+  { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
+      using emeasure_subadditive_countably[OF A] .
+    also have "\<dots> < \<infinity>"
+      using fin by (simp add: less_top)
+    finally have "emeasure M (\<Union>i. A i) \<noteq> top" by simp }
+  then have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
+    by (rule emeasure_eq_ennreal_measure[symmetric])
+  also have "\<dots> \<le> (\<Sum>i. emeasure M (A i))"
+    using emeasure_subadditive_countably[OF A] .
+  also have "\<dots> = ennreal (\<Sum>i. measure M (A i))"
+    using fin unfolding emeasure_eq_ennreal_measure[OF **]
+    by (subst suminf_ennreal) (auto simp: **)
+  finally show ?thesis
+    apply (rule ennreal_le_iff[THEN iffD1, rotated])
+    apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top)
+    using fin
+    apply (simp add: emeasure_eq_ennreal_measure[OF **])
+    done
+qed
+
+lemma measure_eq_setsum_singleton:
+  "finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
+    measure M S = (\<Sum>x\<in>S. measure M {x})"
+  using emeasure_eq_setsum_singleton[of S M]
+  by (intro measure_eq_emeasure_eq_ennreal) (auto simp: setsum_nonneg emeasure_eq_ennreal_measure)
+
+lemma Lim_measure_incseq:
+  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
+proof (rule tendsto_ennrealD)
+  have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
+    using fin by (auto simp: emeasure_eq_ennreal_measure)
+  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
+    using assms emeasure_mono[of "A _" "\<Union>i. A i" M]
+    by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans)
+  ultimately show "(\<lambda>x. ennreal (Sigma_Algebra.measure M (A x))) \<longlonglongrightarrow> ennreal (Sigma_Algebra.measure M (\<Union>i. A i))"
+    using A by (auto intro!: Lim_emeasure_incseq)
+qed auto
+
+lemma Lim_measure_decseq:
+  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
+proof (rule tendsto_ennrealD)
+  have "ennreal (measure M (\<Inter>i. A i)) = emeasure M (\<Inter>i. A i)"
+    using fin[of 0] A emeasure_mono[of "\<Inter>i. A i" "A 0" M]
+    by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans)
+  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
+    using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto
+  ultimately show "(\<lambda>x. ennreal (Sigma_Algebra.measure M (A x))) \<longlonglongrightarrow> ennreal (Sigma_Algebra.measure M (\<Inter>i. A i))"
+    using fin A by (auto intro!: Lim_emeasure_decseq)
+qed auto
+
+subsection \<open>Measure spaces with @{term "emeasure M (space M) < \<infinity>"}\<close>
+
+locale finite_measure = sigma_finite_measure M for M +
+  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> top"
+
+lemma finite_measureI[Pure.intro!]:
+  "emeasure M (space M) \<noteq> \<infinity> \<Longrightarrow> finite_measure M"
+  proof qed (auto intro!: exI[of _ "{space M}"])
+
+lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> top"
+  using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique)
+
+lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)"
+  by (intro emeasure_eq_ennreal_measure) simp
+
+lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ennreal r"
+  using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto
+
+lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
+  using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
+
+lemma (in finite_measure) finite_measure_Diff:
+  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
+  shows "measure M (A - B) = measure M A - measure M B"
+  using measure_Diff[OF _ assms] by simp
+
+lemma (in finite_measure) finite_measure_Union:
+  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
+  shows "measure M (A \<union> B) = measure M A + measure M B"
+  using measure_Union[OF _ _ assms] by simp
+
+lemma (in finite_measure) finite_measure_finite_Union:
+  assumes measurable: "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
+  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
+  using measure_finite_Union[OF assms] by simp
+
+lemma (in finite_measure) finite_measure_UNION:
+  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
+  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
+  using measure_UNION[OF A] by simp
+
+lemma (in finite_measure) finite_measure_mono:
+  assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
+  using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
+
+lemma (in finite_measure) finite_measure_subadditive:
+  assumes m: "A \<in> sets M" "B \<in> sets M"
+  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
+  using measure_subadditive[OF m] by simp
+
+lemma (in finite_measure) finite_measure_subadditive_finite:
+  assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
+  using measure_subadditive_finite[OF assms] by simp
+
+lemma (in finite_measure) finite_measure_subadditive_countably:
+  "range A \<subseteq> sets M \<Longrightarrow> summable (\<lambda>i. measure M (A i)) \<Longrightarrow> measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
+  by (rule measure_subadditive_countably)
+     (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure)
+
+lemma (in finite_measure) finite_measure_eq_setsum_singleton:
+  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
+  using measure_eq_setsum_singleton[OF assms] by simp
+
+lemma (in finite_measure) finite_Lim_measure_incseq:
+  assumes A: "range A \<subseteq> sets M" "incseq A"
+  shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
+  using Lim_measure_incseq[OF A] by simp
+
+lemma (in finite_measure) finite_Lim_measure_decseq:
+  assumes A: "range A \<subseteq> sets M" "decseq A"
+  shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
+  using Lim_measure_decseq[OF A] by simp
+
+lemma (in finite_measure) finite_measure_compl:
+  assumes S: "S \<in> sets M"
+  shows "measure M (space M - S) = measure M (space M) - measure M S"
+  using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp
+
+lemma (in finite_measure) finite_measure_mono_AE:
+  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
+  shows "measure M A \<le> measure M B"
+  using assms emeasure_mono_AE[OF imp B]
+  by (simp add: emeasure_eq_measure)
+
+lemma (in finite_measure) finite_measure_eq_AE:
+  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "measure M A = measure M B"
+  using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
+
+lemma (in finite_measure) measure_increasing: "increasing M (measure M)"
+  by (auto intro!: finite_measure_mono simp: increasing_def)
+
+lemma (in finite_measure) measure_zero_union:
+  assumes "s \<in> sets M" "t \<in> sets M" "measure M t = 0"
+  shows "measure M (s \<union> t) = measure M s"
+using assms
+proof -
+  have "measure M (s \<union> t) \<le> measure M s"
+    using finite_measure_subadditive[of s t] assms by auto
+  moreover have "measure M (s \<union> t) \<ge> measure M s"
+    using assms by (blast intro: finite_measure_mono)
+  ultimately show ?thesis by simp
+qed
+
+lemma (in finite_measure) measure_eq_compl:
+  assumes "s \<in> sets M" "t \<in> sets M"
+  assumes "measure M (space M - s) = measure M (space M - t)"
+  shows "measure M s = measure M t"
+  using assms finite_measure_compl by auto
+
+lemma (in finite_measure) measure_eq_bigunion_image:
+  assumes "range f \<subseteq> sets M" "range g \<subseteq> sets M"
+  assumes "disjoint_family f" "disjoint_family g"
+  assumes "\<And> n :: nat. measure M (f n) = measure M (g n)"
+  shows "measure M (\<Union>i. f i) = measure M (\<Union>i. g i)"
+using assms
+proof -
+  have a: "(\<lambda> i. measure M (f i)) sums (measure M (\<Union>i. f i))"
+    by (rule finite_measure_UNION[OF assms(1,3)])
+  have b: "(\<lambda> i. measure M (g i)) sums (measure M (\<Union>i. g i))"
+    by (rule finite_measure_UNION[OF assms(2,4)])
+  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
+qed
+
+lemma (in finite_measure) measure_countably_zero:
+  assumes "range c \<subseteq> sets M"
+  assumes "\<And> i. measure M (c i) = 0"
+  shows "measure M (\<Union>i :: nat. c i) = 0"
+proof (rule antisym)
+  show "measure M (\<Union>i :: nat. c i) \<le> 0"
+    using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
+qed simp
+
+lemma (in finite_measure) measure_space_inter:
+  assumes events:"s \<in> sets M" "t \<in> sets M"
+  assumes "measure M t = measure M (space M)"
+  shows "measure M (s \<inter> t) = measure M s"
+proof -
+  have "measure M ((space M - s) \<union> (space M - t)) = measure M (space M - s)"
+    using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union)
+  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
+    by blast
+  finally show "measure M (s \<inter> t) = measure M s"
+    using events by (auto intro!: measure_eq_compl[of "s \<inter> t" s])
+qed
+
+lemma (in finite_measure) measure_equiprobable_finite_unions:
+  assumes s: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> sets M"
+  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> measure M {x} = measure M {y}"
+  shows "measure M s = real (card s) * measure M {SOME x. x \<in> s}"
+proof cases
+  assume "s \<noteq> {}"
+  then have "\<exists> x. x \<in> s" by blast
+  from someI_ex[OF this] assms
+  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> measure M {x} = measure M {SOME y. y \<in> s}" by blast
+  have "measure M s = (\<Sum> x \<in> s. measure M {x})"
+    using finite_measure_eq_setsum_singleton[OF s] by simp
+  also have "\<dots> = (\<Sum> x \<in> s. measure M {SOME y. y \<in> s})" using prob_some by auto
+  also have "\<dots> = real (card s) * measure M {(SOME x. x \<in> s)}"
+    using setsum_constant assms by simp
+  finally show ?thesis by simp
+qed simp
+
+lemma (in finite_measure) measure_real_sum_image_fn:
+  assumes "e \<in> sets M"
+  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> sets M"
+  assumes "finite s"
+  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
+  assumes upper: "space M \<subseteq> (\<Union>i \<in> s. f i)"
+  shows "measure M e = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
+proof -
+  have "e \<subseteq> (\<Union>i\<in>s. f i)"
+    using \<open>e \<in> sets M\<close> sets.sets_into_space upper by blast
+  then have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
+    by auto
+  hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp
+  also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
+  proof (rule finite_measure_finite_Union)
+    show "finite s" by fact
+    show "(\<lambda>i. e \<inter> f i)`s \<subseteq> sets M" using assms(2) by auto
+    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
+      using disjoint by (auto simp: disjoint_family_on_def)
+  qed
+  finally show ?thesis .
+qed
+
+lemma (in finite_measure) measure_exclude:
+  assumes "A \<in> sets M" "B \<in> sets M"
+  assumes "measure M A = measure M (space M)" "A \<inter> B = {}"
+  shows "measure M B = 0"
+  using measure_space_inter[of B A] assms by (auto simp: ac_simps)
+lemma (in finite_measure) finite_measure_distr:
+  assumes f: "f \<in> measurable M M'"
+  shows "finite_measure (distr M M' f)"
+proof (rule finite_measureI)
+  have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
+  with f show "emeasure (distr M M' f) (space (distr M M' f)) \<noteq> \<infinity>" by (auto simp: emeasure_distr)
+qed
+
+lemma emeasure_gfp[consumes 1, case_names cont measurable]:
+  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
+  assumes "\<And>s. finite_measure (M s)"
+  assumes cont: "inf_continuous F" "inf_continuous f"
+  assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
+  assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
+  assumes bound: "\<And>P. f P \<le> f (\<lambda>s. emeasure (M s) (space (M s)))"
+  shows "emeasure (M s) {x\<in>space N. gfp F x} = gfp f s"
+proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and
+    P="Measurable.pred N", symmetric])
+  interpret finite_measure "M s" for s by fact
+  fix C assume "decseq C" "\<And>i. Measurable.pred N (C i)"
+  then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (INF i. C i) x}) = (INF i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
+    unfolding INF_apply[abs_def]
+    by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
+next
+  show "f x \<le> (\<lambda>s. emeasure (M s) {x \<in> space N. F top x})" for x
+    using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter)
+qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont)
+
+subsection \<open>Counting space\<close>
+
+lemma strict_monoI_Suc:
+  assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
+  unfolding strict_mono_def
+proof safe
+  fix n m :: nat assume "n < m" then show "f n < f m"
+    by (induct m) (auto simp: less_Suc_eq intro: less_trans ord)
+qed
+
+lemma emeasure_count_space:
+  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \<infinity>)"
+    (is "_ = ?M X")
+  unfolding count_space_def
+proof (rule emeasure_measure_of_sigma)
+  show "X \<in> Pow A" using \<open>X \<subseteq> A\<close> by auto
+  show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
+  show positive: "positive (Pow A) ?M"
+    by (auto simp: positive_def)
+  have additive: "additive (Pow A) ?M"
+    by (auto simp: card_Un_disjoint additive_def)
+
+  interpret ring_of_sets A "Pow A"
+    by (rule ring_of_setsI) auto
+  show "countably_additive (Pow A) ?M"
+    unfolding countably_additive_iff_continuous_from_below[OF positive additive]
+  proof safe
+    fix F :: "nat \<Rightarrow> 'a set" assume "incseq F"
+    show "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
+    proof cases
+      assume "\<exists>i. \<forall>j\<ge>i. F i = F j"
+      then guess i .. note i = this
+      { fix j from i \<open>incseq F\<close> have "F j \<subseteq> F i"
+          by (cases "i \<le> j") (auto simp: incseq_def) }
+      then have eq: "(\<Union>i. F i) = F i"
+        by auto
+      with i show ?thesis
+        by (auto intro!: Lim_transform_eventually[OF _ tendsto_const] eventually_sequentiallyI[where c=i])
+    next
+      assume "\<not> (\<exists>i. \<forall>j\<ge>i. F i = F j)"
+      then obtain f where f: "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis
+      then have "\<And>i. F i \<subseteq> F (f i)" using \<open>incseq F\<close> by (auto simp: incseq_def)
+      with f have *: "\<And>i. F i \<subset> F (f i)" by auto
+
+      have "incseq (\<lambda>i. ?M (F i))"
+        using \<open>incseq F\<close> unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
+      then have "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> (SUP n. ?M (F n))"
+        by (rule LIMSEQ_SUP)
+
+      moreover have "(SUP n. ?M (F n)) = top"
+      proof (rule ennreal_SUP_eq_top)
+        fix n :: nat show "\<exists>k::nat\<in>UNIV. of_nat n \<le> ?M (F k)"
+        proof (induct n)
+          case (Suc n)
+          then guess k .. note k = this
+          moreover have "finite (F k) \<Longrightarrow> finite (F (f k)) \<Longrightarrow> card (F k) < card (F (f k))"
+            using \<open>F k \<subset> F (f k)\<close> by (simp add: psubset_card_mono)
+          moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
+            using \<open>k \<le> f k\<close> \<open>incseq F\<close> by (auto simp: incseq_def dest: finite_subset)
+          ultimately show ?case
+            by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc)
+        qed auto
+      qed
+
+      moreover
+      have "inj (\<lambda>n. F ((f ^^ n) 0))"
+        by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *)
+      then have 1: "infinite (range (\<lambda>i. F ((f ^^ i) 0)))"
+        by (rule range_inj_infinite)
+      have "infinite (Pow (\<Union>i. F i))"
+        by (rule infinite_super[OF _ 1]) auto
+      then have "infinite (\<Union>i. F i)"
+        by auto
+
+      ultimately show ?thesis by auto
+    qed
+  qed
+qed
+
+lemma distr_bij_count_space:
+  assumes f: "bij_betw f A B"
+  shows "distr (count_space A) (count_space B) f = count_space B"
+proof (rule measure_eqI)
+  have f': "f \<in> measurable (count_space A) (count_space B)"
+    using f unfolding Pi_def bij_betw_def by auto
+  fix X assume "X \<in> sets (distr (count_space A) (count_space B) f)"
+  then have X: "X \<in> sets (count_space B)" by auto
+  moreover from X have "f -` X \<inter> A = the_inv_into A f ` X"
+    using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric])
+  moreover have "inj_on (the_inv_into A f) B"
+    using X f by (auto simp: bij_betw_def inj_on_the_inv_into)
+  with X have "inj_on (the_inv_into A f) X"
+    by (auto intro: subset_inj_on)
+  ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X"
+    using f unfolding emeasure_distr[OF f' X]
+    by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD)
+qed simp
+
+lemma emeasure_count_space_finite[simp]:
+  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = of_nat (card X)"
+  using emeasure_count_space[of X A] by simp
+
+lemma emeasure_count_space_infinite[simp]:
+  "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
+  using emeasure_count_space[of X A] by simp
+
+lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then of_nat (card X) else 0)"
+  by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat
+                                    measure_zero_top measure_eq_emeasure_eq_ennreal)
+
+lemma emeasure_count_space_eq_0:
+  "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
+proof cases
+  assume X: "X \<subseteq> A"
+  then show ?thesis
+  proof (intro iffI impI)
+    assume "emeasure (count_space A) X = 0"
+    with X show "X = {}"
+      by (subst (asm) emeasure_count_space) (auto split: if_split_asm)
+  qed simp
+qed (simp add: emeasure_notin_sets)
+
+lemma space_empty: "space M = {} \<Longrightarrow> M = count_space {}"
+  by (rule measure_eqI) (simp_all add: space_empty_iff)
+
+lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
+  unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
+
+lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
+  unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
+
+lemma sigma_finite_measure_count_space_countable:
+  assumes A: "countable A"
+  shows "sigma_finite_measure (count_space A)"
+  proof qed (insert A, auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"])
+
+lemma sigma_finite_measure_count_space:
+  fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)"
+  by (rule sigma_finite_measure_count_space_countable) auto
+
+lemma finite_measure_count_space:
+  assumes [simp]: "finite A"
+  shows "finite_measure (count_space A)"
+  by rule simp
+
+lemma sigma_finite_measure_count_space_finite:
+  assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
+proof -
+  interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
+  show "sigma_finite_measure (count_space A)" ..
+qed
+
+subsection \<open>Measure restricted to space\<close>
+
+lemma emeasure_restrict_space:
+  assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
+  shows "emeasure (restrict_space M \<Omega>) A = emeasure M A"
+proof (cases "A \<in> sets M")
+  case True
+  show ?thesis
+  proof (rule emeasure_measure_of[OF restrict_space_def])
+    show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
+      using \<open>A \<subseteq> \<Omega>\<close> \<open>A \<in> sets M\<close> sets.space_closed by (auto simp: sets_restrict_space)
+    show "positive (sets (restrict_space M \<Omega>)) (emeasure M)"
+      by (auto simp: positive_def)
+    show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)"
+    proof (rule countably_additiveI)
+      fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> sets (restrict_space M \<Omega>)" "disjoint_family A"
+      with assms have "\<And>i. A i \<in> sets M" "\<And>i. A i \<subseteq> space M" "disjoint_family A"
+        by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff
+                      dest: sets.sets_into_space)+
+      then show "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
+        by (subst suminf_emeasure) (auto simp: disjoint_family_subset)
+    qed
+  qed
+next
+  case False
+  with assms have "A \<notin> sets (restrict_space M \<Omega>)"
+    by (simp add: sets_restrict_space_iff)
+  with False show ?thesis
+    by (simp add: emeasure_notin_sets)
+qed
+
+lemma measure_restrict_space:
+  assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
+  shows "measure (restrict_space M \<Omega>) A = measure M A"
+  using emeasure_restrict_space[OF assms] by (simp add: measure_def)
+
+lemma AE_restrict_space_iff:
+  assumes "\<Omega> \<inter> space M \<in> sets M"
+  shows "(AE x in restrict_space M \<Omega>. P x) \<longleftrightarrow> (AE x in M. x \<in> \<Omega> \<longrightarrow> P x)"
+proof -
+  have ex_cong: "\<And>P Q f. (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> P (f x)) \<Longrightarrow> (\<exists>x. P x) \<longleftrightarrow> (\<exists>x. Q x)"
+    by auto
+  { fix X assume X: "X \<in> sets M" "emeasure M X = 0"
+    then have "emeasure M (\<Omega> \<inter> space M \<inter> X) \<le> emeasure M X"
+      by (intro emeasure_mono) auto
+    then have "emeasure M (\<Omega> \<inter> space M \<inter> X) = 0"
+      using X by (auto intro!: antisym) }
+  with assms show ?thesis
+    unfolding eventually_ae_filter
+    by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff
+                       emeasure_restrict_space cong: conj_cong
+             intro!: ex_cong[where f="\<lambda>X. (\<Omega> \<inter> space M) \<inter> X"])
+qed
+
+lemma restrict_restrict_space:
+  assumes "A \<inter> space M \<in> sets M" "B \<inter> space M \<in> sets M"
+  shows "restrict_space (restrict_space M A) B = restrict_space M (A \<inter> B)" (is "?l = ?r")
+proof (rule measure_eqI[symmetric])
+  show "sets ?r = sets ?l"
+    unfolding sets_restrict_space image_comp by (intro image_cong) auto
+next
+  fix X assume "X \<in> sets (restrict_space M (A \<inter> B))"
+  then obtain Y where "Y \<in> sets M" "X = Y \<inter> A \<inter> B"
+    by (auto simp: sets_restrict_space)
+  with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X"
+    by (subst (1 2) emeasure_restrict_space)
+       (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps)
+qed
+
+lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \<inter> B)"
+proof (rule measure_eqI)
+  show "sets (restrict_space (count_space B) A) = sets (count_space (A \<inter> B))"
+    by (subst sets_restrict_space) auto
+  moreover fix X assume "X \<in> sets (restrict_space (count_space B) A)"
+  ultimately have "X \<subseteq> A \<inter> B" by auto
+  then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \<inter> B)) X"
+    by (cases "finite X") (auto simp add: emeasure_restrict_space)
+qed
+
+lemma sigma_finite_measure_restrict_space:
+  assumes "sigma_finite_measure M"
+  and A: "A \<in> sets M"
+  shows "sigma_finite_measure (restrict_space M A)"
+proof -
+  interpret sigma_finite_measure M by fact
+  from sigma_finite_countable obtain C
+    where C: "countable C" "C \<subseteq> sets M" "(\<Union>C) = space M" "\<forall>a\<in>C. emeasure M a \<noteq> \<infinity>"
+    by blast
+  let ?C = "op \<inter> A ` C"
+  from C have "countable ?C" "?C \<subseteq> sets (restrict_space M A)" "(\<Union>?C) = space (restrict_space M A)"
+    by(auto simp add: sets_restrict_space space_restrict_space)
+  moreover {
+    fix a
+    assume "a \<in> ?C"
+    then obtain a' where "a = A \<inter> a'" "a' \<in> C" ..
+    then have "emeasure (restrict_space M A) a \<le> emeasure M a'"
+      using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono)
+    also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by (simp add: less_top)
+    finally have "emeasure (restrict_space M A) a \<noteq> \<infinity>" by simp }
+  ultimately show ?thesis
+    by unfold_locales (rule exI conjI|assumption|blast)+
+qed
+
+lemma finite_measure_restrict_space:
+  assumes "finite_measure M"
+  and A: "A \<in> sets M"
+  shows "finite_measure (restrict_space M A)"
+proof -
+  interpret finite_measure M by fact
+  show ?thesis
+    by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space)
+qed
+
+lemma restrict_distr:
+  assumes [measurable]: "f \<in> measurable M N"
+  assumes [simp]: "\<Omega> \<inter> space N \<in> sets N" and restrict: "f \<in> space M \<rightarrow> \<Omega>"
+  shows "restrict_space (distr M N f) \<Omega> = distr M (restrict_space N \<Omega>) f"
+  (is "?l = ?r")
+proof (rule measure_eqI)
+  fix A assume "A \<in> sets (restrict_space (distr M N f) \<Omega>)"
+  with restrict show "emeasure ?l A = emeasure ?r A"
+    by (subst emeasure_distr)
+       (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr
+             intro!: measurable_restrict_space2)
+qed (simp add: sets_restrict_space)
+
+lemma measure_eqI_restrict_generator:
+  assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
+  assumes sets_eq: "sets M = sets N" and \<Omega>: "\<Omega> \<in> sets M"
+  assumes "sets (restrict_space M \<Omega>) = sigma_sets \<Omega> E"
+  assumes "sets (restrict_space N \<Omega>) = sigma_sets \<Omega> E"
+  assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>"
+  assumes A: "countable A" "A \<noteq> {}" "A \<subseteq> E" "\<Union>A = \<Omega>" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
+  shows "M = N"
+proof (rule measure_eqI)
+  fix X assume X: "X \<in> sets M"
+  then have "emeasure M X = emeasure (restrict_space M \<Omega>) (X \<inter> \<Omega>)"
+    using ae \<Omega> by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE)
+  also have "restrict_space M \<Omega> = restrict_space N \<Omega>"
+  proof (rule measure_eqI_generator_eq)
+    fix X assume "X \<in> E"
+    then show "emeasure (restrict_space M \<Omega>) X = emeasure (restrict_space N \<Omega>) X"
+      using E \<Omega> by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq])
+  next
+    show "range (from_nat_into A) \<subseteq> E" "(\<Union>i. from_nat_into A i) = \<Omega>"
+      using A by (auto cong del: strong_SUP_cong)
+  next
+    fix i
+    have "emeasure (restrict_space M \<Omega>) (from_nat_into A i) = emeasure M (from_nat_into A i)"
+      using A \<Omega> by (subst emeasure_restrict_space)
+                   (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into)
+    with A show "emeasure (restrict_space M \<Omega>) (from_nat_into A i) \<noteq> \<infinity>"
+      by (auto intro: from_nat_into)
+  qed fact+
+  also have "emeasure (restrict_space N \<Omega>) (X \<inter> \<Omega>) = emeasure N X"
+    using X ae \<Omega> by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE)
+  finally show "emeasure M X = emeasure N X" .
+qed fact
+
+subsection \<open>Null measure\<close>
+
+definition "null_measure M = sigma (space M) (sets M)"
+
+lemma space_null_measure[simp]: "space (null_measure M) = space M"
+  by (simp add: null_measure_def)
+
+lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M"
+  by (simp add: null_measure_def)
+
+lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0"
+  by (cases "X \<in> sets M", rule emeasure_measure_of)
+     (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def
+           dest: sets.sets_into_space)
+
+lemma measure_null_measure[simp]: "measure (null_measure M) X = 0"
+  by (intro measure_eq_emeasure_eq_ennreal) auto
+
+lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M"
+  by(rule measure_eqI) simp_all
+
+subsection \<open>Scaling a measure\<close>
+
+definition scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+where
+  "scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. r * emeasure M A)"
+
+lemma space_scale_measure: "space (scale_measure r M) = space M"
+  by (simp add: scale_measure_def)
+
+lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M"
+  by (simp add: scale_measure_def)
+
+lemma emeasure_scale_measure [simp]:
+  "emeasure (scale_measure r M) A = r * emeasure M A"
+  (is "_ = ?\<mu> A")
+proof(cases "A \<in> sets M")
+  case True
+  show ?thesis unfolding scale_measure_def
+  proof(rule emeasure_measure_of_sigma)
+    show "sigma_algebra (space M) (sets M)" ..
+    show "positive (sets M) ?\<mu>" by (simp add: positive_def)
+    show "countably_additive (sets M) ?\<mu>"
+    proof (rule countably_additiveI)
+      fix A :: "nat \<Rightarrow> _"  assume *: "range A \<subseteq> sets M" "disjoint_family A"
+      have "(\<Sum>i. ?\<mu> (A i)) = r * (\<Sum>i. emeasure M (A i))"
+        by simp
+      also have "\<dots> = ?\<mu> (\<Union>i. A i)" using * by(simp add: suminf_emeasure)
+      finally show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)" .
+    qed
+  qed(fact True)
+qed(simp add: emeasure_notin_sets)
+
+lemma scale_measure_1 [simp]: "scale_measure 1 M = M"
+  by(rule measure_eqI) simp_all
+
+lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M"
+  by(rule measure_eqI) simp_all
+
+lemma measure_scale_measure [simp]: "0 \<le> r \<Longrightarrow> measure (scale_measure r M) A = r * measure M A"
+  using emeasure_scale_measure[of r M A]
+    emeasure_eq_ennreal_measure[of M A]
+    measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A]
+  by (cases "emeasure (scale_measure r M) A = top")
+     (auto simp del: emeasure_scale_measure
+           simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric])
+
+lemma scale_scale_measure [simp]:
+  "scale_measure r (scale_measure r' M) = scale_measure (r * r') M"
+  by (rule measure_eqI) (simp_all add: max_def mult.assoc)
+
+lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M"
+  by (rule measure_eqI) simp_all
+
+
+subsection \<open>Complete lattice structure on measures\<close>
+
+lemma (in finite_measure) finite_measure_Diff':
+  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A - B) = measure M A - measure M (A \<inter> B)"
+  using finite_measure_Diff[of A "A \<inter> B"] by (auto simp: Diff_Int)
+
+lemma (in finite_measure) finite_measure_Union':
+  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)"
+  using finite_measure_Union[of A "B - A"] by auto
+
+lemma finite_unsigned_Hahn_decomposition:
+  assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M"
+  shows "\<exists>Y\<in>sets M. (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
+proof -
+  interpret M: finite_measure M by fact
+  interpret N: finite_measure N by fact
+
+  define d where "d X = measure M X - measure N X" for X
+
+  have [intro]: "bdd_above (d`sets M)"
+    using sets.sets_into_space[of _ M]
+    by (intro bdd_aboveI[where M="measure M (space M)"])
+       (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono)
+
+  define \<gamma> where "\<gamma> = (SUP X:sets M. d X)"
+  have le_\<gamma>[intro]: "X \<in> sets M \<Longrightarrow> d X \<le> \<gamma>" for X
+    by (auto simp: \<gamma>_def intro!: cSUP_upper)
+
+  have "\<exists>f. \<forall>n. f n \<in> sets M \<and> d (f n) > \<gamma> - 1 / 2^n"
+  proof (intro choice_iff[THEN iffD1] allI)
+    fix n
+    have "\<exists>X\<in>sets M. \<gamma> - 1 / 2^n < d X"
+      unfolding \<gamma>_def by (intro less_cSUP_iff[THEN iffD1]) auto
+    then show "\<exists>y. y \<in> sets M \<and> \<gamma> - 1 / 2 ^ n < d y"
+      by auto
+  qed
+  then obtain E where [measurable]: "E n \<in> sets M" and E: "d (E n) > \<gamma> - 1 / 2^n" for n
+    by auto
+
+  define F where "F m n = (if m \<le> n then \<Inter>i\<in>{m..n}. E i else space M)" for m n
+
+  have [measurable]: "m \<le> n \<Longrightarrow> F m n \<in> sets M" for m n
+    by (auto simp: F_def)
+
+  have 1: "\<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)" if "m \<le> n" for m n
+    using that
+  proof (induct rule: dec_induct)
+    case base with E[of m] show ?case
+      by (simp add: F_def field_simps)
+  next
+    case (step i)
+    have F_Suc: "F m (Suc i) = F m i \<inter> E (Suc i)"
+      using \<open>m \<le> i\<close> by (auto simp: F_def le_Suc_eq)
+
+    have "\<gamma> + (\<gamma> - 2 / 2^m + 1 / 2 ^ Suc i) \<le> (\<gamma> - 1 / 2^Suc i) + (\<gamma> - 2 / 2^m + 1 / 2^i)"
+      by (simp add: field_simps)
+    also have "\<dots> \<le> d (E (Suc i)) + d (F m i)"
+      using E[of "Suc i"] by (intro add_mono step) auto
+    also have "\<dots> = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))"
+      using \<open>m \<le> i\<close> by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff')
+    also have "\<dots> = d (E (Suc i) \<union> F m i) + d (F m (Suc i))"
+      using \<open>m \<le> i\<close> by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union')
+    also have "\<dots> \<le> \<gamma> + d (F m (Suc i))"
+      using \<open>m \<le> i\<close> by auto
+    finally show ?case
+      by auto
+  qed
+
+  define F' where "F' m = (\<Inter>i\<in>{m..}. E i)" for m
+  have F'_eq: "F' m = (\<Inter>i. F m (i + m))" for m
+    by (fastforce simp: le_iff_add[of m] F'_def F_def)
+
+  have [measurable]: "F' m \<in> sets M" for m
+    by (auto simp: F'_def)
+
+  have \<gamma>_le: "\<gamma> - 0 \<le> d (\<Union>m. F' m)"
+  proof (rule LIMSEQ_le)
+    show "(\<lambda>n. \<gamma> - 2 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 0"
+      by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto
+    have "incseq F'"
+      by (auto simp: incseq_def F'_def)
+    then show "(\<lambda>m. d (F' m)) \<longlonglongrightarrow> d (\<Union>m. F' m)"
+      unfolding d_def
+      by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto
+
+    have "\<gamma> - 2 / 2 ^ m + 0 \<le> d (F' m)" for m
+    proof (rule LIMSEQ_le)
+      have *: "decseq (\<lambda>n. F m (n + m))"
+        by (auto simp: decseq_def F_def)
+      show "(\<lambda>n. d (F m n)) \<longlonglongrightarrow> d (F' m)"
+        unfolding d_def F'_eq
+        by (rule LIMSEQ_offset[where k=m])
+           (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *)
+      show "(\<lambda>n. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 2 / 2 ^ m + 0"
+        by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto
+      show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)"
+        using 1[of m] by (intro exI[of _ m]) auto
+    qed
+    then show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ n \<le> d (F' n)"
+      by auto
+  qed
+
+  show ?thesis
+  proof (safe intro!: bexI[of _ "\<Union>m. F' m"])
+    fix X assume [measurable]: "X \<in> sets M" and X: "X \<subseteq> (\<Union>m. F' m)"
+    have "d (\<Union>m. F' m) - d X = d ((\<Union>m. F' m) - X)"
+      using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff)
+    also have "\<dots> \<le> \<gamma>"
+      by auto
+    finally have "0 \<le> d X"
+      using \<gamma>_le by auto
+    then show "emeasure N X \<le> emeasure M X"
+      by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
+  next
+    fix X assume [measurable]: "X \<in> sets M" and X: "X \<inter> (\<Union>m. F' m) = {}"
+    then have "d (\<Union>m. F' m) + d X = d (X \<union> (\<Union>m. F' m))"
+      by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union)
+    also have "\<dots> \<le> \<gamma>"
+      by auto
+    finally have "d X \<le> 0"
+      using \<gamma>_le by auto
+    then show "emeasure M X \<le> emeasure N X"
+      by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
+  qed auto
+qed
+
+lemma unsigned_Hahn_decomposition:
+  assumes [simp]: "sets N = sets M" and [measurable]: "A \<in> sets M"
+    and [simp]: "emeasure M A \<noteq> top" "emeasure N A \<noteq> top"
+  shows "\<exists>Y\<in>sets M. Y \<subseteq> A \<and> (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<subseteq> A \<longrightarrow> X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
+proof -
+  have "\<exists>Y\<in>sets (restrict_space M A).
+    (\<forall>X\<in>sets (restrict_space M A). X \<subseteq> Y \<longrightarrow> (restrict_space N A) X \<le> (restrict_space M A) X) \<and>
+    (\<forall>X\<in>sets (restrict_space M A). X \<inter> Y = {} \<longrightarrow> (restrict_space M A) X \<le> (restrict_space N A) X)"
+  proof (rule finite_unsigned_Hahn_decomposition)
+    show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)"
+      by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI)
+  qed (simp add: sets_restrict_space)
+  then guess Y ..
+  then show ?thesis
+    apply (intro bexI[of _ Y] conjI ballI conjI)
+    apply (simp_all add: sets_restrict_space emeasure_restrict_space)
+    apply safe
+    subgoal for X Z
+      by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1)
+    subgoal for X Z
+      by (erule ballE[of _ _ X])  (auto simp add: Int_absorb1 ac_simps)
+    apply auto
+    done
+qed
+
+text \<open>
+  Define a lexicographical order on @{type measure}, in the order space, sets and measure. The parts
+  of the lexicographical order are point-wise ordered.
+\<close>
+
+instantiation measure :: (type) order_bot
+begin
+
+inductive less_eq_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
+  "space M \<subset> space N \<Longrightarrow> less_eq_measure M N"
+| "space M = space N \<Longrightarrow> sets M \<subset> sets N \<Longrightarrow> less_eq_measure M N"
+| "space M = space N \<Longrightarrow> sets M = sets N \<Longrightarrow> emeasure M \<le> emeasure N \<Longrightarrow> less_eq_measure M N"
+
+lemma le_measure_iff:
+  "M \<le> N \<longleftrightarrow> (if space M = space N then
+    if sets M = sets N then emeasure M \<le> emeasure N else sets M \<subseteq> sets N else space M \<subseteq> space N)"
+  by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros)
+
+definition less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
+  "less_measure M N \<longleftrightarrow> (M \<le> N \<and> \<not> N \<le> M)"
+
+definition bot_measure :: "'a measure" where
+  "bot_measure = sigma {} {}"
+
+lemma
+  shows space_bot[simp]: "space bot = {}"
+    and sets_bot[simp]: "sets bot = {{}}"
+    and emeasure_bot[simp]: "emeasure bot X = 0"
+  by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma)
+
+instance
+proof standard
+  show "bot \<le> a" for a :: "'a measure"
+    by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def)
+qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI)
+
+end
+
+lemma le_measure: "sets M = sets N \<Longrightarrow> M \<le> N \<longleftrightarrow> (\<forall>A\<in>sets M. emeasure M A \<le> emeasure N A)"
+  apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq)
+  subgoal for X
+    by (cases "X \<in> sets M") (auto simp: emeasure_notin_sets)
+  done
+
+definition sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+where
+  "sup_measure' A B = measure_of (space A) (sets A) (\<lambda>X. SUP Y:sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
+
+lemma assumes [simp]: "sets B = sets A"
+  shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A"
+    and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A"
+  using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def)
+
+lemma emeasure_sup_measure':
+  assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \<in> sets A"
+  shows "emeasure (sup_measure' A B) X = (SUP Y:sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
+    (is "_ = ?S X")
+proof -
+  note sets_eq_imp_space_eq[OF sets_eq, simp]
+  show ?thesis
+    using sup_measure'_def
+  proof (rule emeasure_measure_of)
+    let ?d = "\<lambda>X Y. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)"
+    show "countably_additive (sets (sup_measure' A B)) (\<lambda>X. SUP Y : sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
+    proof (rule countably_additiveI, goal_cases)
+      case (1 X)
+      then have [measurable]: "\<And>i. X i \<in> sets A" and "disjoint_family X"
+        by auto
+      have "(\<Sum>i. ?S (X i)) = (SUP Y:sets A. \<Sum>i. ?d (X i) Y)"
+      proof (rule ennreal_suminf_SUP_eq_directed)
+        fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \<in> sets A" "b \<in> sets A"
+        have "\<exists>c\<in>sets A. c \<subseteq> X i \<and> (\<forall>a\<in>sets A. ?d (X i) a \<le> ?d (X i) c)" for i
+        proof cases
+          assume "emeasure A (X i) = top \<or> emeasure B (X i) = top"
+          then show ?thesis
+          proof
+            assume "emeasure A (X i) = top" then show ?thesis
+              by (intro bexI[of _ "X i"]) auto
+          next
+            assume "emeasure B (X i) = top" then show ?thesis
+              by (intro bexI[of _ "{}"]) auto
+          qed
+        next
+          assume finite: "\<not> (emeasure A (X i) = top \<or> emeasure B (X i) = top)"
+          then have "\<exists>Y\<in>sets A. Y \<subseteq> X i \<and> (\<forall>C\<in>sets A. C \<subseteq> Y \<longrightarrow> B C \<le> A C) \<and> (\<forall>C\<in>sets A. C \<subseteq> X i \<longrightarrow> C \<inter> Y = {} \<longrightarrow> A C \<le> B C)"
+            using unsigned_Hahn_decomposition[of B A "X i"] by simp
+          then obtain Y where [measurable]: "Y \<in> sets A" and [simp]: "Y \<subseteq> X i"
+            and B_le_A: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> B C \<le> A C"
+            and A_le_B: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> X i \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> A C \<le> B C"
+            by auto
+
+          show ?thesis
+          proof (intro bexI[of _ Y] ballI conjI)
+            fix a assume [measurable]: "a \<in> sets A"
+            have *: "(X i \<inter> a \<inter> Y \<union> (X i \<inter> a - Y)) = X i \<inter> a" "(X i - a) \<inter> Y \<union> (X i - a - Y) = X i \<inter> - a"
+              for a Y by auto
+            then have "?d (X i) a =
+              (A (X i \<inter> a \<inter> Y) + A (X i \<inter> a \<inter> - Y)) + (B (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))"
+              by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric])
+            also have "\<dots> \<le> (A (X i \<inter> a \<inter> Y) + B (X i \<inter> a \<inter> - Y)) + (A (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))"
+              by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric])
+            also have "\<dots> \<le> (A (X i \<inter> Y \<inter> a) + A (X i \<inter> Y \<inter> - a)) + (B (X i \<inter> - Y \<inter> a) + B (X i \<inter> - Y \<inter> - a))"
+              by (simp add: ac_simps)
+            also have "\<dots> \<le> A (X i \<inter> Y) + B (X i \<inter> - Y)"
+              by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *)
+            finally show "?d (X i) a \<le> ?d (X i) Y" .
+          qed auto
+        qed
+        then obtain C where [measurable]: "C i \<in> sets A" and "C i \<subseteq> X i"
+          and C: "\<And>a. a \<in> sets A \<Longrightarrow> ?d (X i) a \<le> ?d (X i) (C i)" for i
+          by metis
+        have *: "X i \<inter> (\<Union>i. C i) = X i \<inter> C i" for i
+        proof safe
+          fix x j assume "x \<in> X i" "x \<in> C j"
+          moreover have "i = j \<or> X i \<inter> X j = {}"
+            using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
+          ultimately show "x \<in> C i"
+            using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
+        qed auto
+        have **: "X i \<inter> - (\<Union>i. C i) = X i \<inter> - C i" for i
+        proof safe
+          fix x j assume "x \<in> X i" "x \<notin> C i" "x \<in> C j"
+          moreover have "i = j \<or> X i \<inter> X j = {}"
+            using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
+          ultimately show False
+            using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
+        qed auto
+        show "\<exists>c\<in>sets A. \<forall>i\<in>J. ?d (X i) a \<le> ?d (X i) c \<and> ?d (X i) b \<le> ?d (X i) c"
+          apply (intro bexI[of _ "\<Union>i. C i"])
+          unfolding * **
+          apply (intro C ballI conjI)
+          apply auto
+          done
+      qed
+      also have "\<dots> = ?S (\<Union>i. X i)"
+        unfolding UN_extend_simps(4)
+        by (auto simp add: suminf_add[symmetric] Diff_eq[symmetric] simp del: UN_simps
+                 intro!: SUP_cong arg_cong2[where f="op +"] suminf_emeasure
+                         disjoint_family_on_bisimulation[OF \<open>disjoint_family X\<close>])
+      finally show "(\<Sum>i. ?S (X i)) = ?S (\<Union>i. X i)" .
+    qed
+  qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const)
+qed
+
+lemma le_emeasure_sup_measure'1:
+  assumes "sets B = sets A" "X \<in> sets A" shows "emeasure A X \<le> emeasure (sup_measure' A B) X"
+  by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms)
+
+lemma le_emeasure_sup_measure'2:
+  assumes "sets B = sets A" "X \<in> sets A" shows "emeasure B X \<le> emeasure (sup_measure' A B) X"
+  by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms)
+
+lemma emeasure_sup_measure'_le2:
+  assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \<in> sets C"
+  assumes A: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure A Y \<le> emeasure C Y"
+  assumes B: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure B Y \<le> emeasure C Y"
+  shows "emeasure (sup_measure' A B) X \<le> emeasure C X"
+proof (subst emeasure_sup_measure')
+  show "(SUP Y:sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)) \<le> emeasure C X"
+    unfolding \<open>sets A = sets C\<close>
+  proof (intro SUP_least)
+    fix Y assume [measurable]: "Y \<in> sets C"
+    have [simp]: "X \<inter> Y \<union> (X - Y) = X"
+      by auto
+    have "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C (X \<inter> Y) + emeasure C (X \<inter> - Y)"
+      by (intro add_mono A B) (auto simp: Diff_eq[symmetric])
+    also have "\<dots> = emeasure C X"
+      by (subst plus_emeasure) (auto simp: Diff_eq[symmetric])
+    finally show "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C X" .
+  qed
+qed simp_all
+
+definition sup_lexord :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+where
+  "sup_lexord A B k s c =
+    (if k A = k B then c else if \<not> k A \<le> k B \<and> \<not> k B \<le> k A then s else if k B \<le> k A then A else B)"
+
+lemma sup_lexord:
+  "(k A < k B \<Longrightarrow> P B) \<Longrightarrow> (k B < k A \<Longrightarrow> P A) \<Longrightarrow> (k A = k B \<Longrightarrow> P c) \<Longrightarrow>
+    (\<not> k B \<le> k A \<Longrightarrow> \<not> k A \<le> k B \<Longrightarrow> P s) \<Longrightarrow> P (sup_lexord A B k s c)"
+  by (auto simp: sup_lexord_def)
+
+lemmas le_sup_lexord = sup_lexord[where P="\<lambda>a. c \<le> a" for c]
+
+lemma sup_lexord1: "k A = k B \<Longrightarrow> sup_lexord A B k s c = c"
+  by (simp add: sup_lexord_def)
+
+lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c"
+  by (auto simp: sup_lexord_def)
+
+lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \<A> \<subseteq> sets x) = (\<A> \<subseteq> sets x)"
+  using sets.sigma_sets_subset[of \<A> x] by auto
+
+lemma sigma_le_iff: "\<A> \<subseteq> Pow \<Omega> \<Longrightarrow> sigma \<Omega> \<A> \<le> x \<longleftrightarrow> (\<Omega> \<subseteq> space x \<and> (space x = \<Omega> \<longrightarrow> \<A> \<subseteq> sets x))"
+  by (cases "\<Omega> = space x")
+     (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def
+                    sigma_sets_superset_generator sigma_sets_le_sets_iff)
+
+instantiation measure :: (type) semilattice_sup
+begin
+
+definition sup_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+where
+  "sup_measure A B =
+    sup_lexord A B space (sigma (space A \<union> space B) {})
+      (sup_lexord A B sets (sigma (space A) (sets A \<union> sets B)) (sup_measure' A B))"
+
+instance
+proof
+  fix x y z :: "'a measure"
+  show "x \<le> sup x y"
+    unfolding sup_measure_def
+  proof (intro le_sup_lexord)
+    assume "space x = space y"
+    then have *: "sets x \<union> sets y \<subseteq> Pow (space x)"
+      using sets.space_closed by auto
+    assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y"
+    then have "sets x \<subset> sets x \<union> sets y"
+      by auto
+    also have "\<dots> \<le> sigma (space x) (sets x \<union> sets y)"
+      by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
+    finally show "x \<le> sigma (space x) (sets x \<union> sets y)"
+      by (simp add: space_measure_of[OF *] less_eq_measure.intros(2))
+  next
+    assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y"
+    then show "x \<le> sigma (space x \<union> space y) {}"
+      by (intro less_eq_measure.intros) auto
+  next
+    assume "sets x = sets y" then show "x \<le> sup_measure' x y"
+      by (simp add: le_measure le_emeasure_sup_measure'1)
+  qed (auto intro: less_eq_measure.intros)
+  show "y \<le> sup x y"
+    unfolding sup_measure_def
+  proof (intro le_sup_lexord)
+    assume **: "space x = space y"
+    then have *: "sets x \<union> sets y \<subseteq> Pow (space y)"
+      using sets.space_closed by auto
+    assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y"
+    then have "sets y \<subset> sets x \<union> sets y"
+      by auto
+    also have "\<dots> \<le> sigma (space y) (sets x \<union> sets y)"
+      by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
+    finally show "y \<le> sigma (space x) (sets x \<union> sets y)"
+      by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2))
+  next
+    assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y"
+    then show "y \<le> sigma (space x \<union> space y) {}"
+      by (intro less_eq_measure.intros) auto
+  next
+    assume "sets x = sets y" then show "y \<le> sup_measure' x y"
+      by (simp add: le_measure le_emeasure_sup_measure'2)
+  qed (auto intro: less_eq_measure.intros)
+  show "x \<le> y \<Longrightarrow> z \<le> y \<Longrightarrow> sup x z \<le> y"
+    unfolding sup_measure_def
+  proof (intro sup_lexord[where P="\<lambda>x. x \<le> y"])
+    assume "x \<le> y" "z \<le> y" and [simp]: "space x = space z" "sets x = sets z"
+    from \<open>x \<le> y\<close> show "sup_measure' x z \<le> y"
+    proof cases
+      case 1 then show ?thesis
+        by (intro less_eq_measure.intros(1)) simp
+    next
+      case 2 then show ?thesis
+        by (intro less_eq_measure.intros(2)) simp_all
+    next
+      case 3 with \<open>z \<le> y\<close> \<open>x \<le> y\<close> show ?thesis
+        by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2)
+    qed
+  next
+    assume **: "x \<le> y" "z \<le> y" "space x = space z" "\<not> sets z \<subseteq> sets x" "\<not> sets x \<subseteq> sets z"
+    then have *: "sets x \<union> sets z \<subseteq> Pow (space x)"
+      using sets.space_closed by auto
+    show "sigma (space x) (sets x \<union> sets z) \<le> y"
+      unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm)
+  next
+    assume "x \<le> y" "z \<le> y" "\<not> space z \<subseteq> space x" "\<not> space x \<subseteq> space z"
+    then have "space x \<subseteq> space y" "space z \<subseteq> space y"
+      by (auto simp: le_measure_iff split: if_split_asm)
+    then show "sigma (space x \<union> space z) {} \<le> y"
+      by (simp add: sigma_le_iff)
+  qed
+qed
+
+end
+
+lemma space_empty_eq_bot: "space a = {} \<longleftrightarrow> a = bot"
+  using space_empty[of a] by (auto intro!: measure_eqI)
+
+interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure"
+  proof qed (auto intro!: antisym)
+
+lemma sup_measure_F_mono':
+  "finite J \<Longrightarrow> finite I \<Longrightarrow> sup_measure.F id I \<le> sup_measure.F id (I \<union> J)"
+proof (induction J rule: finite_induct)
+  case empty then show ?case
+    by simp
+next
+  case (insert i J)
+  show ?case
+  proof cases
+    assume "i \<in> I" with insert show ?thesis
+      by (auto simp: insert_absorb)
+  next
+    assume "i \<notin> I"
+    have "sup_measure.F id I \<le> sup_measure.F id (I \<union> J)"
+      by (intro insert)
+    also have "\<dots> \<le> sup_measure.F id (insert i (I \<union> J))"
+      using insert \<open>i \<notin> I\<close> by (subst sup_measure.insert) auto
+    finally show ?thesis
+      by auto
+  qed
+qed
+
+lemma sup_measure_F_mono: "finite I \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sup_measure.F id J \<le> sup_measure.F id I"
+  using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1)
+
+lemma sets_eq_iff_bounded: "A \<le> B \<Longrightarrow> B \<le> C \<Longrightarrow> sets A = sets C \<Longrightarrow> sets B = sets A"
+  by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm)
+
+lemma sets_sup: "sets A = sets M \<Longrightarrow> sets B = sets M \<Longrightarrow> sets (sup A B) = sets M"
+  by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq)
+
+lemma sets_sup_measure_F:
+  "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> sets i = sets M) \<Longrightarrow> sets (sup_measure.F id I) = sets M"
+  by (induction I rule: finite_ne_induct) (simp_all add: sets_sup)
+
+lemma le_measureD1: "A \<le> B \<Longrightarrow> space A \<le> space B"
+  by (auto simp: le_measure_iff split: if_split_asm)
+
+lemma le_measureD2: "A \<le> B \<Longrightarrow> space A = space B \<Longrightarrow> sets A \<le> sets B"
+  by (auto simp: le_measure_iff split: if_split_asm)
+
+lemma le_measureD3: "A \<le> B \<Longrightarrow> sets A = sets B \<Longrightarrow> emeasure A X \<le> emeasure B X"
+  by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm)
+
+definition Sup_measure' :: "'a measure set \<Rightarrow> 'a measure"
+where
+  "Sup_measure' M = measure_of (\<Union>a\<in>M. space a) (\<Union>a\<in>M. sets a)
+    (\<lambda>X. (SUP P:{P. finite P \<and> P \<subseteq> M }. sup_measure.F id P X))"
+
+lemma UN_space_closed: "UNION S sets \<subseteq> Pow (UNION S space)"
+  using sets.space_closed by auto
+
+lemma space_Sup_measure'2: "space (Sup_measure' M) = (\<Union>m\<in>M. space m)"
+  unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed])
+
+lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)"
+  unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed])
+
+lemma sets_Sup_measure':
+  assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
+  shows "sets (Sup_measure' M) = sets A"
+  using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close> by (simp add: Sup_measure'_def)
+
+lemma space_Sup_measure':
+  assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
+  shows "space (Sup_measure' M) = space A"
+  using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close>
+  by (simp add: Sup_measure'_def )
+
+lemma emeasure_Sup_measure':
+  assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "X \<in> sets A" "M \<noteq> {}"
+  shows "emeasure (Sup_measure' M) X = (SUP P:{P. finite P \<and> P \<subseteq> M}. sup_measure.F id P X)"
+    (is "_ = ?S X")
+  using Sup_measure'_def
+proof (rule emeasure_measure_of)
+  note sets_eq[THEN sets_eq_imp_space_eq, simp]
+  have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A"
+    using \<open>M \<noteq> {}\<close> by (simp_all add: Sup_measure'_def)
+  let ?\<mu> = "sup_measure.F id"
+  show "countably_additive (sets (Sup_measure' M)) ?S"
+  proof (rule countably_additiveI, goal_cases)
+    case (1 F)
+    then have **: "range F \<subseteq> sets A"
+      by (auto simp: *)
+    show "(\<Sum>i. ?S (F i)) = ?S (\<Union>i. F i)"
+    proof (subst ennreal_suminf_SUP_eq_directed)
+      fix i j and N :: "nat set" assume ij: "i \<in> {P. finite P \<and> P \<subseteq> M}" "j \<in> {P. finite P \<and> P \<subseteq> M}"
+      have "(i \<noteq> {} \<longrightarrow> sets (?\<mu> i) = sets A) \<and> (j \<noteq> {} \<longrightarrow> sets (?\<mu> j) = sets A) \<and>
+        (i \<noteq> {} \<or> j \<noteq> {} \<longrightarrow> sets (?\<mu> (i \<union> j)) = sets A)"
+        using ij by (intro impI sets_sup_measure_F conjI) auto
+      then have "?\<mu> j (F n) \<le> ?\<mu> (i \<union> j) (F n) \<and> ?\<mu> i (F n) \<le> ?\<mu> (i \<union> j) (F n)" for n
+        using ij
+        by (cases "i = {}"; cases "j = {}")
+           (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F
+                 simp del: id_apply)
+      with ij show "\<exists>k\<in>{P. finite P \<and> P \<subseteq> M}. \<forall>n\<in>N. ?\<mu> i (F n) \<le> ?\<mu> k (F n) \<and> ?\<mu> j (F n) \<le> ?\<mu> k (F n)"
+        by (safe intro!: bexI[of _ "i \<union> j"]) auto
+    next
+      show "(SUP P : {P. finite P \<and> P \<subseteq> M}. \<Sum>n. ?\<mu> P (F n)) = (SUP P : {P. finite P \<and> P \<subseteq> M}. ?\<mu> P (UNION UNIV F))"
+      proof (intro SUP_cong refl)
+        fix i assume i: "i \<in> {P. finite P \<and> P \<subseteq> M}"
+        show "(\<Sum>n. ?\<mu> i (F n)) = ?\<mu> i (UNION UNIV F)"
+        proof cases
+          assume "i \<noteq> {}" with i ** show ?thesis
+            apply (intro suminf_emeasure \<open>disjoint_family F\<close>)
+            apply (subst sets_sup_measure_F[OF _ _ sets_eq])
+            apply auto
+            done
+        qed simp
+      qed
+    qed
+  qed
+  show "positive (sets (Sup_measure' M)) ?S"
+    by (auto simp: positive_def bot_ennreal[symmetric])
+  show "X \<in> sets (Sup_measure' M)"
+    using assms * by auto
+qed (rule UN_space_closed)
+
+definition Sup_lexord :: "('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> 'a"
+where
+  "Sup_lexord k c s A = (let U = (SUP a:A. k a) in if \<exists>a\<in>A. k a = U then c {a\<in>A. k a = U} else s A)"
+
+lemma Sup_lexord:
+  "(\<And>a S. a \<in> A \<Longrightarrow> k a = (SUP a:A. k a) \<Longrightarrow> S = {a'\<in>A. k a' = k a} \<Longrightarrow> P (c S)) \<Longrightarrow> ((\<And>a. a \<in> A \<Longrightarrow> k a \<noteq> (SUP a:A. k a)) \<Longrightarrow> P (s A)) \<Longrightarrow>
+    P (Sup_lexord k c s A)"
+  by (auto simp: Sup_lexord_def Let_def)
+
+lemma Sup_lexord1:
+  assumes A: "A \<noteq> {}" "(\<And>a. a \<in> A \<Longrightarrow> k a = (\<Union>a\<in>A. k a))" "P (c A)"
+  shows "P (Sup_lexord k c s A)"
+  unfolding Sup_lexord_def Let_def
+proof (clarsimp, safe)
+  show "\<forall>a\<in>A. k a \<noteq> (\<Union>x\<in>A. k x) \<Longrightarrow> P (s A)"
+    by (metis assms(1,2) ex_in_conv)
+next
+  fix a assume "a \<in> A" "k a = (\<Union>x\<in>A. k x)"
+  then have "{a \<in> A. k a = (\<Union>x\<in>A. k x)} = {a \<in> A. k a = k a}"
+    by (metis A(2)[symmetric])
+  then show "P (c {a \<in> A. k a = (\<Union>x\<in>A. k x)})"
+    by (simp add: A(3))
+qed
+
+instantiation measure :: (type) complete_lattice
+begin
+
+definition Sup_measure :: "'a measure set \<Rightarrow> 'a measure"
+where
+  "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure'
+    (\<lambda>U. sigma (\<Union>u\<in>U. space u) (\<Union>u\<in>U. sets u))) (\<lambda>U. sigma (\<Union>u\<in>U. space u) {})"
+
+definition Inf_measure :: "'a measure set \<Rightarrow> 'a measure"
+where
+  "Inf_measure A = Sup {x. \<forall>a\<in>A. x \<le> a}"
+
+definition inf_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+where
+  "inf_measure a b = Inf {a, b}"
+
+definition top_measure :: "'a measure"
+where
+  "top_measure = Inf {}"
+
+instance
+proof
+  note UN_space_closed [simp]
+  show upper: "x \<le> Sup A" if x: "x \<in> A" for x :: "'a measure" and A
+    unfolding Sup_measure_def
+  proof (intro Sup_lexord[where P="\<lambda>y. x \<le> y"])
+    assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)"
+    from this[OF \<open>x \<in> A\<close>] \<open>x \<in> A\<close> show "x \<le> sigma (\<Union>a\<in>A. space a) {}"
+      by (intro less_eq_measure.intros) auto
+  next
+    fix a S assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
+      and neq: "\<And>aa. aa \<in> S \<Longrightarrow> sets aa \<noteq> (\<Union>a\<in>S. sets a)"
+    have sp_a: "space a = (UNION S space)"
+      using \<open>a\<in>A\<close> by (auto simp: S)
+    show "x \<le> sigma (UNION S space) (UNION S sets)"
+    proof cases
+      assume [simp]: "space x = space a"
+      have "sets x \<subset> (\<Union>a\<in>S. sets a)"
+        using \<open>x\<in>A\<close> neq[of x] by (auto simp: S)
+      also have "\<dots> \<subseteq> sigma_sets (\<Union>x\<in>S. space x) (\<Union>x\<in>S. sets x)"
+        by (rule sigma_sets_superset_generator)
+      finally show ?thesis
+        by (intro less_eq_measure.intros(2)) (simp_all add: sp_a)
+    next
+      assume "space x \<noteq> space a"
+      moreover have "space x \<le> space a"
+        unfolding a using \<open>x\<in>A\<close> by auto
+      ultimately show ?thesis
+        by (intro less_eq_measure.intros) (simp add: less_le sp_a)
+    qed
+  next
+    fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
+      and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
+    then have "S' \<noteq> {}" "space b = space a"
+      by auto
+    have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b"
+      by (auto simp: S')
+    note sets_eq[THEN sets_eq_imp_space_eq, simp]
+    have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
+      using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
+    show "x \<le> Sup_measure' S'"
+    proof cases
+      assume "x \<in> S"
+      with \<open>b \<in> S\<close> have "space x = space b"
+        by (simp add: S)
+      show ?thesis
+      proof cases
+        assume "x \<in> S'"
+        show "x \<le> Sup_measure' S'"
+        proof (intro le_measure[THEN iffD2] ballI)
+          show "sets x = sets (Sup_measure' S')"
+            using \<open>x\<in>S'\<close> * by (simp add: S')
+          fix X assume "X \<in> sets x"
+          show "emeasure x X \<le> emeasure (Sup_measure' S') X"
+          proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets x\<close>])
+            show "emeasure x X \<le> (SUP P : {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X)"
+              using \<open>x\<in>S'\<close> by (intro SUP_upper2[where i="{x}"]) auto
+          qed (insert \<open>x\<in>S'\<close> S', auto)
+        qed
+      next
+        assume "x \<notin> S'"
+        then have "sets x \<noteq> sets b"
+          using \<open>x\<in>S\<close> by (auto simp: S')
+        moreover have "sets x \<le> sets b"
+          using \<open>x\<in>S\<close> unfolding b by auto
+        ultimately show ?thesis
+          using * \<open>x \<in> S\<close>
+          by (intro less_eq_measure.intros(2))
+             (simp_all add: * \<open>space x = space b\<close> less_le)
+      qed
+    next
+      assume "x \<notin> S"
+      with \<open>x\<in>A\<close> \<open>x \<notin> S\<close> \<open>space b = space a\<close> show ?thesis
+        by (intro less_eq_measure.intros)
+           (simp_all add: * less_le a SUP_upper S)
+    qed
+  qed
+  show least: "Sup A \<le> x" if x: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x" for x :: "'a measure" and A
+    unfolding Sup_measure_def
+  proof (intro Sup_lexord[where P="\<lambda>y. y \<le> x"])
+    assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)"
+    show "sigma (UNION A space) {} \<le> x"
+      using x[THEN le_measureD1] by (subst sigma_le_iff) auto
+  next
+    fix a S assume "a \<in> A" "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
+      "\<And>a. a \<in> S \<Longrightarrow> sets a \<noteq> (\<Union>a\<in>S. sets a)"
+    have "UNION S space \<subseteq> space x"
+      using S le_measureD1[OF x] by auto
+    moreover
+    have "UNION S space = space a"
+      using \<open>a\<in>A\<close> S by auto
+    then have "space x = UNION S space \<Longrightarrow> UNION S sets \<subseteq> sets x"
+      using \<open>a \<in> A\<close> le_measureD2[OF x] by (auto simp: S)
+    ultimately show "sigma (UNION S space) (UNION S sets) \<le> x"
+      by (subst sigma_le_iff) simp_all
+  next
+    fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
+      and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
+    then have "S' \<noteq> {}" "space b = space a"
+      by auto
+    have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b"
+      by (auto simp: S')
+    note sets_eq[THEN sets_eq_imp_space_eq, simp]
+    have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
+      using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
+    show "Sup_measure' S' \<le> x"
+    proof cases
+      assume "space x = space a"
+      show ?thesis
+      proof cases
+        assume **: "sets x = sets b"
+        show ?thesis
+        proof (intro le_measure[THEN iffD2] ballI)
+          show ***: "sets (Sup_measure' S') = sets x"
+            by (simp add: * **)
+          fix X assume "X \<in> sets (Sup_measure' S')"
+          show "emeasure (Sup_measure' S') X \<le> emeasure x X"
+            unfolding ***
+          proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets (Sup_measure' S')\<close>])
+            show "(SUP P : {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X) \<le> emeasure x X"
+            proof (safe intro!: SUP_least)
+              fix P assume P: "finite P" "P \<subseteq> S'"
+              show "emeasure (sup_measure.F id P) X \<le> emeasure x X"
+              proof cases
+                assume "P = {}" then show ?thesis
+                  by auto
+              next
+                assume "P \<noteq> {}"
+                from P have "finite P" "P \<subseteq> A"
+                  unfolding S' S by (simp_all add: subset_eq)
+                then have "sup_measure.F id P \<le> x"
+                  by (induction P) (auto simp: x)
+                moreover have "sets (sup_measure.F id P) = sets x"
+                  using \<open>finite P\<close> \<open>P \<noteq> {}\<close> \<open>P \<subseteq> S'\<close> \<open>sets x = sets b\<close>
+                  by (intro sets_sup_measure_F) (auto simp: S')
+                ultimately show "emeasure (sup_measure.F id P) X \<le> emeasure x X"
+                  by (rule le_measureD3)
+              qed
+            qed
+            show "m \<in> S' \<Longrightarrow> sets m = sets (Sup_measure' S')" for m
+              unfolding * by (simp add: S')
+          qed fact
+        qed
+      next
+        assume "sets x \<noteq> sets b"
+        moreover have "sets b \<le> sets x"
+          unfolding b S using x[THEN le_measureD2] \<open>space x = space a\<close> by auto
+        ultimately show "Sup_measure' S' \<le> x"
+          using \<open>space x = space a\<close> \<open>b \<in> S\<close>
+          by (intro less_eq_measure.intros(2)) (simp_all add: * S)
+      qed
+    next
+      assume "space x \<noteq> space a"
+      then have "space a < space x"
+        using le_measureD1[OF x[OF \<open>a\<in>A\<close>]] by auto
+      then show "Sup_measure' S' \<le> x"
+        by (intro less_eq_measure.intros) (simp add: * \<open>space b = space a\<close>)
+    qed
+  qed
+  show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)"
+    by (auto intro!: antisym least simp: top_measure_def)
+  show lower: "x \<in> A \<Longrightarrow> Inf A \<le> x" for x :: "'a measure" and A
+    unfolding Inf_measure_def by (intro least) auto
+  show greatest: "(\<And>z. z \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> x \<le> Inf A" for x :: "'a measure" and A
+    unfolding Inf_measure_def by (intro upper) auto
+  show "inf x y \<le> x" "inf x y \<le> y" "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> inf y z" for x y z :: "'a measure"
+    by (auto simp: inf_measure_def intro!: lower greatest)
+qed
+
+end
+
+lemma sets_SUP:
+  assumes "\<And>x. x \<in> I \<Longrightarrow> sets (M x) = sets N"
+  shows "I \<noteq> {} \<Longrightarrow> sets (SUP i:I. M i) = sets N"
+  unfolding Sup_measure_def
+  using assms assms[THEN sets_eq_imp_space_eq]
+    sets_Sup_measure'[where A=N and M="M`I"]
+  by (intro Sup_lexord1[where P="\<lambda>x. sets x = sets N"]) auto
+
+lemma emeasure_SUP:
+  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N" "I \<noteq> {}"
+  shows "emeasure (SUP i:I. M i) X = (SUP J:{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emeasure (SUP i:J. M i) X)"
+proof -
+  have eq: "finite J \<Longrightarrow> sup_measure.F id J = (SUP i:J. i)" for J :: "'b measure set"
+    by (induction J rule: finite_induct) auto
+  have 1: "J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sets (SUP x:J. M x) = sets N" for J
+    by (intro sets_SUP sets) (auto )
+
+  from \<open>I \<noteq> {}\<close> obtain i where "i\<in>I" by auto
+  have "Sup_measure' (M`I) X = (SUP P:{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X)"
+    using sets by (intro emeasure_Sup_measure') auto
+  also have "Sup_measure' (M`I) = (SUP i:I. M i)"
+    unfolding Sup_measure_def using \<open>I \<noteq> {}\<close> sets sets(1)[THEN sets_eq_imp_space_eq]
+    by (intro Sup_lexord1[where P="\<lambda>x. _ = x"]) auto
+  also have "(SUP P:{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X) =
+    (SUP J:{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. (SUP i:J. M i) X)"
+  proof (intro SUP_eq)
+    fix J assume "J \<in> {P. finite P \<and> P \<subseteq> M`I}"
+    then obtain J' where J': "J' \<subseteq> I" "finite J'" and J: "J = M`J'" and "finite J"
+      using finite_subset_image[of J M I] by auto
+    show "\<exists>j\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. sup_measure.F id J X \<le> (SUP i:j. M i) X"
+    proof cases
+      assume "J' = {}" with \<open>i \<in> I\<close> show ?thesis
+        by (auto simp add: J)
+    next
+      assume "J' \<noteq> {}" with J J' show ?thesis
+        by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply)
+    qed
+  next
+    fix J assume J: "J \<in> {P. P \<noteq> {} \<and> finite P \<and> P \<subseteq> I}"
+    show "\<exists>J'\<in>{J. finite J \<and> J \<subseteq> M`I}. (SUP i:J. M i) X \<le> sup_measure.F id J' X"
+      using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply)
+  qed
+  finally show ?thesis .
+qed
+
+lemma emeasure_SUP_chain:
+  assumes sets: "\<And>i. i \<in> A \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N"
+  assumes ch: "Complete_Partial_Order.chain op \<le> (M ` A)" and "A \<noteq> {}"
+  shows "emeasure (SUP i:A. M i) X = (SUP i:A. emeasure (M i) X)"
+proof (subst emeasure_SUP[OF sets \<open>A \<noteq> {}\<close>])
+  show "(SUP J:{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (SUPREMUM J M) X) = (SUP i:A. emeasure (M i) X)"
+  proof (rule SUP_eq)
+    fix J assume "J \<in> {J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}"
+    then have J: "Complete_Partial_Order.chain op \<le> (M ` J)" "finite J" "J \<noteq> {}" and "J \<subseteq> A"
+      using ch[THEN chain_subset, of "M`J"] by auto
+    with in_chain_finite[OF J(1)] obtain j where "j \<in> J" "(SUP j:J. M j) = M j"
+      by auto
+    with \<open>J \<subseteq> A\<close> show "\<exists>j\<in>A. emeasure (SUPREMUM J M) X \<le> emeasure (M j) X"
+      by auto
+  next
+    fix j assume "j\<in>A" then show "\<exists>i\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (M j) X \<le> emeasure (SUPREMUM i M) X"
+      by (intro bexI[of _ "{j}"]) auto
+  qed
+qed
+
+subsubsection \<open>Supremum of a set of $\sigma$-algebras\<close>
+
+lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)"
+  unfolding Sup_measure_def
+  apply (intro Sup_lexord[where P="\<lambda>x. space x = _"])
+  apply (subst space_Sup_measure'2)
+  apply auto []
+  apply (subst space_measure_of[OF UN_space_closed])
+  apply auto
+  done
+
+lemma sets_Sup_eq:
+  assumes *: "\<And>m. m \<in> M \<Longrightarrow> space m = X" and "M \<noteq> {}"
+  shows "sets (Sup M) = sigma_sets X (\<Union>x\<in>M. sets x)"
+  unfolding Sup_measure_def
+  apply (rule Sup_lexord1)
+  apply fact
+  apply (simp add: assms)
+  apply (rule Sup_lexord)
+  subgoal premises that for a S
+    unfolding that(3) that(2)[symmetric]
+    using that(1)
+    apply (subst sets_Sup_measure'2)
+    apply (intro arg_cong2[where f=sigma_sets])
+    apply (auto simp: *)
+    done
+  apply (subst sets_measure_of[OF UN_space_closed])
+  apply (simp add:  assms)
+  done
+
+lemma in_sets_Sup: "(\<And>m. m \<in> M \<Longrightarrow> space m = X) \<Longrightarrow> m \<in> M \<Longrightarrow> A \<in> sets m \<Longrightarrow> A \<in> sets (Sup M)"
+  by (subst sets_Sup_eq[where X=X]) auto
+
+lemma Sup_lexord_rel:
+  assumes "\<And>i. i \<in> I \<Longrightarrow> k (A i) = k (B i)"
+    "R (c (A ` {a \<in> I. k (B a) = (SUP x:I. k (B x))})) (c (B ` {a \<in> I. k (B a) = (SUP x:I. k (B x))}))"
+    "R (s (A`I)) (s (B`I))"
+  shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))"
+proof -
+  have "A ` {a \<in> I. k (B a) = (SUP x:I. k (B x))} =  {a \<in> A ` I. k a = (SUP x:I. k (B x))}"
+    using assms(1) by auto
+  moreover have "B ` {a \<in> I. k (B a) = (SUP x:I. k (B x))} =  {a \<in> B ` I. k a = (SUP x:I. k (B x))}"
+    by auto
+  ultimately show ?thesis
+    using assms by (auto simp: Sup_lexord_def Let_def)
+qed
+
+lemma sets_SUP_cong:
+  assumes eq: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (SUP i:I. M i) = sets (SUP i:I. N i)"
+  unfolding Sup_measure_def
+  using eq eq[THEN sets_eq_imp_space_eq]
+  apply (intro Sup_lexord_rel[where R="\<lambda>x y. sets x = sets y"])
+  apply simp
+  apply simp
+  apply (simp add: sets_Sup_measure'2)
+  apply (intro arg_cong2[where f="\<lambda>x y. sets (sigma x y)"])
+  apply auto
+  done
+
+lemma sets_Sup_in_sets:
+  assumes "M \<noteq> {}"
+  assumes "\<And>m. m \<in> M \<Longrightarrow> space m = space N"
+  assumes "\<And>m. m \<in> M \<Longrightarrow> sets m \<subseteq> sets N"
+  shows "sets (Sup M) \<subseteq> sets N"
+proof -
+  have *: "UNION M space = space N"
+    using assms by auto
+  show ?thesis
+    unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset)
+qed
+
+lemma measurable_Sup1:
+  assumes m: "m \<in> M" and f: "f \<in> measurable m N"
+    and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n"
+  shows "f \<in> measurable (Sup M) N"
+proof -
+  have "space (Sup M) = space m"
+    using m by (auto simp add: space_Sup_eq_UN dest: const_space)
+  then show ?thesis
+    using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space])
+qed
+
+lemma measurable_Sup2:
+  assumes M: "M \<noteq> {}"
+  assumes f: "\<And>m. m \<in> M \<Longrightarrow> f \<in> measurable N m"
+    and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n"
+  shows "f \<in> measurable N (Sup M)"
+proof -
+  from M obtain m where "m \<in> M" by auto
+  have space_eq: "\<And>n. n \<in> M \<Longrightarrow> space n = space m"
+    by (intro const_space \<open>m \<in> M\<close>)
+  have "f \<in> measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m))"
+  proof (rule measurable_measure_of)
+    show "f \<in> space N \<rightarrow> UNION M space"
+      using measurable_space[OF f] M by auto
+  qed (auto intro: measurable_sets f dest: sets.sets_into_space)
+  also have "measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)) = measurable N (Sup M)"
+    apply (intro measurable_cong_sets refl)
+    apply (subst sets_Sup_eq[OF space_eq M])
+    apply simp
+    apply (subst sets_measure_of[OF UN_space_closed])
+    apply (simp add: space_eq M)
+    done
+  finally show ?thesis .
+qed
+
+lemma sets_Sup_sigma:
+  assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
+  shows "sets (SUP m:M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
+proof -
+  { fix a m assume "a \<in> sigma_sets \<Omega> m" "m \<in> M"
+    then have "a \<in> sigma_sets \<Omega> (\<Union>M)"
+     by induction (auto intro: sigma_sets.intros) }
+  then show "sets (SUP m:M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
+    apply (subst sets_Sup_eq[where X="\<Omega>"])
+    apply (auto simp add: M) []
+    apply auto []
+    apply (simp add: space_measure_of_conv M Union_least)
+    apply (rule sigma_sets_eqI)
+    apply auto
+    done
+qed
+
+lemma Sup_sigma:
+  assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
+  shows "(SUP m:M. sigma \<Omega> m) = (sigma \<Omega> (\<Union>M))"
+proof (intro antisym SUP_least)
+  have *: "\<Union>M \<subseteq> Pow \<Omega>"
+    using M by auto
+  show "sigma \<Omega> (\<Union>M) \<le> (SUP m:M. sigma \<Omega> m)"
+  proof (intro less_eq_measure.intros(3))
+    show "space (sigma \<Omega> (\<Union>M)) = space (SUP m:M. sigma \<Omega> m)"
+      "sets (sigma \<Omega> (\<Union>M)) = sets (SUP m:M. sigma \<Omega> m)"
+      using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq]
+      by auto
+  qed (simp add: emeasure_sigma le_fun_def)
+  fix m assume "m \<in> M" then show "sigma \<Omega> m \<le> sigma \<Omega> (\<Union>M)"
+    by (subst sigma_le_iff) (auto simp add: M *)
+qed
+
+lemma SUP_sigma_sigma:
+  "M \<noteq> {} \<Longrightarrow> (\<And>m. m \<in> M \<Longrightarrow> f m \<subseteq> Pow \<Omega>) \<Longrightarrow> (SUP m:M. sigma \<Omega> (f m)) = sigma \<Omega> (\<Union>m\<in>M. f m)"
+  using Sup_sigma[of "f`M" \<Omega>] by auto
+
+lemma sets_vimage_Sup_eq:
+  assumes *: "M \<noteq> {}" "f \<in> X \<rightarrow> Y" "\<And>m. m \<in> M \<Longrightarrow> space m = Y"
+  shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m : M. vimage_algebra X f m)"
+  (is "?IS = ?SI")
+proof
+  show "?IS \<subseteq> ?SI"
+    apply (intro sets_image_in_sets measurable_Sup2)
+    apply (simp add: space_Sup_eq_UN *)
+    apply (simp add: *)
+    apply (intro measurable_Sup1)
+    apply (rule imageI)
+    apply assumption
+    apply (rule measurable_vimage_algebra1)
+    apply (auto simp: *)
+    done
+  show "?SI \<subseteq> ?IS"
+    apply (intro sets_Sup_in_sets)
+    apply (auto simp: *) []
+    apply (auto simp: *) []
+    apply (elim imageE)
+    apply simp
+    apply (rule sets_image_in_sets)
+    apply simp
+    apply (simp add: measurable_def)
+    apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2)
+    apply (auto intro: in_sets_Sup[OF *(3)])
+    done
+qed
+
+lemma restrict_space_eq_vimage_algebra':
+  "sets (restrict_space M \<Omega>) = sets (vimage_algebra (\<Omega> \<inter> space M) (\<lambda>x. x) M)"
+proof -
+  have *: "{A \<inter> (\<Omega> \<inter> space M) |A. A \<in> sets M} = {A \<inter> \<Omega> |A. A \<in> sets M}"
+    using sets.sets_into_space[of _ M] by blast
+
+  show ?thesis
+    unfolding restrict_space_def
+    by (subst sets_measure_of)
+       (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets])
+qed
+
+lemma sigma_le_sets:
+  assumes [simp]: "A \<subseteq> Pow X" shows "sets (sigma X A) \<subseteq> sets N \<longleftrightarrow> X \<in> sets N \<and> A \<subseteq> sets N"
+proof
+  have "X \<in> sigma_sets X A" "A \<subseteq> sigma_sets X A"
+    by (auto intro: sigma_sets_top)
+  moreover assume "sets (sigma X A) \<subseteq> sets N"
+  ultimately show "X \<in> sets N \<and> A \<subseteq> sets N"
+    by auto
+next
+  assume *: "X \<in> sets N \<and> A \<subseteq> sets N"
+  { fix Y assume "Y \<in> sigma_sets X A" from this * have "Y \<in> sets N"
+      by induction auto }
+  then show "sets (sigma X A) \<subseteq> sets N"
+    by auto
+qed
+
+lemma measurable_iff_sets:
+  "f \<in> measurable M N \<longleftrightarrow> (f \<in> space M \<rightarrow> space N \<and> sets (vimage_algebra (space M) f N) \<subseteq> sets M)"
+proof -
+  have *: "{f -` A \<inter> space M |A. A \<in> sets N} \<subseteq> Pow (space M)"
+    by auto
+  show ?thesis
+    unfolding measurable_def
+    by (auto simp add: vimage_algebra_def sigma_le_sets[OF *])
+qed
+
+lemma sets_vimage_algebra_space: "X \<in> sets (vimage_algebra X f M)"
+  using sets.top[of "vimage_algebra X f M"] by simp
+
+lemma measurable_mono:
+  assumes N: "sets N' \<le> sets N" "space N = space N'"
+  assumes M: "sets M \<le> sets M'" "space M = space M'"
+  shows "measurable M N \<subseteq> measurable M' N'"
+  unfolding measurable_def
+proof safe
+  fix f A assume "f \<in> space M \<rightarrow> space N" "A \<in> sets N'"
+  moreover assume "\<forall>y\<in>sets N. f -` y \<inter> space M \<in> sets M" note this[THEN bspec, of A]
+  ultimately show "f -` A \<inter> space M' \<in> sets M'"
+    using assms by auto
+qed (insert N M, auto)
+
+lemma measurable_Sup_measurable:
+  assumes f: "f \<in> space N \<rightarrow> A"
+  shows "f \<in> measurable N (Sup {M. space M = A \<and> f \<in> measurable N M})"
+proof (rule measurable_Sup2)
+  show "{M. space M = A \<and> f \<in> measurable N M} \<noteq> {}"
+    using f unfolding ex_in_conv[symmetric]
+    by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of)
+qed auto
+
+lemma (in sigma_algebra) sigma_sets_subset':
+  assumes a: "a \<subseteq> M" "\<Omega>' \<in> M"
+  shows "sigma_sets \<Omega>' a \<subseteq> M"
+proof
+  show "x \<in> M" if x: "x \<in> sigma_sets \<Omega>' a" for x
+    using x by (induct rule: sigma_sets.induct) (insert a, auto)
+qed
+
+lemma in_sets_SUP: "i \<in> I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> space (M i) = Y) \<Longrightarrow> X \<in> sets (M i) \<Longrightarrow> X \<in> sets (SUP i:I. M i)"
+  by (intro in_sets_Sup[where X=Y]) auto
+
+lemma measurable_SUP1:
+  "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<And>m n. m \<in> I \<Longrightarrow> n \<in> I \<Longrightarrow> space (M m) = space (M n)) \<Longrightarrow>
+    f \<in> measurable (SUP i:I. M i) N"
+  by (auto intro: measurable_Sup1)
+
+lemma sets_image_in_sets':
+  assumes X: "X \<in> sets N"
+  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> f -` A \<inter> X \<in> sets N"
+  shows "sets (vimage_algebra X f M) \<subseteq> sets N"
+  unfolding sets_vimage_algebra
+  by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f)
+
+lemma mono_vimage_algebra:
+  "sets M \<le> sets N \<Longrightarrow> sets (vimage_algebra X f M) \<subseteq> sets (vimage_algebra X f N)"
+  using sets.top[of "sigma X {f -` A \<inter> X |A. A \<in> sets N}"]
+  unfolding vimage_algebra_def
+  apply (subst (asm) space_measure_of)
+  apply auto []
+  apply (subst sigma_le_sets)
+  apply auto
+  done
+
+lemma mono_restrict_space: "sets M \<le> sets N \<Longrightarrow> sets (restrict_space M X) \<subseteq> sets (restrict_space N X)"
+  unfolding sets_restrict_space by (rule image_mono)
+
+lemma sets_eq_bot: "sets M = {{}} \<longleftrightarrow> M = bot"
+  apply safe
+  apply (intro measure_eqI)
+  apply auto
+  done
+
+lemma sets_eq_bot2: "{{}} = sets M \<longleftrightarrow> M = bot"
+  using sets_eq_bot[of M] by blast
+
+
+lemma (in finite_measure) countable_support:
+  "countable {x. measure M {x} \<noteq> 0}"
+proof cases
+  assume "measure M (space M) = 0"
+  with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
+    by auto
+  then show ?thesis
+    by simp
+next
+  let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
+  assume "?M \<noteq> 0"
+  then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
+    using reals_Archimedean[of "?m x / ?M" for x]
+    by (auto simp: field_simps not_le[symmetric] measure_nonneg divide_le_0_iff measure_le_0_iff)
+  have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
+  proof (rule ccontr)
+    fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
+    then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X"
+      by (metis infinite_arbitrarily_large)
+    from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
+      by auto
+    { fix x assume "x \<in> X"
+      from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
+      then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
+    note singleton_sets = this
+    have "?M < (\<Sum>x\<in>X. ?M / Suc n)"
+      using \<open>?M \<noteq> 0\<close>
+      by (simp add: \<open>card X = Suc (Suc n)\<close> of_nat_Suc field_simps less_le measure_nonneg)
+    also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)"
+      by (rule setsum_mono) fact
+    also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
+      using singleton_sets \<open>finite X\<close>
+      by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
+    finally have "?M < measure M (\<Union>x\<in>X. {x})" .
+    moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
+      using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto
+    ultimately show False by simp
+  qed
+  show ?thesis
+    unfolding * by (intro countable_UN countableI_type countable_finite[OF **])
+qed
+
+end
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -1,10 +1,13 @@
 theory Multivariate_Analysis
 imports
+  Regularity
+  Lebesgue_Integral_Substitution
+  Embed_Measure
+  Complete_Measure
+  Radon_Nikodym
   Fashoda_Theorem
-  Extended_Real_Limits
   Determinants
   Homeomorphism
-  Ordered_Euclidean_Space
   Bounded_Continuous_Function
   Weierstrass_Theorems
   Polytope
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Nonnegative_Lebesgue_Integration.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,2401 @@
+(*  Title:      HOL/Probability/Nonnegative_Lebesgue_Integration.thy
+    Author:     Johannes Hölzl, TU München
+    Author:     Armin Heller, TU München
+*)
+
+section \<open>Lebesgue Integration for Nonnegative Functions\<close>
+
+theory Nonnegative_Lebesgue_Integration
+  imports Measure_Space Borel_Space
+begin
+
+subsection "Simple function"
+
+text \<open>
+
+Our simple functions are not restricted to nonnegative real numbers. Instead
+they are just functions with a finite range and are measurable when singleton
+sets are measurable.
+
+\<close>
+
+definition "simple_function M g \<longleftrightarrow>
+    finite (g ` space M) \<and>
+    (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
+
+lemma simple_functionD:
+  assumes "simple_function M g"
+  shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M"
+proof -
+  show "finite (g ` space M)"
+    using assms unfolding simple_function_def by auto
+  have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto
+  also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
+  finally show "g -` X \<inter> space M \<in> sets M" using assms
+    by (auto simp del: UN_simps simp: simple_function_def)
+qed
+
+lemma measurable_simple_function[measurable_dest]:
+  "simple_function M f \<Longrightarrow> f \<in> measurable M (count_space UNIV)"
+  unfolding simple_function_def measurable_def
+proof safe
+  fix A assume "finite (f ` space M)" "\<forall>x\<in>f ` space M. f -` {x} \<inter> space M \<in> sets M"
+  then have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) \<in> sets M"
+    by (intro sets.finite_UN) auto
+  also have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) = f -` A \<inter> space M"
+    by (auto split: if_split_asm)
+  finally show "f -` A \<inter> space M \<in> sets M" .
+qed simp
+
+lemma borel_measurable_simple_function:
+  "simple_function M f \<Longrightarrow> f \<in> borel_measurable M"
+  by (auto dest!: measurable_simple_function simp: measurable_def)
+
+lemma simple_function_measurable2[intro]:
+  assumes "simple_function M f" "simple_function M g"
+  shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
+proof -
+  have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
+    by auto
+  then show ?thesis using assms[THEN simple_functionD(2)] by auto
+qed
+
+lemma simple_function_indicator_representation:
+  fixes f ::"'a \<Rightarrow> ennreal"
+  assumes f: "simple_function M f" and x: "x \<in> space M"
+  shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
+  (is "?l = ?r")
+proof -
+  have "?r = (\<Sum>y \<in> f ` space M.
+    (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
+    by (auto intro!: setsum.cong)
+  also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
+    using assms by (auto dest: simple_functionD simp: setsum.delta)
+  also have "... = f x" using x by (auto simp: indicator_def)
+  finally show ?thesis by auto
+qed
+
+lemma simple_function_notspace:
+  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h")
+proof -
+  have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
+  hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
+  have "?h -` {0} \<inter> space M = space M" by auto
+  thus ?thesis unfolding simple_function_def by auto
+qed
+
+lemma simple_function_cong:
+  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
+  shows "simple_function M f \<longleftrightarrow> simple_function M g"
+proof -
+  have "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
+    using assms by auto
+  with assms show ?thesis
+    by (simp add: simple_function_def cong: image_cong)
+qed
+
+lemma simple_function_cong_algebra:
+  assumes "sets N = sets M" "space N = space M"
+  shows "simple_function M f \<longleftrightarrow> simple_function N f"
+  unfolding simple_function_def assms ..
+
+lemma simple_function_borel_measurable:
+  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
+  assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
+  shows "simple_function M f"
+  using assms unfolding simple_function_def
+  by (auto intro: borel_measurable_vimage)
+
+lemma simple_function_iff_borel_measurable:
+  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
+  shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M"
+  by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable)
+
+lemma simple_function_eq_measurable:
+  "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)"
+  using measurable_simple_function[of M f] by (fastforce simp: simple_function_def)
+
+lemma simple_function_const[intro, simp]:
+  "simple_function M (\<lambda>x. c)"
+  by (auto intro: finite_subset simp: simple_function_def)
+lemma simple_function_compose[intro, simp]:
+  assumes "simple_function M f"
+  shows "simple_function M (g \<circ> f)"
+  unfolding simple_function_def
+proof safe
+  show "finite ((g \<circ> f) ` space M)"
+    using assms unfolding simple_function_def by (auto simp: image_comp [symmetric])
+next
+  fix x assume "x \<in> space M"
+  let ?G = "g -` {g (f x)} \<inter> (f`space M)"
+  have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
+    (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
+  show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
+    using assms unfolding simple_function_def *
+    by (rule_tac sets.finite_UN) auto
+qed
+
+lemma simple_function_indicator[intro, simp]:
+  assumes "A \<in> sets M"
+  shows "simple_function M (indicator A)"
+proof -
+  have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
+    by (auto simp: indicator_def)
+  hence "finite ?S" by (rule finite_subset) simp
+  moreover have "- A \<inter> space M = space M - A" by auto
+  ultimately show ?thesis unfolding simple_function_def
+    using assms by (auto simp: indicator_def [abs_def])
+qed
+
+lemma simple_function_Pair[intro, simp]:
+  assumes "simple_function M f"
+  assumes "simple_function M g"
+  shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p")
+  unfolding simple_function_def
+proof safe
+  show "finite (?p ` space M)"
+    using assms unfolding simple_function_def
+    by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
+next
+  fix x assume "x \<in> space M"
+  have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
+      (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
+    by auto
+  with \<open>x \<in> space M\<close> show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
+    using assms unfolding simple_function_def by auto
+qed
+
+lemma simple_function_compose1:
+  assumes "simple_function M f"
+  shows "simple_function M (\<lambda>x. g (f x))"
+  using simple_function_compose[OF assms, of g]
+  by (simp add: comp_def)
+
+lemma simple_function_compose2:
+  assumes "simple_function M f" and "simple_function M g"
+  shows "simple_function M (\<lambda>x. h (f x) (g x))"
+proof -
+  have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
+    using assms by auto
+  thus ?thesis by (simp_all add: comp_def)
+qed
+
+lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"]
+  and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
+  and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
+  and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
+  and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
+  and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
+  and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
+
+lemma simple_function_setsum[intro, simp]:
+  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
+  shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
+proof cases
+  assume "finite P" from this assms show ?thesis by induct auto
+qed auto
+
+lemma simple_function_ennreal[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
+  shows "simple_function M (\<lambda>x. ennreal (f x))"
+  by (rule simple_function_compose1[OF sf])
+
+lemma simple_function_real_of_nat[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
+  shows "simple_function M (\<lambda>x. real (f x))"
+  by (rule simple_function_compose1[OF sf])
+
+lemma borel_measurable_implies_simple_function_sequence:
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u[measurable]: "u \<in> borel_measurable M"
+  shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)"
+proof -
+  define f where [abs_def]:
+    "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x
+
+  have [simp]: "0 \<le> f i x" for i x
+    by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg)
+
+  have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
+    by simp
+
+  have "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = real_of_int \<lfloor>i * 2 ^ i\<rfloor>" for i
+    by (intro arg_cong[where f=real_of_int]) simp
+  then have [simp]: "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = i * 2 ^ i" for i
+    unfolding floor_of_nat by simp
+
+  have "incseq f"
+  proof (intro monoI le_funI)
+    fix m n :: nat and x assume "m \<le> n"
+    moreover
+    { fix d :: nat
+      have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * enn2real (min (of_nat m) (u x))\<rfloor> \<le>
+        \<lfloor>2^d * (2^m * enn2real (min (of_nat m) (u x)))\<rfloor>"
+        by (rule le_mult_floor) (auto simp: enn2real_nonneg)
+      also have "\<dots> \<le> \<lfloor>2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\<rfloor>"
+        by (intro floor_mono mult_mono enn2real_mono min.mono)
+           (auto simp: enn2real_nonneg min_less_iff_disj of_nat_less_top)
+      finally have "f m x \<le> f (m + d) x"
+        unfolding f_def
+        by (auto simp: field_simps power_add * simp del: of_int_mult) }
+    ultimately show "f m x \<le> f n x"
+      by (auto simp add: le_iff_add)
+  qed
+  then have inc_f: "incseq (\<lambda>i. ennreal (f i x))" for x
+    by (auto simp: incseq_def le_fun_def)
+  then have "incseq (\<lambda>i x. ennreal (f i x))"
+    by (auto simp: incseq_def le_fun_def)
+  moreover
+  have "simple_function M (f i)" for i
+  proof (rule simple_function_borel_measurable)
+    have "\<lfloor>enn2real (min (of_nat i) (u x)) * 2 ^ i\<rfloor> \<le> \<lfloor>int i * 2 ^ i\<rfloor>" for x
+      by (cases "u x" rule: ennreal_cases)
+         (auto split: split_min intro!: floor_mono)
+    then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
+      unfolding floor_of_int by (auto simp: f_def enn2real_nonneg intro!: imageI)
+    then show "finite (f i ` space M)"
+      by (rule finite_subset) auto
+    show "f i \<in> borel_measurable M"
+      unfolding f_def enn2real_def by measurable
+  qed
+  moreover
+  { fix x
+    have "(SUP i. ennreal (f i x)) = u x"
+    proof (cases "u x" rule: ennreal_cases)
+      case top then show ?thesis
+        by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric]
+                      ennreal_SUP_of_nat_eq_top)
+    next
+      case (real r)
+      obtain n where "r \<le> of_nat n" using real_arch_simple by auto
+      then have min_eq_r: "\<forall>\<^sub>F x in sequentially. min (real x) r = r"
+        by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)
+
+      have "(\<lambda>i. real_of_int \<lfloor>min (real i) r * 2^i\<rfloor> / 2^i) \<longlonglongrightarrow> r"
+      proof (rule tendsto_sandwich)
+        show "(\<lambda>n. r - (1/2)^n) \<longlonglongrightarrow> r"
+          by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
+        show "\<forall>\<^sub>F n in sequentially. real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n \<le> r"
+          using min_eq_r by eventually_elim (auto simp: field_simps)
+        have *: "r * (2 ^ n * 2 ^ n) \<le> 2^n + 2^n * real_of_int \<lfloor>r * 2 ^ n\<rfloor>" for n
+          using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"]
+          by (auto simp: field_simps)
+        show "\<forall>\<^sub>F n in sequentially. r - (1/2)^n \<le> real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n"
+          using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
+      qed auto
+      then have "(\<lambda>i. ennreal (f i x)) \<longlonglongrightarrow> ennreal r"
+        by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal)
+      from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this]
+      show ?thesis
+        by (simp add: real)
+    qed }
+  ultimately show ?thesis
+    by (intro exI[of _ "\<lambda>i x. ennreal (f i x)"]) auto
+qed
+
+lemma borel_measurable_implies_simple_function_sequence':
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u: "u \<in> borel_measurable M"
+  obtains f where
+    "\<And>i. simple_function M (f i)" "incseq f" "\<And>i x. f i x < top" "\<And>x. (SUP i. f i x) = u x"
+  using borel_measurable_implies_simple_function_sequence[OF u] by (auto simp: fun_eq_iff) blast
+
+lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]:
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u: "simple_function M u"
+  assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
+  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
+  assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+  assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
+  shows "P u"
+proof (rule cong)
+  from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
+  proof eventually_elim
+    fix x assume x: "x \<in> space M"
+    from simple_function_indicator_representation[OF u x]
+    show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
+  qed
+next
+  from u have "finite (u ` space M)"
+    unfolding simple_function_def by auto
+  then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
+  proof induct
+    case empty show ?case
+      using set[of "{}"] by (simp add: indicator_def[abs_def])
+  qed (auto intro!: add mult set simple_functionD u)
+next
+  show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
+    apply (subst simple_function_cong)
+    apply (rule simple_function_indicator_representation[symmetric])
+    apply (auto intro: u)
+    done
+qed fact
+
+lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]:
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u: "simple_function M u"
+  assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
+  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
+  assumes mult: "\<And>u c. simple_function M u \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+  assumes add: "\<And>u v. simple_function M u \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<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)"
+  shows "P u"
+proof -
+  show ?thesis
+  proof (rule cong)
+    fix x assume x: "x \<in> space M"
+    from simple_function_indicator_representation[OF u x]
+    show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
+  next
+    show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
+      apply (subst simple_function_cong)
+      apply (rule simple_function_indicator_representation[symmetric])
+      apply (auto intro: u)
+      done
+  next
+    from u have "finite (u ` space M)"
+      unfolding simple_function_def by auto
+    then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
+    proof induct
+      case empty show ?case
+        using set[of "{}"] by (simp add: indicator_def[abs_def])
+    next
+      case (insert x S)
+      { fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
+          x * indicator (u -` {x} \<inter> space M) z = 0"
+          using insert by (subst setsum_eq_0_iff) (auto simp: indicator_def) }
+      note disj = this
+      from insert show ?case
+        by (auto intro!: add mult set simple_functionD u simple_function_setsum disj)
+    qed
+  qed fact
+qed
+
+lemma borel_measurable_induct[consumes 1, case_names cong set mult add seq, induct set: borel_measurable]:
+  fixes u :: "'a \<Rightarrow> ennreal"
+  assumes u: "u \<in> borel_measurable M"
+  assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f"
+  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
+  assumes mult': "\<And>u c. c < top \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+  assumes add: "\<And>u v. u \<in> borel_measurable M\<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < top) \<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. x \<in> space M \<Longrightarrow> U i x < top) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)"
+  shows "P u"
+  using u
+proof (induct rule: borel_measurable_implies_simple_function_sequence')
+  fix U assume 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) = u x"
+  have u_eq: "u = (SUP i. U i)"
+    using u sup by auto
+
+  have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < top"
+    using U by (auto simp: image_iff eq_commute)
+
+  from U have "\<And>i. U i \<in> borel_measurable M"
+    by (simp add: borel_measurable_simple_function)
+
+  show "P u"
+    unfolding u_eq
+  proof (rule seq)
+    fix i show "P (U i)"
+      using \<open>simple_function M (U i)\<close> not_inf[of _ i]
+    proof (induct rule: simple_function_induct_nn)
+      case (mult u c)
+      show ?case
+      proof cases
+        assume "c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0)"
+        with mult(1) show ?thesis
+          by (intro cong[of "\<lambda>x. c * u x" "indicator {}"] set)
+             (auto dest!: borel_measurable_simple_function)
+      next
+        assume "\<not> (c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0))"
+        then obtain x where "space M \<noteq> {}" and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0"
+          by auto
+        with mult(3)[of x] have "c < top"
+          by (auto simp: ennreal_mult_less_top)
+        then have u_fin: "x' \<in> space M \<Longrightarrow> u x' < top" for x'
+          using mult(3)[of x'] \<open>c \<noteq> 0\<close> by (auto simp: ennreal_mult_less_top)
+        then have "P u"
+          by (rule mult)
+        with u_fin \<open>c < top\<close> mult(1) show ?thesis
+          by (intro mult') (auto dest!: borel_measurable_simple_function)
+      qed
+    qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function)
+  qed fact+
+qed
+
+lemma simple_function_If_set:
+  assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
+  shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
+proof -
+  define F where "F x = f -` {x} \<inter> space M" for x
+  define G where "G x = g -` {x} \<inter> space M" for x
+  show ?thesis unfolding simple_function_def
+  proof safe
+    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
+    from finite_subset[OF this] assms
+    show "finite (?IF ` space M)" unfolding simple_function_def by auto
+  next
+    fix x assume "x \<in> space M"
+    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
+      then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M))))
+      else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))"
+      using sets.sets_into_space[OF A] by (auto split: if_split_asm simp: G_def F_def)
+    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
+      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
+    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
+  qed
+qed
+
+lemma simple_function_If:
+  assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
+  shows "simple_function M (\<lambda>x. if P x then f x else g x)"
+proof -
+  have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
+  with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
+qed
+
+lemma simple_function_subalgebra:
+  assumes "simple_function N f"
+  and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M"
+  shows "simple_function M f"
+  using assms unfolding simple_function_def by auto
+
+lemma simple_function_comp:
+  assumes T: "T \<in> measurable M M'"
+    and f: "simple_function M' f"
+  shows "simple_function M (\<lambda>x. f (T x))"
+proof (intro simple_function_def[THEN iffD2] conjI ballI)
+  have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
+    using T unfolding measurable_def by auto
+  then show "finite ((\<lambda>x. f (T x)) ` space M)"
+    using f unfolding simple_function_def by (auto intro: finite_subset)
+  fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M"
+  then have "i \<in> f ` space M'"
+    using T unfolding measurable_def by auto
+  then have "f -` {i} \<inter> space M' \<in> sets M'"
+    using f unfolding simple_function_def by auto
+  then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
+    using T unfolding measurable_def by auto
+  also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
+    using T unfolding measurable_def by auto
+  finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
+qed
+
+subsection "Simple integral"
+
+definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
+  "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
+
+syntax
+  "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
+
+translations
+  "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
+
+lemma simple_integral_cong:
+  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
+  shows "integral\<^sup>S M f = integral\<^sup>S M g"
+proof -
+  have "f ` space M = g ` space M"
+    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
+    using assms by (auto intro!: image_eqI)
+  thus ?thesis unfolding simple_integral_def by simp
+qed
+
+lemma simple_integral_const[simp]:
+  "(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)"
+proof (cases "space M = {}")
+  case True thus ?thesis unfolding simple_integral_def by simp
+next
+  case False hence "(\<lambda>x. c) ` space M = {c}" by auto
+  thus ?thesis unfolding simple_integral_def by simp
+qed
+
+lemma simple_function_partition:
+  assumes f: "simple_function M f" and g: "simple_function M g"
+  assumes sub: "\<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 "integral\<^sup>S M f = (\<Sum>y\<in>g ` space M. v y * emeasure M {x\<in>space M. g x = y})"
+    (is "_ = ?r")
+proof -
+  from f g have [simp]: "finite (f`space M)" "finite (g`space M)"
+    by (auto simp: simple_function_def)
+  from f g have [measurable]: "f \<in> measurable M (count_space UNIV)" "g \<in> measurable M (count_space UNIV)"
+    by (auto intro: measurable_simple_function)
+
+  { 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 "integral\<^sup>S M f =
+    (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M.
+      if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
+    unfolding simple_integral_def
+  proof (safe intro!: setsum.cong ennreal_mult_left_cong)
+    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:"(\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i}) =
+        f -` {f y} \<inter> space M"
+      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
+    then show "emeasure M (f -` {f y} \<inter> space M) =
+      (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then emeasure M {x \<in> space M. g x = z} else 0)"
+      apply (simp add: setsum.If_cases)
+      apply (subst setsum_emeasure)
+      apply (auto simp: disjoint_family_on_def eq)
+      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 y * emeasure M {x\<in>space M. g x = z} else 0))"
+    by (auto intro!: setsum.cong simp: setsum_right_distrib)
+  also have "\<dots> = ?r"
+    by (subst setsum.commute)
+       (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
+  finally show "integral\<^sup>S M f = ?r" .
+qed
+
+lemma simple_integral_add[simp]:
+  assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
+  shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g"
+proof -
+  have "(\<integral>\<^sup>Sx. f x + g x \<partial>M) =
+    (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\<in>space M. (f x, g x) = y})"
+    by (intro simple_function_partition) (auto intro: f g)
+  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
+    (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
+    using assms(2,4) by (auto intro!: setsum.cong distrib_right simp: setsum.distrib[symmetric])
+  also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
+    by (intro simple_function_partition[symmetric]) (auto intro: f g)
+  also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
+    by (intro simple_function_partition[symmetric]) (auto intro: f g)
+  finally show ?thesis .
+qed
+
+lemma simple_integral_setsum[simp]:
+  assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
+  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
+  shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))"
+proof cases
+  assume "finite P"
+  from this assms show ?thesis
+    by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
+qed auto
+
+lemma simple_integral_mult[simp]:
+  assumes f: "simple_function M f"
+  shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f"
+proof -
+  have "(\<integral>\<^sup>Sx. c * f x \<partial>M) = (\<Sum>y\<in>f ` space M. (c * y) * emeasure M {x\<in>space M. f x = y})"
+    using f by (intro simple_function_partition) auto
+  also have "\<dots> = c * integral\<^sup>S M f"
+    using f unfolding simple_integral_def
+    by (subst setsum_right_distrib) (auto simp: mult.assoc Int_def conj_commute)
+  finally show ?thesis .
+qed
+
+lemma simple_integral_mono_AE:
+  assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g"
+  and mono: "AE x in M. f x \<le> g x"
+  shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
+proof -
+  let ?\<mu> = "\<lambda>P. emeasure M {x\<in>space M. P x}"
+  have "integral\<^sup>S M f = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
+    using f g by (intro simple_function_partition) auto
+  also have "\<dots> \<le> (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
+  proof (clarsimp intro!: setsum_mono)
+    fix x assume "x \<in> space M"
+    let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)"
+    show "f x * ?M \<le> g x * ?M"
+    proof cases
+      assume "?M \<noteq> 0"
+      then have "0 < ?M"
+        by (simp add: less_le)
+      also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)"
+        using mono by (intro emeasure_mono_AE) auto
+      finally have "\<not> \<not> f x \<le> g x"
+        by (intro notI) auto
+      then show ?thesis
+        by (intro mult_right_mono) auto
+    qed simp
+  qed
+  also have "\<dots> = integral\<^sup>S M g"
+    using f g by (intro simple_function_partition[symmetric]) auto
+  finally show ?thesis .
+qed
+
+lemma simple_integral_mono:
+  assumes "simple_function M f" and "simple_function M g"
+  and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
+  shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
+  using assms by (intro simple_integral_mono_AE) auto
+
+lemma simple_integral_cong_AE:
+  assumes "simple_function M f" and "simple_function M g"
+  and "AE x in M. f x = g x"
+  shows "integral\<^sup>S M f = integral\<^sup>S M g"
+  using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
+
+lemma simple_integral_cong':
+  assumes sf: "simple_function M f" "simple_function M g"
+  and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
+  shows "integral\<^sup>S M f = integral\<^sup>S M g"
+proof (intro simple_integral_cong_AE sf AE_I)
+  show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
+  show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
+    using sf[THEN borel_measurable_simple_function] by auto
+qed simp
+
+lemma simple_integral_indicator:
+  assumes A: "A \<in> sets M"
+  assumes f: "simple_function M f"
+  shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
+    (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
+proof -
+  have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ennreal))`A"
+    using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm)
+  have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
+    by (auto simp: image_iff)
+
+  have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
+    (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\<in>space M. (f x, indicator A x) = y})"
+    using assms by (intro simple_function_partition) auto
+  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ennreal))`space M.
+    if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
+    by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong)
+  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
+    using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
+  also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
+    by (subst setsum.reindex [of fst]) (auto simp: inj_on_def)
+  also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
+    using A[THEN sets.sets_into_space]
+    by (intro setsum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
+  finally show ?thesis .
+qed
+
+lemma simple_integral_indicator_only[simp]:
+  assumes "A \<in> sets M"
+  shows "integral\<^sup>S M (indicator A) = emeasure M A"
+  using simple_integral_indicator[OF assms, of "\<lambda>x. 1"] sets.sets_into_space[OF assms]
+  by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm)
+
+lemma simple_integral_null_set:
+  assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M"
+  shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0"
+proof -
+  have "AE x in M. indicator N x = (0 :: ennreal)"
+    using \<open>N \<in> null_sets M\<close> by (auto simp: indicator_def intro!: AE_I[of _ _ N])
+  then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)"
+    using assms apply (intro simple_integral_cong_AE) by auto
+  then show ?thesis by simp
+qed
+
+lemma simple_integral_cong_AE_mult_indicator:
+  assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M"
+  shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)"
+  using assms by (intro simple_integral_cong_AE) auto
+
+lemma simple_integral_cmult_indicator:
+  assumes A: "A \<in> sets M"
+  shows "(\<integral>\<^sup>Sx. c * indicator A x \<partial>M) = c * emeasure M A"
+  using simple_integral_mult[OF simple_function_indicator[OF A]]
+  unfolding simple_integral_indicator_only[OF A] by simp
+
+lemma simple_integral_nonneg:
+  assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x"
+  shows "0 \<le> integral\<^sup>S M f"
+proof -
+  have "integral\<^sup>S M (\<lambda>x. 0) \<le> integral\<^sup>S M f"
+    using simple_integral_mono_AE[OF _ f ae] by auto
+  then show ?thesis by simp
+qed
+
+subsection \<open>Integral on nonnegative functions\<close>
+
+definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
+  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
+
+syntax
+  "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
+
+translations
+  "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)"
+
+lemma nn_integral_def_finite:
+  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> f \<and> (\<forall>x. g x < top)}. integral\<^sup>S M g)"
+    (is "_ = SUPREMUM ?A ?f")
+  unfolding nn_integral_def
+proof (safe intro!: antisym SUP_least)
+  fix g assume g[measurable]: "simple_function M g" "g \<le> f"
+
+  show "integral\<^sup>S M g \<le> SUPREMUM ?A ?f"
+  proof cases
+    assume ae: "AE x in M. g x \<noteq> top"
+    let ?G = "{x \<in> space M. g x \<noteq> top}"
+    have "integral\<^sup>S M g = integral\<^sup>S M (\<lambda>x. g x * indicator ?G x)"
+    proof (rule simple_integral_cong_AE)
+      show "AE x in M. g x = g x * indicator ?G x"
+        using ae AE_space by eventually_elim auto
+    qed (insert g, auto)
+    also have "\<dots> \<le> SUPREMUM ?A ?f"
+      using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator)
+    finally show ?thesis .
+  next
+    assume nAE: "\<not> (AE x in M. g x \<noteq> top)"
+    then have "emeasure M {x\<in>space M. g x = top} \<noteq> 0" (is "emeasure M ?G \<noteq> 0")
+      by (subst (asm) AE_iff_measurable[OF _ refl]) auto
+    then have "top = (SUP n. (\<integral>\<^sup>Sx. of_nat n * indicator ?G x \<partial>M))"
+      by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric])
+    also have "\<dots> \<le> SUPREMUM ?A ?f"
+      using g
+      by (safe intro!: SUP_least SUP_upper)
+         (auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator
+               intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]])
+    finally show ?thesis
+      by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff)
+  qed
+qed (auto intro: SUP_upper)
+
+lemma nn_integral_mono_AE:
+  assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v"
+  unfolding nn_integral_def
+proof (safe intro!: SUP_mono)
+  fix n assume n: "simple_function M n" "n \<le> u"
+  from ae[THEN AE_E] guess N . note N = this
+  then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in)
+  let ?n = "\<lambda>x. n x * indicator (space M - N) x"
+  have "AE x in M. n x \<le> ?n x" "simple_function M ?n"
+    using n N ae_N by auto
+  moreover
+  { fix x have "?n x \<le> v x"
+    proof cases
+      assume x: "x \<in> space M - N"
+      with N have "u x \<le> v x" by auto
+      with n(2)[THEN le_funD, of x] x show ?thesis
+        by (auto simp: max_def split: if_split_asm)
+    qed simp }
+  then have "?n \<le> v" by (auto simp: le_funI)
+  moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n"
+    using ae_N N n by (auto intro!: simple_integral_mono_AE)
+  ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
+    by force
+qed
+
+lemma nn_integral_mono:
+  "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>N M u \<le> integral\<^sup>N M v"
+  by (auto intro: nn_integral_mono_AE)
+
+lemma mono_nn_integral: "mono F \<Longrightarrow> mono (\<lambda>x. integral\<^sup>N M (F x))"
+  by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono)
+
+lemma nn_integral_cong_AE:
+  "AE x in M. u x = v x \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
+  by (auto simp: eq_iff intro!: nn_integral_mono_AE)
+
+lemma nn_integral_cong:
+  "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
+  by (auto intro: nn_integral_cong_AE)
+
+lemma nn_integral_cong_simp:
+  "(\<And>x. x \<in> space M =simp=> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
+  by (auto intro: nn_integral_cong simp: simp_implies_def)
+
+lemma nn_integral_cong_strong:
+  "M = N \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N N v"
+  by (auto intro: nn_integral_cong)
+
+lemma incseq_nn_integral:
+  assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))"
+proof -
+  have "\<And>i x. f i x \<le> f (Suc i) x"
+    using assms by (auto dest!: incseq_SucD simp: le_fun_def)
+  then show ?thesis
+    by (auto intro!: incseq_SucI nn_integral_mono)
+qed
+
+lemma nn_integral_eq_simple_integral:
+  assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f"
+proof -
+  let ?f = "\<lambda>x. f x * indicator (space M) x"
+  have f': "simple_function M ?f" using f by auto
+  have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f'
+    by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def)
+  moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f"
+    unfolding nn_integral_def
+    using f' by (auto intro!: SUP_upper)
+  ultimately show ?thesis
+    by (simp cong: nn_integral_cong simple_integral_cong)
+qed
+
+text \<open>Beppo-Levi monotone convergence theorem\<close>
+lemma nn_integral_monotone_convergence_SUP:
+  assumes f: "incseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
+proof (rule antisym)
+  show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
+    unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"]
+  proof (safe intro!: SUP_least)
+    fix u assume sf_u[simp]: "simple_function M u" and
+      u: "u \<le> (\<lambda>x. SUP i. f i x)" and u_range: "\<forall>x. u x < top"
+    note sf_u[THEN borel_measurable_simple_function, measurable]
+    show "integral\<^sup>S M u \<le> (SUP j. \<integral>\<^sup>+x. f j x \<partial>M)"
+    proof (rule ennreal_approx_unit)
+      fix a :: ennreal assume "a < 1"
+      let ?au = "\<lambda>x. a * u x"
+
+      let ?B = "\<lambda>c i. {x\<in>space M. ?au x = c \<and> c \<le> f i x}"
+      have "integral\<^sup>S M ?au = (\<Sum>c\<in>?au`space M. c * (SUP i. emeasure M (?B c i)))"
+        unfolding simple_integral_def
+      proof (intro setsum.cong ennreal_mult_left_cong refl)
+        fix c assume "c \<in> ?au ` space M" "c \<noteq> 0"
+        { fix x' assume x': "x' \<in> space M" "?au x' = c"
+          with \<open>c \<noteq> 0\<close> u_range have "?au x' < 1 * u x'"
+            by (intro ennreal_mult_strict_right_mono \<open>a < 1\<close>) (auto simp: less_le)
+          also have "\<dots> \<le> (SUP i. f i x')"
+            using u by (auto simp: le_fun_def)
+          finally have "\<exists>i. ?au x' \<le> f i x'"
+            by (auto simp: less_SUP_iff intro: less_imp_le) }
+        then have *: "?au -` {c} \<inter> space M = (\<Union>i. ?B c i)"
+          by auto
+        show "emeasure M (?au -` {c} \<inter> space M) = (SUP i. emeasure M (?B c i))"
+          unfolding * using f
+          by (intro SUP_emeasure_incseq[symmetric])
+             (auto simp: incseq_def le_fun_def intro: order_trans)
+      qed
+      also have "\<dots> = (SUP i. \<Sum>c\<in>?au`space M. c * emeasure M (?B c i))"
+        unfolding SUP_mult_left_ennreal using f
+        by (intro ennreal_SUP_setsum[symmetric])
+           (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans)
+      also have "\<dots> \<le> (SUP i. integral\<^sup>N M (f i))"
+      proof (intro SUP_subset_mono order_refl)
+        fix i
+        have "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) =
+          (\<integral>\<^sup>Sx. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
+          by (subst simple_integral_indicator)
+             (auto intro!: setsum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
+        also have "\<dots> = (\<integral>\<^sup>+x. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
+          by (rule nn_integral_eq_simple_integral[symmetric]) simp
+        also have "\<dots> \<le> (\<integral>\<^sup>+x. f i x \<partial>M)"
+          by (intro nn_integral_mono) (auto split: split_indicator)
+        finally show "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) \<le> (\<integral>\<^sup>+x. f i x \<partial>M)" .
+      qed
+      finally show "a * integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))"
+        by simp
+    qed
+  qed
+qed (auto intro!: SUP_least SUP_upper nn_integral_mono)
+
+lemma sup_continuous_nn_integral[order_continuous_intros]:
+  assumes f: "\<And>y. sup_continuous (f y)"
+  assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
+  shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
+  unfolding sup_continuous_def
+proof safe
+  fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C"
+  with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (SUPREMUM UNIV C) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
+    unfolding sup_continuousD[OF f C]
+    by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def)
+qed
+
+lemma nn_integral_monotone_convergence_SUP_AE:
+  assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x" "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
+proof -
+  from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x"
+    by (simp add: AE_all_countable)
+  from this[THEN AE_E] guess N . note N = this
+  let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0"
+  have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N])
+  then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)"
+    by (auto intro!: nn_integral_cong_AE)
+  also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))"
+  proof (rule nn_integral_monotone_convergence_SUP)
+    show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
+    { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
+        using f N(3) by (intro measurable_If_set) auto }
+  qed
+  also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
+    using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] nn_integral_cong_AE ext)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_monotone_convergence_simple:
+  "incseq f \<Longrightarrow> (\<And>i. simple_function M (f i)) \<Longrightarrow> (SUP i. \<integral>\<^sup>S x. f i x \<partial>M) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
+  using nn_integral_monotone_convergence_SUP[of f M]
+  by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function)
+
+lemma SUP_simple_integral_sequences:
+  assumes f: "incseq f" "\<And>i. simple_function M (f i)"
+  and g: "incseq g" "\<And>i. simple_function M (g i)"
+  and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)"
+  shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))"
+    (is "SUPREMUM _ ?F = SUPREMUM _ ?G")
+proof -
+  have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
+    using f by (rule nn_integral_monotone_convergence_simple)
+  also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)"
+    unfolding eq[THEN nn_integral_cong_AE] ..
+  also have "\<dots> = (SUP i. ?G i)"
+    using g by (rule nn_integral_monotone_convergence_simple[symmetric])
+  finally show ?thesis by simp
+qed
+
+lemma nn_integral_const[simp]: "(\<integral>\<^sup>+ x. c \<partial>M) = c * emeasure M (space M)"
+  by (subst nn_integral_eq_simple_integral) auto
+
+lemma nn_integral_linear:
+  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g"
+    (is "integral\<^sup>N M ?L = _")
+proof -
+  from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
+  note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this
+  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
+  note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this
+  let ?L' = "\<lambda>i x. a * u i x + v i x"
+
+  have "?L \<in> borel_measurable M" using assms by auto
+  from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
+  note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this
+
+  have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))"
+    using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono)
+
+  have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))"
+  proof (rule SUP_simple_integral_sequences[OF l(3,2)])
+    show "incseq ?L'" "\<And>i. simple_function M (?L' i)"
+      using u v unfolding incseq_Suc_iff le_fun_def
+      by (auto intro!: add_mono mult_left_mono)
+    { fix x
+      have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
+        using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal
+        by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) }
+    then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)"
+      unfolding l(5) using u(5) v(5) by (intro AE_I2) auto
+  qed
+  also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))"
+    using u(2) v(2) by auto
+  finally show ?thesis
+    unfolding l(5)[symmetric] l(1)[symmetric]
+    by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric])
+qed
+
+lemma nn_integral_cmult: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f"
+  using nn_integral_linear[of f M "\<lambda>x. 0" c] by simp
+
+lemma nn_integral_multc: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c"
+  unfolding mult.commute[of _ c] nn_integral_cmult by simp
+
+lemma nn_integral_divide: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c"
+   unfolding divide_ennreal_def by (rule nn_integral_multc)
+
+lemma nn_integral_indicator[simp]: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A"
+  by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator)
+
+lemma nn_integral_cmult_indicator: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * emeasure M A"
+  by (subst nn_integral_eq_simple_integral)
+     (auto simp: simple_function_indicator simple_integral_indicator)
+
+lemma nn_integral_indicator':
+  assumes [measurable]: "A \<inter> space M \<in> sets M"
+  shows "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)"
+proof -
+  have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)"
+    by (intro nn_integral_cong) (simp split: split_indicator)
+  also have "\<dots> = emeasure M (A \<inter> space M)"
+    by simp
+  finally show ?thesis .
+qed
+
+lemma nn_integral_indicator_singleton[simp]:
+  assumes [measurable]: "{y} \<in> sets M" shows "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = f y * emeasure M {y}"
+proof -
+  have "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = (\<integral>\<^sup>+x. f y * indicator {y} x \<partial>M)"
+    by (auto intro!: nn_integral_cong split: split_indicator)
+  then show ?thesis
+    by (simp add: nn_integral_cmult)
+qed
+
+lemma nn_integral_set_ennreal:
+  "(\<integral>\<^sup>+x. ennreal (f x) * indicator A x \<partial>M) = (\<integral>\<^sup>+x. ennreal (f x * indicator A x) \<partial>M)"
+  by (rule nn_integral_cong) (simp split: split_indicator)
+
+lemma nn_integral_indicator_singleton'[simp]:
+  assumes [measurable]: "{y} \<in> sets M"
+  shows "(\<integral>\<^sup>+x. ennreal (f x * indicator {y} x) \<partial>M) = f y * emeasure M {y}"
+  by (subst nn_integral_set_ennreal[symmetric]) (simp add: nn_integral_indicator_singleton)
+
+lemma nn_integral_add:
+  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
+  using nn_integral_linear[of f M g 1] by simp
+
+lemma nn_integral_setsum:
+  "(\<And>i. i \<in> P \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))"
+  by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add)
+
+lemma nn_integral_suminf:
+  assumes f: "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))"
+proof -
+  have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
+    using assms by (auto simp: AE_all_countable)
+  have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))"
+    by (rule suminf_eq_SUP)
+  also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
+    unfolding nn_integral_setsum[OF f] ..
+  also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
+    by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
+       (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
+  also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
+    by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP)
+  finally show ?thesis by simp
+qed
+
+lemma nn_integral_bound_simple_function:
+  assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>"
+  assumes f[measurable]: "simple_function M f"
+  assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
+  shows "nn_integral M f < \<infinity>"
+proof cases
+  assume "space M = {}"
+  then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)"
+    by (intro nn_integral_cong) auto
+  then show ?thesis by simp
+next
+  assume "space M \<noteq> {}"
+  with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>"
+    by (subst Max_less_iff) (auto simp: Max_ge_iff)
+
+  have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)"
+  proof (rule nn_integral_mono)
+    fix x assume "x \<in> space M"
+    with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x"
+      by (auto split: split_indicator intro!: Max_ge simple_functionD)
+  qed
+  also have "\<dots> < \<infinity>"
+    using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_Markov_inequality:
+  assumes u: "u \<in> borel_measurable M" and "A \<in> sets M"
+  shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
+    (is "(emeasure M) ?A \<le> _ * ?PI")
+proof -
+  have "?A \<in> sets M"
+    using \<open>A \<in> sets M\<close> u by auto
+  hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)"
+    using nn_integral_indicator by simp
+  also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)"
+    using u by (auto intro!: nn_integral_mono_AE simp: indicator_def)
+  also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
+    using assms by (auto intro!: nn_integral_cmult)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_noteq_infinite:
+  assumes g: "g \<in> borel_measurable M" and "integral\<^sup>N M g \<noteq> \<infinity>"
+  shows "AE x in M. g x \<noteq> \<infinity>"
+proof (rule ccontr)
+  assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)"
+  have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
+    using c g by (auto simp add: AE_iff_null)
+  then have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}"
+    by (auto simp: zero_less_iff_neq_zero)
+  then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}"
+    by (auto simp: ennreal_top_eq_mult_iff)
+  also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
+    using g by (subst nn_integral_cmult_indicator) auto
+  also have "\<dots> \<le> integral\<^sup>N M g"
+    using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def)
+  finally show False
+    using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by (auto simp: top_unique)
+qed
+
+lemma nn_integral_PInf:
+  assumes f: "f \<in> borel_measurable M" and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>"
+  shows "emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
+proof -
+  have "\<infinity> * emeasure M (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
+    using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets)
+  also have "\<dots> \<le> integral\<^sup>N M f"
+    by (auto intro!: nn_integral_mono simp: indicator_def)
+  finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
+    by simp
+  then show ?thesis
+    using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm)
+qed
+
+lemma simple_integral_PInf:
+  "simple_function M f \<Longrightarrow> integral\<^sup>S M f \<noteq> \<infinity> \<Longrightarrow> emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
+  by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function)
+
+lemma nn_integral_PInf_AE:
+  assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
+proof (rule AE_I)
+  show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
+    by (rule nn_integral_PInf[OF assms])
+  show "f -` {\<infinity>} \<inter> space M \<in> sets M"
+    using assms by (auto intro: borel_measurable_vimage)
+qed auto
+
+lemma nn_integral_diff:
+  assumes f: "f \<in> borel_measurable M"
+  and g: "g \<in> borel_measurable M"
+  and fin: "integral\<^sup>N M g \<noteq> \<infinity>"
+  and mono: "AE x in M. g x \<le> f x"
+  shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g"
+proof -
+  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+    using assms by auto
+  have "AE x in M. f x = f x - g x + g x"
+    using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto
+  then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g"
+    unfolding nn_integral_add[OF diff g, symmetric]
+    by (rule nn_integral_cong_AE)
+  show ?thesis unfolding **
+    using fin
+    by (cases rule: ennreal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto
+qed
+
+lemma nn_integral_mult_bounded_inf:
+  assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and c: "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x"
+  shows "(\<integral>\<^sup>+x. g x \<partial>M) < \<infinity>"
+proof -
+  have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)"
+    by (intro nn_integral_mono_AE ae)
+  also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>"
+    using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less)
+  finally show ?thesis .
+qed
+
+text \<open>Fatou's lemma: convergence theorem on limes inferior\<close>
+
+lemma nn_integral_monotone_convergence_INF_AE':
+  assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+    and *: "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
+proof (rule ennreal_minus_cancel)
+  have "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+x. f 0 x - (INF i. f i x) \<partial>M)"
+  proof (rule nn_integral_diff[symmetric])
+    have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
+      by (intro nn_integral_mono INF_lower) simp
+    with * show "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<noteq> \<infinity>"
+      by simp
+  qed (auto intro: INF_lower)
+  also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. f 0 x - f i x) \<partial>M)"
+    by (simp add: ennreal_INF_const_minus)
+  also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f 0 x - f i x \<partial>M))"
+  proof (intro nn_integral_monotone_convergence_SUP_AE)
+    show "AE x in M. f 0 x - f i x \<le> f 0 x - f (Suc i) x" for i
+      using f[of i] by eventually_elim (auto simp: ennreal_mono_minus)
+  qed simp
+  also have "\<dots> = (SUP i. nn_integral M (f 0) - (\<integral>\<^sup>+x. f i x \<partial>M))"
+  proof (subst nn_integral_diff[symmetric])
+    fix i
+    have dec: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
+      unfolding AE_all_countable using f by auto
+    then show "AE x in M. f i x \<le> f 0 x"
+      using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\<lambda>i. f i x" 0 i for x])
+    then have "(\<integral>\<^sup>+ x. f i x \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
+      by (rule nn_integral_mono_AE)
+    with * show "(\<integral>\<^sup>+ x. f i x \<partial>M) \<noteq> \<infinity>"
+      by simp
+  qed (insert f, auto simp: decseq_def le_fun_def)
+  finally show "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) =
+    integral\<^sup>N M (f 0) - (INF i. \<integral>\<^sup>+ x. f i x \<partial>M)"
+    by (simp add: ennreal_INF_const_minus)
+qed (insert *, auto intro!: nn_integral_mono intro: INF_lower)
+
+lemma nn_integral_monotone_convergence_INF_AE:
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
+  assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x"
+    and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+    and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
+proof -
+  { fix f :: "nat \<Rightarrow> ennreal" and j assume "decseq f"
+    then have "(INF i. f i) = (INF i. f (i + j))"
+      apply (intro INF_eq)
+      apply (rule_tac x="i" in bexI)
+      apply (auto simp: decseq_def le_fun_def)
+      done }
+  note INF_shift = this
+  have mono: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
+    using f by (auto simp: AE_all_countable)
+  then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)"
+    by eventually_elim (auto intro!: decseq_SucI INF_shift)
+  then have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF n. f (n + i) x) \<partial>M)"
+    by (rule nn_integral_cong_AE)
+  also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f (n + i) x \<partial>M))"
+    by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto)
+  also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f n x \<partial>M))"
+    by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_monotone_convergence_INF_decseq:
+  assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
+  using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (auto simp: decseq_Suc_iff le_fun_def)
+
+lemma nn_integral_liminf:
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
+  assumes u: "\<And>i. u i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
+proof -
+  have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (SUP n. \<integral>\<^sup>+ x. (INF i:{n..}. u i x) \<partial>M)"
+    unfolding liminf_SUP_INF using u
+    by (intro nn_integral_monotone_convergence_SUP_AE)
+       (auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
+  also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
+    by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_limsup:
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
+  assumes [measurable]: "\<And>i. u i \<in> borel_measurable M" "w \<in> borel_measurable M"
+  assumes bounds: "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
+  shows "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
+proof -
+  have bnd: "AE x in M. \<forall>i. u i x \<le> w x"
+    using bounds by (auto simp: AE_all_countable)
+  then have "(\<integral>\<^sup>+ x. (SUP n. u n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. w x \<partial>M)"
+    by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least)
+  then have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (INF n. \<integral>\<^sup>+ x. (SUP i:{n..}. u i x) \<partial>M)"
+    unfolding limsup_INF_SUP using bnd w
+    by (intro nn_integral_monotone_convergence_INF_AE')
+       (auto intro!: AE_I2 intro: SUP_least SUP_subset_mono)
+  also have "\<dots> \<ge> limsup (\<lambda>n. integral\<^sup>N M (u n))"
+    by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper)
+  finally (xtrans) show ?thesis .
+qed
+
+lemma nn_integral_LIMSEQ:
+  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M"
+    and u: "\<And>x. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
+  shows "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> integral\<^sup>N M u"
+proof -
+  have "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> (SUP n. integral\<^sup>N M (f n))"
+    using f by (intro LIMSEQ_SUP[of "\<lambda>n. integral\<^sup>N M (f n)"] incseq_nn_integral)
+  also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)"
+    using f by (intro nn_integral_monotone_convergence_SUP[symmetric])
+  also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)"
+    using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_dominated_convergence:
+  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. 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. u i x \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. u' x \<partial>M)"
+proof -
+  have "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
+    by (intro nn_integral_limsup[OF _ _ bound w]) auto
+  moreover have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
+    using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
+  moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
+    using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
+  moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
+    by (intro nn_integral_liminf) auto
+  moreover have "liminf (\<lambda>n. integral\<^sup>N M (u n)) \<le> limsup (\<lambda>n. integral\<^sup>N M (u n))"
+    by (intro Liminf_le_Limsup sequentially_bot)
+  ultimately show ?thesis
+    by (intro Liminf_eq_Limsup) auto
+qed
+
+lemma inf_continuous_nn_integral[order_continuous_intros]:
+  assumes f: "\<And>y. inf_continuous (f y)"
+  assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
+  assumes bnd: "\<And>x. (\<integral>\<^sup>+ y. f y x \<partial>M) \<noteq> \<infinity>"
+  shows "inf_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
+  unfolding inf_continuous_def
+proof safe
+  fix C :: "nat \<Rightarrow> 'b" assume C: "decseq C"
+  then show "(\<integral>\<^sup>+ y. f y (INFIMUM UNIV C) \<partial>M) = (INF i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
+    using inf_continuous_mono[OF f] bnd
+    by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def less_top
+             intro!: nn_integral_monotone_convergence_INF_decseq)
+qed
+
+lemma nn_integral_null_set:
+  assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0"
+proof -
+  have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
+  proof (intro nn_integral_cong_AE AE_I)
+    show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
+      by (auto simp: indicator_def)
+    show "(emeasure M) N = 0" "N \<in> sets M"
+      using assms by auto
+  qed
+  then show ?thesis by simp
+qed
+
+lemma nn_integral_0_iff:
+  assumes u: "u \<in> borel_measurable M"
+  shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
+    (is "_ \<longleftrightarrow> (emeasure M) ?A = 0")
+proof -
+  have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>N M u"
+    by (auto intro!: nn_integral_cong simp: indicator_def)
+  show ?thesis
+  proof
+    assume "(emeasure M) ?A = 0"
+    with nn_integral_null_set[of ?A M u] u
+    show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def)
+  next
+    assume *: "integral\<^sup>N M u = 0"
+    let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
+    have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))"
+    proof -
+      { fix n :: nat
+        from nn_integral_Markov_inequality[OF u, of ?A "of_nat n"] u
+        have "(emeasure M) (?M n \<inter> ?A) \<le> 0"
+          by (simp add: ennreal_of_nat_eq_real_of_nat u_eq *)
+        moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto
+        ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto }
+      thus ?thesis by simp
+    qed
+    also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)"
+    proof (safe intro!: SUP_emeasure_incseq)
+      fix n show "?M n \<inter> ?A \<in> sets M"
+        using u by (auto intro!: sets.Int)
+    next
+      show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
+      proof (safe intro!: incseq_SucI)
+        fix n :: nat and x
+        assume *: "1 \<le> real n * u x"
+        also have "real n * u x \<le> real (Suc n) * u x"
+          by (auto intro!: mult_right_mono)
+        finally show "1 \<le> real (Suc n) * u x" by auto
+      qed
+    qed
+    also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
+    proof (safe intro!: arg_cong[where f="(emeasure M)"])
+      fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
+      show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
+      proof (cases "u x" rule: ennreal_cases)
+        case (real r) with \<open>0 < u x\<close> have "0 < r" by auto
+        obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
+        hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto
+        hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto
+        thus ?thesis using \<open>0 < r\<close> real
+          by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric]
+                   simp del: ennreal_1)
+      qed (insert \<open>0 < u x\<close>, auto simp: ennreal_mult_top)
+    qed (auto simp: zero_less_iff_neq_zero)
+    finally show "emeasure M ?A = 0"
+      by (simp add: zero_less_iff_neq_zero)
+  qed
+qed
+
+lemma nn_integral_0_iff_AE:
+  assumes u: "u \<in> borel_measurable M"
+  shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
+proof -
+  have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
+    using u by auto
+  show "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
+    using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto
+qed
+
+lemma AE_iff_nn_integral:
+  "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>N M (indicator {x. \<not> P x}) = 0"
+  by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def])
+
+lemma nn_integral_less:
+  assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  assumes f: "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>"
+  assumes ord: "AE x in M. f x \<le> g x" "\<not> (AE x in M. g x \<le> f x)"
+  shows "(\<integral>\<^sup>+x. f x \<partial>M) < (\<integral>\<^sup>+x. g x \<partial>M)"
+proof -
+  have "0 < (\<integral>\<^sup>+x. g x - f x \<partial>M)"
+  proof (intro order_le_neq_trans notI)
+    assume "0 = (\<integral>\<^sup>+x. g x - f x \<partial>M)"
+    then have "AE x in M. g x - f x = 0"
+      using nn_integral_0_iff_AE[of "\<lambda>x. g x - f x" M] by simp
+    with ord(1) have "AE x in M. g x \<le> f x"
+      by eventually_elim (auto simp: ennreal_minus_eq_0)
+    with ord show False
+      by simp
+  qed simp
+  also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M) - (\<integral>\<^sup>+x. f x \<partial>M)"
+    using f by (subst nn_integral_diff) (auto simp: ord)
+  finally show ?thesis
+    using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top)
+qed
+
+lemma nn_integral_subalgebra:
+  assumes f: "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>N N f = integral\<^sup>N M f"
+proof -
+  have [simp]: "\<And>f :: 'a \<Rightarrow> ennreal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M"
+    using N by (auto simp: measurable_def)
+  have [simp]: "\<And>P. (AE x in N. P x) \<Longrightarrow> (AE x in M. P x)"
+    using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq)
+  have [simp]: "\<And>A. A \<in> sets N \<Longrightarrow> A \<in> sets M"
+    using N by auto
+  from f show ?thesis
+    apply induct
+    apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N)
+    apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric])
+    done
+qed
+
+lemma nn_integral_nat_function:
+  fixes f :: "'a \<Rightarrow> nat"
+  assumes "f \<in> measurable M (count_space UNIV)"
+  shows "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
+proof -
+  define F where "F i = {x\<in>space M. i < f x}" for i
+  with assms have [measurable]: "\<And>i. F i \<in> sets M"
+    by auto
+
+  { fix x assume "x \<in> space M"
+    have "(\<lambda>i. if i < f x then 1 else 0) sums (of_nat (f x)::real)"
+      using sums_If_finite[of "\<lambda>i. i < f x" "\<lambda>_. 1::real"] by simp
+    then have "(\<lambda>i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)"
+      unfolding ennreal_of_nat_eq_real_of_nat
+      by (subst sums_ennreal) auto
+    moreover have "\<And>i. ennreal (if i < f x then 1 else 0) = indicator (F i) x"
+      using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def)
+    ultimately have "of_nat (f x) = (\<Sum>i. indicator (F i) x :: ennreal)"
+      by (simp add: sums_iff) }
+  then have "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
+    by (simp cong: nn_integral_cong)
+  also have "\<dots> = (\<Sum>i. emeasure M (F i))"
+    by (simp add: nn_integral_suminf)
+  finally show ?thesis
+    by (simp add: F_def)
+qed
+
+lemma nn_integral_lfp:
+  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
+  assumes f: "sup_continuous f"
+  assumes g: "sup_continuous g"
+  assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
+  assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
+  shows "(\<integral>\<^sup>+\<omega>. lfp f \<omega> \<partial>M s) = lfp g s"
+proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f and P="\<lambda>f. f \<in> borel_measurable N", symmetric])
+  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N"
+  then show "(\<lambda>s. \<integral>\<^sup>+x. (SUP i. C i) x \<partial>M s) = (SUP i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
+    unfolding SUP_apply[abs_def]
+    by (subst nn_integral_monotone_convergence_SUP)
+       (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
+qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g)
+
+lemma nn_integral_gfp:
+  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
+  assumes f: "inf_continuous f" and g: "inf_continuous g"
+  assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
+  assumes bound: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f F x \<partial>M s) < \<infinity>"
+  assumes non_zero: "\<And>s. emeasure (M s) (space (M s)) \<noteq> 0"
+  assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
+  shows "(\<integral>\<^sup>+\<omega>. gfp f \<omega> \<partial>M s) = gfp g s"
+proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f
+    and P="\<lambda>F. F \<in> borel_measurable N \<and> (\<forall>s. (\<integral>\<^sup>+x. F x \<partial>M s) < \<infinity>)", symmetric])
+  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)"
+  then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
+    unfolding INF_apply[abs_def]
+    by (subst nn_integral_monotone_convergence_INF_decseq)
+       (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
+next
+  show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))"
+    by (subst step)
+       (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult
+             cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD])
+next
+  fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C"
+  with bound show "INFIMUM UNIV C \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (INFIMUM UNIV C) < \<infinity>)"
+    unfolding INF_apply[abs_def]
+    by (subst nn_integral_monotone_convergence_INF_decseq)
+       (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF)
+next
+  show "\<And>x. x \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) x < \<infinity>) \<Longrightarrow>
+         (\<lambda>s. integral\<^sup>N (M s) (f x)) = g (\<lambda>s. integral\<^sup>N (M s) x)"
+    by (subst step) auto
+qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g)
+
+subsection \<open>Integral under concrete measures\<close>
+
+lemma nn_integral_mono_measure:
+  assumes "sets M = sets N" "M \<le> N" shows "nn_integral M f \<le> nn_integral N f"
+  unfolding nn_integral_def
+proof (intro SUP_subset_mono)
+  note \<open>sets M = sets N\<close>[simp]  \<open>sets M = sets N\<close>[THEN sets_eq_imp_space_eq, simp]
+  show "{g. simple_function M g \<and> g \<le> f} \<subseteq> {g. simple_function N g \<and> g \<le> f}"
+    by (simp add: simple_function_def)
+  show "integral\<^sup>S M x \<le> integral\<^sup>S N x" for x
+    using le_measureD3[OF \<open>M \<le> N\<close>]
+    by (auto simp add: simple_integral_def intro!: setsum_mono mult_mono)
+qed
+
+lemma nn_integral_empty:
+  assumes "space M = {}"
+  shows "nn_integral M f = 0"
+proof -
+  have "(\<integral>\<^sup>+ x. f x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
+    by(rule nn_integral_cong)(simp add: assms)
+  thus ?thesis by simp
+qed
+
+lemma nn_integral_bot[simp]: "nn_integral bot f = 0"
+  by (simp add: nn_integral_empty)
+
+subsubsection \<open>Distributions\<close>
+
+lemma nn_integral_distr:
+  assumes T: "T \<in> measurable M M'" and f: "f \<in> borel_measurable (distr M M' T)"
+  shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
+  using f
+proof induct
+  case (cong f g)
+  with T show ?case
+    apply (subst nn_integral_cong[of _ f g])
+    apply simp
+    apply (subst nn_integral_cong[of _ "\<lambda>x. f (T x)" "\<lambda>x. g (T x)"])
+    apply (simp add: measurable_def Pi_iff)
+    apply simp
+    done
+next
+  case (set A)
+  then have eq: "\<And>x. x \<in> space M \<Longrightarrow> indicator A (T x) = indicator (T -` A \<inter> space M) x"
+    by (auto simp: indicator_def)
+  from set T show ?case
+    by (subst nn_integral_cong[OF eq])
+       (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets)
+qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add
+                   nn_integral_monotone_convergence_SUP le_fun_def incseq_def)
+
+subsubsection \<open>Counting space\<close>
+
+lemma simple_function_count_space[simp]:
+  "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
+  unfolding simple_function_def by simp
+
+lemma nn_integral_count_space:
+  assumes A: "finite {a\<in>A. 0 < f a}"
+  shows "integral\<^sup>N (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
+proof -
+  have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) =
+    (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
+    by (auto intro!: nn_integral_cong
+             simp add: indicator_def if_distrib setsum.If_cases[OF A] max_def le_less)
+  also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
+    by (subst nn_integral_setsum) (simp_all add: AE_count_space  less_imp_le)
+  also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
+    by (auto intro!: setsum.cong simp: one_ennreal_def[symmetric] max_def)
+  finally show ?thesis by (simp add: max.absorb2)
+qed
+
+lemma nn_integral_count_space_finite:
+    "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
+  by (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le)
+
+lemma nn_integral_count_space':
+  assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "A \<subseteq> B"
+  shows "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>x\<in>A. f x)"
+proof -
+  have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 0 < f a. f a)"
+    using assms(2,3)
+    by (intro nn_integral_count_space finite_subset[OF _ \<open>finite A\<close>]) (auto simp: less_le)
+  also have "\<dots> = (\<Sum>a\<in>A. f a)"
+    using assms by (intro setsum.mono_neutral_cong_left) (auto simp: less_le)
+  finally show ?thesis .
+qed
+
+lemma nn_integral_bij_count_space:
+  assumes g: "bij_betw g A B"
+  shows "(\<integral>\<^sup>+x. f (g x) \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
+  using g[THEN bij_betw_imp_funcset]
+  by (subst distr_bij_count_space[OF g, symmetric])
+     (auto intro!: nn_integral_distr[symmetric])
+
+lemma nn_integral_indicator_finite:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  assumes f: "finite A" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
+  shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
+proof -
+  from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
+    by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] setsum.If_cases)
+  also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
+    by (subst nn_integral_setsum) auto
+  finally show ?thesis .
+qed
+
+lemma nn_integral_count_space_nat:
+  fixes f :: "nat \<Rightarrow> ennreal"
+  shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)"
+proof -
+  have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) =
+    (\<integral>\<^sup>+i. (\<Sum>j. f j * indicator {j} i) \<partial>count_space UNIV)"
+  proof (intro nn_integral_cong)
+    fix i
+    have "f i = (\<Sum>j\<in>{i}. f j * indicator {j} i)"
+      by simp
+    also have "\<dots> = (\<Sum>j. f j * indicator {j} i)"
+      by (rule suminf_finite[symmetric]) auto
+    finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
+  qed
+  also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
+    by (rule nn_integral_suminf) auto
+  finally show ?thesis
+    by simp
+qed
+
+lemma nn_integral_enat_function:
+  assumes f: "f \<in> measurable M (count_space UNIV)"
+  shows "(\<integral>\<^sup>+ x. ennreal_of_enat (f x) \<partial>M) = (\<Sum>t. emeasure M {x \<in> space M. t < f x})"
+proof -
+  define F where "F i = {x\<in>space M. i < f x}" for i :: nat
+  with assms have [measurable]: "\<And>i. F i \<in> sets M"
+    by auto
+
+  { fix x assume "x \<in> space M"
+    have "(\<lambda>i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)"
+      using sums_If_finite[of "\<lambda>r. r < f x" "\<lambda>_. 1 :: ennreal"]
+      by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal)
+    also have "(\<lambda>i. (if i < f x then 1 else 0)) = (\<lambda>i. indicator (F i) x)"
+      using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def fun_eq_iff)
+    finally have "ennreal_of_enat (f x) = (\<Sum>i. indicator (F i) x)"
+      by (simp add: sums_iff) }
+  then have "(\<integral>\<^sup>+x. ennreal_of_enat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
+    by (simp cong: nn_integral_cong)
+  also have "\<dots> = (\<Sum>i. emeasure M (F i))"
+    by (simp add: nn_integral_suminf)
+  finally show ?thesis
+    by (simp add: F_def)
+qed
+
+lemma nn_integral_count_space_nn_integral:
+  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ennreal"
+  assumes "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)"
+proof cases
+  assume "finite I" then show ?thesis
+    by (simp add: nn_integral_count_space_finite nn_integral_setsum)
+next
+  assume "infinite I"
+  then have [simp]: "I \<noteq> {}"
+    by auto
+  note * = bij_betw_from_nat_into[OF \<open>countable I\<close> \<open>infinite I\<close>]
+  have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>n. f (from_nat_into I n))"
+    by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat)
+  show ?thesis
+    by (simp add: ** nn_integral_suminf from_nat_into)
+qed
+
+lemma emeasure_UN_countable:
+  assumes sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I[simp]: "countable I"
+  assumes disj: "disjoint_family_on X I"
+  shows "emeasure M (UNION I X) = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)"
+proof -
+  have eq: "\<And>x. indicator (UNION I X) x = \<integral>\<^sup>+ i. indicator (X i) x \<partial>count_space I"
+  proof cases
+    fix x assume x: "x \<in> UNION I X"
+    then obtain j where j: "x \<in> X j" "j \<in> I"
+      by auto
+    with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ennreal)"
+      by (auto simp: disjoint_family_on_def split: split_indicator)
+    with x j show "?thesis x"
+      by (simp cong: nn_integral_cong_simp)
+  qed (auto simp: nn_integral_0_iff_AE)
+
+  note sets.countable_UN'[unfolded subset_eq, measurable]
+  have "emeasure M (UNION I X) = (\<integral>\<^sup>+x. indicator (UNION I X) x \<partial>M)"
+    by simp
+  also have "\<dots> = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. indicator (X i) x \<partial>M \<partial>count_space I)"
+    by (simp add: eq nn_integral_count_space_nn_integral)
+  finally show ?thesis
+    by (simp cong: nn_integral_cong_simp)
+qed
+
+lemma emeasure_countable_singleton:
+  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" and X: "countable X"
+  shows "emeasure M X = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
+proof -
+  have "emeasure M (\<Union>i\<in>X. {i}) = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
+    using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def)
+  also have "(\<Union>i\<in>X. {i}) = X" by auto
+  finally show ?thesis .
+qed
+
+lemma measure_eqI_countable:
+  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A"
+  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
+  shows "M = N"
+proof (rule measure_eqI)
+  fix X assume "X \<in> sets M"
+  then have X: "X \<subseteq> A" by auto
+  moreover from A X have "countable X" by (auto dest: countable_subset)
+  ultimately have
+    "emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
+    "emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
+    by (auto intro!: emeasure_countable_singleton)
+  moreover have "(\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X) = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
+    using X by (intro nn_integral_cong eq) auto
+  ultimately show "emeasure M X = emeasure N X"
+    by simp
+qed simp
+
+lemma measure_eqI_countable_AE:
+  assumes [simp]: "sets M = UNIV" "sets N = UNIV"
+  assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>" and [simp]: "countable \<Omega>"
+  assumes eq: "\<And>x. x \<in> \<Omega> \<Longrightarrow> emeasure M {x} = emeasure N {x}"
+  shows "M = N"
+proof (rule measure_eqI)
+  fix A
+  have "emeasure N A = emeasure N {x\<in>\<Omega>. x \<in> A}"
+    using ae by (intro emeasure_eq_AE) auto
+  also have "\<dots> = (\<integral>\<^sup>+x. emeasure N {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
+    by (intro emeasure_countable_singleton) auto
+  also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
+    by (intro nn_integral_cong eq[symmetric]) auto
+  also have "\<dots> = emeasure M {x\<in>\<Omega>. x \<in> A}"
+    by (intro emeasure_countable_singleton[symmetric]) auto
+  also have "\<dots> = emeasure M A"
+    using ae by (intro emeasure_eq_AE) auto
+  finally show "emeasure M A = emeasure N A" ..
+qed simp
+
+lemma nn_integral_monotone_convergence_SUP_nat:
+  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
+  assumes chain: "Complete_Partial_Order.chain op \<le> (f ` Y)"
+  and nonempty: "Y \<noteq> {}"
+  shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space UNIV) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))"
+  (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _")
+proof (rule order_class.order.antisym)
+  show "?rhs \<le> ?lhs"
+    by (auto intro!: SUP_least SUP_upper nn_integral_mono)
+next
+  have "\<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i:Y. f i x) = (SUP i. g i)" for x
+    by (rule ennreal_Sup_countable_SUP) (simp add: nonempty)
+  then obtain g where incseq: "\<And>x. incseq (g x)"
+    and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y"
+    and sup: "\<And>x. (SUP i:Y. f i x) = (SUP i. g x i)" by moura
+  from incseq have incseq': "incseq (\<lambda>i x. g x i)"
+    by(blast intro: incseq_SucI le_funI dest: incseq_SucD)
+
+  have "?lhs = \<integral>\<^sup>+ x. (SUP i. g x i) \<partial>?M" by(simp add: sup)
+  also have "\<dots> = (SUP i. \<integral>\<^sup>+ x. g x i \<partial>?M)" using incseq'
+    by(rule nn_integral_monotone_convergence_SUP) simp
+  also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
+  proof(rule SUP_least)
+    fix n
+    have "\<And>x. \<exists>i. g x n = f i x \<and> i \<in> Y" using range by blast
+    then obtain I where I: "\<And>x. g x n = f (I x) x" "\<And>x. I x \<in> Y" by moura
+
+    have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)"
+      by(rule nn_integral_count_space_nat)
+    also have "\<dots> = (SUP m. \<Sum>x<m. g x n)"
+      by(rule suminf_eq_SUP)
+    also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
+    proof(rule SUP_mono)
+      fix m
+      show "\<exists>m'\<in>Y. (\<Sum>x<m. g x n) \<le> (\<integral>\<^sup>+ x. f m' x \<partial>?M)"
+      proof(cases "m > 0")
+        case False
+        thus ?thesis using nonempty by auto
+      next
+        case True
+        let ?Y = "I ` {..<m}"
+        have "f ` ?Y \<subseteq> f ` Y" using I by auto
+        with chain have chain': "Complete_Partial_Order.chain op \<le> (f ` ?Y)" by(rule chain_subset)
+        hence "Sup (f ` ?Y) \<in> f ` ?Y"
+          by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff)
+        then obtain m' where "m' < m" and m': "(SUP i:?Y. f i) = f (I m')" by auto
+        have "I m' \<in> Y" using I by blast
+        have "(\<Sum>x<m. g x n) \<le> (\<Sum>x<m. f (I m') x)"
+        proof(rule setsum_mono)
+          fix x
+          assume "x \<in> {..<m}"
+          hence "x < m" by simp
+          have "g x n = f (I x) x" by(simp add: I)
+          also have "\<dots> \<le> (SUP i:?Y. f i) x" unfolding Sup_fun_def image_image
+            using \<open>x \<in> {..<m}\<close> by (rule Sup_upper [OF imageI])
+          also have "\<dots> = f (I m') x" unfolding m' by simp
+          finally show "g x n \<le> f (I m') x" .
+        qed
+        also have "\<dots> \<le> (SUP m. (\<Sum>x<m. f (I m') x))"
+          by(rule SUP_upper) simp
+        also have "\<dots> = (\<Sum>x. f (I m') x)"
+          by(rule suminf_eq_SUP[symmetric])
+        also have "\<dots> = (\<integral>\<^sup>+ x. f (I m') x \<partial>?M)"
+          by(rule nn_integral_count_space_nat[symmetric])
+        finally show ?thesis using \<open>I m' \<in> Y\<close> by blast
+      qed
+    qed
+    finally show "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) \<le> \<dots>" .
+  qed
+  finally show "?lhs \<le> ?rhs" .
+qed
+
+lemma power_series_tendsto_at_left:
+  assumes nonneg: "\<And>i. 0 \<le> f i" and summable: "\<And>z. 0 \<le> z \<Longrightarrow> z < 1 \<Longrightarrow> summable (\<lambda>n. f n * z^n)"
+  shows "((\<lambda>z. ennreal (\<Sum>n. f n * z^n)) \<longlongrightarrow> (\<Sum>n. ennreal (f n))) (at_left (1::real))"
+proof (intro tendsto_at_left_sequentially)
+  show "0 < (1::real)" by simp
+  fix S :: "nat \<Rightarrow> real" assume S: "\<And>n. S n < 1" "\<And>n. 0 < S n" "S \<longlonglongrightarrow> 1" "incseq S"
+  then have S_nonneg: "\<And>i. 0 \<le> S i" by (auto intro: less_imp_le)
+
+  have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) \<longlonglongrightarrow> (\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV)"
+  proof (rule nn_integral_LIMSEQ)
+    show "incseq (\<lambda>i n. ennreal (f n * S i^n))"
+      using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI
+                       simp: incseq_def le_fun_def less_imp_le)
+    fix n have "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n * 1^n)"
+      by (intro tendsto_intros tendsto_ennrealI S)
+    then show "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n)"
+      by simp
+  qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg)
+  also have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) = (\<lambda>i. \<Sum>n. f n * S i^n)"
+    by (subst nn_integral_count_space_nat)
+       (intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg
+              zero_le_power summable S)+
+  also have "(\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV) = (\<Sum>n. ennreal (f n))"
+    by (simp add: nn_integral_count_space_nat nonneg)
+  finally show "(\<lambda>n. ennreal (\<Sum>na. f na * S n ^ na)) \<longlonglongrightarrow> (\<Sum>n. ennreal (f n))" .
+qed
+
+subsubsection \<open>Measures with Restricted Space\<close>
+
+lemma simple_function_restrict_space_ennreal:
+  fixes f :: "'a \<Rightarrow> ennreal"
+  assumes "\<Omega> \<inter> space M \<in> sets M"
+  shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)"
+proof -
+  { assume "finite (f ` space (restrict_space M \<Omega>))"
+    then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
+    then have "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
+      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
+  moreover
+  { assume "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
+    then have "finite (f ` space (restrict_space M \<Omega>))"
+      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
+  ultimately show ?thesis
+    unfolding
+      simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms]
+    by auto
+qed
+
+lemma simple_function_restrict_space:
+  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
+  assumes "\<Omega> \<inter> space M \<in> sets M"
+  shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
+proof -
+  { assume "finite (f ` space (restrict_space M \<Omega>))"
+    then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
+    then have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
+      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
+  moreover
+  { assume "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
+    then have "finite (f ` space (restrict_space M \<Omega>))"
+      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
+  ultimately show ?thesis
+    unfolding simple_function_iff_borel_measurable
+      borel_measurable_restrict_space_iff[OF assms]
+    by auto
+qed
+
+lemma simple_integral_restrict_space:
+  assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f"
+  shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)"
+  using simple_function_restrict_space_ennreal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
+  by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def
+           split: split_indicator split_indicator_asm
+           intro!: setsum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])
+
+lemma nn_integral_restrict_space:
+  assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
+  shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)"
+proof -
+  let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> f \<and> (\<forall>x. s x < top)}"
+  have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)"
+  proof (safe intro!: image_eqI)
+    fix s assume s: "simple_function ?R s" "s \<le> f" "\<forall>x. s x < top"
+    from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)"
+      by (intro simple_integral_restrict_space) auto
+    from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
+      by (simp add: simple_function_restrict_space_ennreal)
+    from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> (\<lambda>x. f x * indicator \<Omega> x)"
+      "\<And>x. s x * indicator \<Omega> x < top"
+      by (auto split: split_indicator simp: le_fun_def image_subset_iff)
+  next
+    fix s assume s: "simple_function M s" "s \<le> (\<lambda>x. f x * indicator \<Omega> x)" "\<forall>x. s x < top"
+    then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s')
+      by (intro simple_function_mult simple_function_indicator) auto
+    also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
+      by (rule simple_function_cong) (auto split: split_indicator)
+    finally show sf: "simple_function (restrict_space M \<Omega>) s"
+      by (simp add: simple_function_restrict_space_ennreal)
+
+    from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)"
+      by (auto simp add: fun_eq_iff le_fun_def image_subset_iff
+                  split: split_indicator split_indicator_asm
+                  intro: antisym)
+
+    show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s"
+      by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf])
+    show "\<And>x. s x < top"
+      using s by (auto simp: image_subset_iff)
+    from s show "s \<le> f"
+      by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm)
+  qed
+  then show ?thesis
+    unfolding nn_integral_def_finite by (simp cong del: strong_SUP_cong)
+qed
+
+lemma nn_integral_count_space_indicator:
+  assumes "NO_MATCH (UNIV::'a set) (X::'a set)"
+  shows "(\<integral>\<^sup>+x. f x \<partial>count_space X) = (\<integral>\<^sup>+x. f x * indicator X x \<partial>count_space UNIV)"
+  by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space)
+
+lemma nn_integral_count_space_eq:
+  "(\<And>x. x \<in> A - B \<Longrightarrow> f x = 0) \<Longrightarrow> (\<And>x. x \<in> B - A \<Longrightarrow> f x = 0) \<Longrightarrow>
+    (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
+  by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
+
+lemma nn_integral_ge_point:
+  assumes "x \<in> A"
+  shows "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
+proof -
+  from assms have "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space {x}"
+    by(auto simp add: nn_integral_count_space_finite max_def)
+  also have "\<dots> = \<integral>\<^sup>+ x'. p x' * indicator {x} x' \<partial>count_space A"
+    using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong)
+  also have "\<dots> \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
+    by(rule nn_integral_mono)(simp add: indicator_def)
+  finally show ?thesis .
+qed
+
+subsubsection \<open>Measure spaces with an associated density\<close>
+
+definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
+  "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
+
+lemma
+  shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M"
+    and space_density[simp]: "space (density M f) = space M"
+  by (auto simp: density_def)
+
+(* FIXME: add conversion to simplify space, sets and measurable *)
+lemma space_density_imp[measurable_dest]:
+  "\<And>x M f. x \<in> space (density M f) \<Longrightarrow> x \<in> space M" by auto
+
+lemma
+  shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'"
+    and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'"
+    and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u"
+  unfolding measurable_def simple_function_def by simp_all
+
+lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow>
+  (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'"
+  unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed)
+
+lemma emeasure_density:
+  assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M"
+  shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
+    (is "_ = ?\<mu> A")
+  unfolding density_def
+proof (rule emeasure_measure_of_sigma)
+  show "sigma_algebra (space M) (sets M)" ..
+  show "positive (sets M) ?\<mu>"
+    using f by (auto simp: positive_def)
+  show "countably_additive (sets M) ?\<mu>"
+  proof (intro countably_additiveI)
+    fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M"
+    then have "\<And>i. A i \<in> sets M" by auto
+    then have *: "\<And>i. (\<lambda>x. f x * indicator (A i) x) \<in> borel_measurable M"
+      by auto
+    assume disj: "disjoint_family A"
+    then have "(\<Sum>n. ?\<mu> (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M)"
+       using f * by (subst nn_integral_suminf) auto
+    also have "(\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M) = (\<integral>\<^sup>+ x. f x * (\<Sum>n. indicator (A n) x) \<partial>M)"
+      using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE)
+    also have "\<dots> = (\<integral>\<^sup>+ x. f x * indicator (\<Union>n. A n) x \<partial>M)"
+      unfolding suminf_indicator[OF disj] ..
+    finally show "(\<Sum>i. \<integral>\<^sup>+ x. f x * indicator (A i) x \<partial>M) = \<integral>\<^sup>+ x. f x * indicator (\<Union>i. A i) x \<partial>M" .
+  qed
+qed fact
+
+lemma null_sets_density_iff:
+  assumes f: "f \<in> borel_measurable M"
+  shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x = 0)"
+proof -
+  { assume "A \<in> sets M"
+    have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> emeasure M {x \<in> space M. f x * indicator A x \<noteq> 0} = 0"
+      using f \<open>A \<in> sets M\<close> by (intro nn_integral_0_iff) auto
+    also have "\<dots> \<longleftrightarrow> (AE x in M. f x * indicator A x = 0)"
+      using f \<open>A \<in> sets M\<close> by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
+    also have "(AE x in M. f x * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
+      by (auto simp add: indicator_def max_def split: if_split_asm)
+    finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
+  with f show ?thesis
+    by (simp add: null_sets_def emeasure_density cong: conj_cong)
+qed
+
+lemma AE_density:
+  assumes f: "f \<in> borel_measurable M"
+  shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)"
+proof
+  assume "AE x in density M f. P x"
+  with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x = 0"
+    by (auto simp: eventually_ae_filter null_sets_density_iff)
+  then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto
+  with ae show "AE x in M. 0 < f x \<longrightarrow> P x"
+    by (rule eventually_elim2) auto
+next
+  fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x"
+  then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
+    by (auto simp: eventually_ae_filter)
+  then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. f x = 0}"
+    "N \<union> {x\<in>space M. f x = 0} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
+    using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in)
+  show "AE x in density M f. P x"
+    using ae2
+    unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f]
+    by (intro exI[of _ "N \<union> {x\<in>space M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
+qed
+
+lemma nn_integral_density:
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)"
+using g proof induct
+  case (cong u v)
+  then show ?case
+    apply (subst nn_integral_cong[OF cong(3)])
+    apply (simp_all cong: nn_integral_cong)
+    done
+next
+  case (set A) then show ?case
+    by (simp add: emeasure_density f)
+next
+  case (mult u c)
+  moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps)
+  ultimately show ?case
+    using f by (simp add: nn_integral_cmult)
+next
+  case (add u v)
+  then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x"
+    by (simp add: distrib_left)
+  with add f show ?case
+    by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric])
+next
+  case (seq U)
+  have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)"
+    by eventually_elim (simp add: SUP_mult_left_ennreal seq)
+  from seq f show ?case
+    apply (simp add: nn_integral_monotone_convergence_SUP)
+    apply (subst nn_integral_cong_AE[OF eq])
+    apply (subst nn_integral_monotone_convergence_SUP_AE)
+    apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono)
+    done
+qed
+
+lemma density_distr:
+  assumes [measurable]: "f \<in> borel_measurable N" "X \<in> measurable M N"
+  shows "density (distr M N X) f = distr (density M (\<lambda>x. f (X x))) N X"
+  by (intro measure_eqI)
+     (auto simp add: emeasure_density nn_integral_distr emeasure_distr
+           split: split_indicator intro!: nn_integral_cong)
+
+lemma emeasure_restricted:
+  assumes S: "S \<in> sets M" and X: "X \<in> sets M"
+  shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)"
+proof -
+  have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)"
+    using S X by (simp add: emeasure_density)
+  also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)"
+    by (auto intro!: nn_integral_cong simp: indicator_def)
+  also have "\<dots> = emeasure M (S \<inter> X)"
+    using S X by (simp add: sets.Int)
+  finally show ?thesis .
+qed
+
+lemma measure_restricted:
+  "S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)"
+  by (simp add: emeasure_restricted measure_def)
+
+lemma (in finite_measure) finite_measure_restricted:
+  "S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))"
+  by standard (simp add: emeasure_restricted)
+
+lemma emeasure_density_const:
+  "A \<in> sets M \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A"
+  by (auto simp: nn_integral_cmult_indicator emeasure_density)
+
+lemma measure_density_const:
+  "A \<in> sets M \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = enn2real c * measure M A"
+  by (auto simp: emeasure_density_const measure_def enn2real_mult)
+
+lemma density_density_eq:
+   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
+   density (density M f) g = density M (\<lambda>x. f x * g x)"
+  by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps)
+
+lemma distr_density_distr:
+  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
+    and inv: "\<forall>x\<in>space M. T' (T x) = x"
+  assumes f: "f \<in> borel_measurable M'"
+  shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L")
+proof (rule measure_eqI)
+  fix A assume A: "A \<in> sets ?R"
+  { fix x assume "x \<in> space M"
+    with sets.sets_into_space[OF A]
+    have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ennreal)"
+      using T inv by (auto simp: indicator_def measurable_space) }
+  with A T T' f show "emeasure ?R A = emeasure ?L A"
+    by (simp add: measurable_comp emeasure_density emeasure_distr
+                  nn_integral_distr measurable_sets cong: nn_integral_cong)
+qed simp
+
+lemma density_density_divide:
+  fixes f g :: "'a \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
+  assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
+  shows "density (density M f) (\<lambda>x. g x / f x) = density M g"
+proof -
+  have "density M g = density M (\<lambda>x. ennreal (f x) * ennreal (g x / f x))"
+    using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric])
+  then show ?thesis
+    using f g by (subst density_density_eq) auto
+qed
+
+lemma density_1: "density M (\<lambda>_. 1) = M"
+  by (intro measure_eqI) (auto simp: emeasure_density)
+
+lemma emeasure_density_add:
+  assumes X: "X \<in> sets M"
+  assumes Mf[measurable]: "f \<in> borel_measurable M"
+  assumes Mg[measurable]: "g \<in> borel_measurable M"
+  shows "emeasure (density M f) X + emeasure (density M g) X =
+           emeasure (density M (\<lambda>x. f x + g x)) X"
+  using assms
+  apply (subst (1 2 3) emeasure_density, simp_all) []
+  apply (subst nn_integral_add[symmetric], simp_all) []
+  apply (intro nn_integral_cong, simp split: split_indicator)
+  done
+
+subsubsection \<open>Point measure\<close>
+
+definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
+  "point_measure A f = density (count_space A) f"
+
+lemma
+  shows space_point_measure: "space (point_measure A f) = A"
+    and sets_point_measure: "sets (point_measure A f) = Pow A"
+  by (auto simp: point_measure_def)
+
+lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)"
+  by (simp add: sets_point_measure)
+
+lemma measurable_point_measure_eq1[simp]:
+  "g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M"
+  unfolding point_measure_def by simp
+
+lemma measurable_point_measure_eq2_finite[simp]:
+  "finite A \<Longrightarrow>
+   g \<in> measurable M (point_measure A f) \<longleftrightarrow>
+    (g \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. g -` {a} \<inter> space M \<in> sets M))"
+  unfolding point_measure_def by (simp add: measurable_count_space_eq2)
+
+lemma simple_function_point_measure[simp]:
+  "simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)"
+  by (simp add: point_measure_def)
+
+lemma emeasure_point_measure:
+  assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
+  shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)"
+proof -
+  have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
+    using \<open>X \<subseteq> A\<close> by auto
+  with A show ?thesis
+    by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def)
+qed
+
+lemma emeasure_point_measure_finite:
+  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
+  by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
+
+lemma emeasure_point_measure_finite2:
+  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
+  by (subst emeasure_point_measure)
+     (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
+
+lemma null_sets_point_measure_iff:
+  "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x = 0)"
+ by (auto simp: AE_count_space null_sets_density_iff point_measure_def)
+
+lemma AE_point_measure:
+  "(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)"
+  unfolding point_measure_def
+  by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def)
+
+lemma nn_integral_point_measure:
+  "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
+    integral\<^sup>N (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
+  unfolding point_measure_def
+  by (subst nn_integral_density)
+     (simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff)
+
+lemma nn_integral_point_measure_finite:
+  "finite A \<Longrightarrow> integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
+  by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le)
+
+subsubsection \<open>Uniform measure\<close>
+
+definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
+
+lemma
+  shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M"
+    and space_uniform_measure[simp]: "space (uniform_measure M A) = space M"
+  by (auto simp: uniform_measure_def)
+
+lemma emeasure_uniform_measure[simp]:
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A"
+proof -
+  from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
+    by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator
+             intro!: nn_integral_cong)
+  also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A"
+    using A B
+    by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute)
+  finally show ?thesis .
+qed
+
+lemma measure_uniform_measure[simp]:
+  assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M"
+  shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A"
+  using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A
+  by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ennreal2_cases)
+     (simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide)
+
+lemma AE_uniform_measureI:
+  "A \<in> sets M \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x) \<Longrightarrow> (AE x in uniform_measure M A. P x)"
+  unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def)
+
+lemma emeasure_uniform_measure_1:
+  "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> emeasure (uniform_measure M S) S = 1"
+  by (subst emeasure_uniform_measure)
+     (simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal
+                    zero_less_iff_neq_zero[symmetric])
+
+lemma nn_integral_uniform_measure:
+  assumes f[measurable]: "f \<in> borel_measurable M" and S[measurable]: "S \<in> sets M"
+  shows "(\<integral>\<^sup>+x. f x \<partial>uniform_measure M S) = (\<integral>\<^sup>+x. f x * indicator S x \<partial>M) / emeasure M S"
+proof -
+  { assume "emeasure M S = \<infinity>"
+    then have ?thesis
+      by (simp add: uniform_measure_def nn_integral_density f) }
+  moreover
+  { assume [simp]: "emeasure M S = 0"
+    then have ae: "AE x in M. x \<notin> S"
+      using sets.sets_into_space[OF S]
+      by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong)
+    from ae have "(\<integral>\<^sup>+ x. indicator S x / 0 * f x \<partial>M) = 0"
+      by (subst nn_integral_0_iff_AE) auto
+    moreover from ae have "(\<integral>\<^sup>+ x. f x * indicator S x \<partial>M) = 0"
+      by (subst nn_integral_0_iff_AE) auto
+    ultimately have ?thesis
+      by (simp add: uniform_measure_def nn_integral_density f) }
+  moreover have "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> ?thesis"
+    unfolding uniform_measure_def
+    by (subst nn_integral_density)
+       (auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute)
+  ultimately show ?thesis by blast
+qed
+
+lemma AE_uniform_measure:
+  assumes "emeasure M A \<noteq> 0" "emeasure M A < \<infinity>"
+  shows "(AE x in uniform_measure M A. P x) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x)"
+proof -
+  have "A \<in> sets M"
+    using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets)
+  moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> A"
+    using assms
+    by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide)
+  ultimately show ?thesis
+    unfolding uniform_measure_def by (simp add: AE_density)
+qed
+
+subsubsection \<open>Null measure\<close>
+
+lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 0)"
+  by (intro measure_eqI) (simp_all add: emeasure_density)
+
+lemma nn_integral_null_measure[simp]: "(\<integral>\<^sup>+x. f x \<partial>null_measure M) = 0"
+  by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def
+           intro!: exI[of _ "\<lambda>x. 0"])
+
+lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M"
+proof (intro measure_eqI)
+  fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A"
+    by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure)
+qed simp
+
+subsubsection \<open>Uniform count measure\<close>
+
+definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
+
+lemma
+  shows space_uniform_count_measure: "space (uniform_count_measure A) = A"
+    and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A"
+    unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure)
+
+lemma sets_uniform_count_measure_count_space[measurable_cong]:
+  "sets (uniform_count_measure A) = sets (count_space A)"
+  by (simp add: sets_uniform_count_measure)
+
+lemma emeasure_uniform_count_measure:
+  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A"
+  by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult
+                ennreal_of_nat_eq_real_of_nat)
+
+lemma measure_uniform_count_measure:
+  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A"
+  by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult)
+
+lemma space_uniform_count_measure_empty_iff [simp]:
+  "space (uniform_count_measure X) = {} \<longleftrightarrow> X = {}"
+by(simp add: space_uniform_count_measure)
+
+lemma sets_uniform_count_measure_eq_UNIV [simp]:
+  "sets (uniform_count_measure UNIV) = UNIV \<longleftrightarrow> True"
+  "UNIV = sets (uniform_count_measure UNIV) \<longleftrightarrow> True"
+by(simp_all add: sets_uniform_count_measure)
+
+subsubsection \<open>Scaled measure\<close>
+
+lemma nn_integral_scale_measure:
+  assumes f: "f \<in> borel_measurable M"
+  shows "nn_integral (scale_measure r M) f = r * nn_integral M f"
+  using f
+proof induction
+  case (cong f g)
+  thus ?case
+    by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp)
+next
+  case (mult f c)
+  thus ?case
+    by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute)
+next
+  case (add f g)
+  thus ?case
+    by(simp add: nn_integral_add distrib_left)
+next
+  case (seq U)
+  thus ?case
+    by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal)
+qed simp
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Radon_Nikodym.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,1064 @@
+(*  Title:      HOL/Probability/Radon_Nikodym.thy
+    Author:     Johannes Hölzl, TU München
+*)
+
+section \<open>Radon-Nikod{\'y}m derivative\<close>
+
+theory Radon_Nikodym
+imports Bochner_Integration
+begin
+
+definition diff_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
+where
+  "diff_measure M N = measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)"
+
+lemma
+  shows space_diff_measure[simp]: "space (diff_measure M N) = space M"
+    and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M"
+  by (auto simp: diff_measure_def)
+
+lemma emeasure_diff_measure:
+  assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N"
+  assumes pos: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure N A \<le> emeasure M A" and A: "A \<in> sets M"
+  shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\<mu> A")
+  unfolding diff_measure_def
+proof (rule emeasure_measure_of_sigma)
+  show "sigma_algebra (space M) (sets M)" ..
+  show "positive (sets M) ?\<mu>"
+    using pos by (simp add: positive_def)
+  show "countably_additive (sets M) ?\<mu>"
+  proof (rule countably_additiveI)
+    fix A :: "nat \<Rightarrow> _"  assume A: "range A \<subseteq> sets M" and "disjoint_family A"
+    then have suminf:
+      "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
+      "(\<Sum>i. emeasure N (A i)) = emeasure N (\<Union>i. A i)"
+      by (simp_all add: suminf_emeasure sets_eq)
+    with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
+      (\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))"
+      using fin pos[of "A _"]
+      by (intro ennreal_suminf_minus)
+         (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure)
+    then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
+      emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) "
+      by (simp add: suminf)
+  qed
+qed fact
+
+lemma (in sigma_finite_measure) Ex_finite_integrable_function:
+  "\<exists>h\<in>borel_measurable M. integral\<^sup>N M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>)"
+proof -
+  obtain A :: "nat \<Rightarrow> 'a set" where
+    range[measurable]: "range A \<subseteq> sets M" and
+    space: "(\<Union>i. A i) = space M" and
+    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and
+    disjoint: "disjoint_family A"
+    using sigma_finite_disjoint by blast
+  let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)"
+  have [measurable]: "\<And>i. A i \<in> sets M"
+    using range by fastforce+
+  have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
+  proof
+    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
+      using measure[of i]
+      by (auto intro!: dense simp: ennreal_inverse_positive ennreal_mult_eq_top_iff power_eq_top_ennreal)
+  qed
+  from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
+    "\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto
+  { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
+  let ?h = "\<lambda>x. \<Sum>i. n i * indicator (A i) x"
+  show ?thesis
+  proof (safe intro!: bexI[of _ ?h] del: notI)
+    have "integral\<^sup>N M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
+      by (simp add: nn_integral_suminf nn_integral_cmult_indicator)
+    also have "\<dots> \<le> (\<Sum>i. ennreal ((1/2)^Suc i))"
+    proof (intro suminf_le allI)
+      fix N
+      have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)"
+        using n[of N] by (intro mult_right_mono) auto
+      also have "\<dots> = (1/2)^Suc N * (inverse (emeasure M (A N)) * emeasure M (A N))"
+        using measure[of N]
+        by (simp add: ennreal_inverse_power divide_ennreal_def ennreal_inverse_mult
+                      power_eq_top_ennreal less_top[symmetric] mult_ac
+                 del: power_Suc)
+      also have "\<dots> \<le> inverse (ennreal 2) ^ Suc N"
+        using measure[of N]
+        by (cases "emeasure M (A N)"; cases "emeasure M (A N) = 0")
+           (auto simp: inverse_ennreal ennreal_mult[symmetric] divide_ennreal_def simp del: power_Suc)
+      also have "\<dots> = ennreal (inverse 2 ^ Suc N)"
+        by (subst ennreal_power[symmetric], simp) (simp add: inverse_ennreal)
+      finally show "n N * emeasure M (A N) \<le> ennreal ((1/2)^Suc N)"
+        by simp
+    qed auto
+    also have "\<dots> < top"
+      unfolding less_top[symmetric]
+      by (rule ennreal_suminf_neq_top)
+         (auto simp: summable_geometric summable_Suc_iff simp del: power_Suc)
+    finally show "integral\<^sup>N M ?h \<noteq> \<infinity>"
+      by (auto simp: top_unique)
+  next
+    { fix x assume "x \<in> space M"
+      then obtain i where "x \<in> A i" using space[symmetric] by auto
+      with disjoint n have "?h x = n i"
+        by (auto intro!: suminf_cmult_indicator intro: less_imp_le)
+      then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by (auto simp: less_top[symmetric]) }
+    note pos = this
+  qed measurable
+qed
+
+subsection "Absolutely continuous"
+
+definition absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
+  "absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N"
+
+lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M"
+  unfolding absolutely_continuous_def by (auto simp: null_sets_count_space)
+
+lemma absolutely_continuousI_density:
+  "f \<in> borel_measurable M \<Longrightarrow> absolutely_continuous M (density M f)"
+  by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in)
+
+lemma absolutely_continuousI_point_measure_finite:
+  "(\<And>x. \<lbrakk> x \<in> A ; f x \<le> 0 \<rbrakk> \<Longrightarrow> g x \<le> 0) \<Longrightarrow> absolutely_continuous (point_measure A f) (point_measure A g)"
+  unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff)
+
+lemma absolutely_continuousD:
+  "absolutely_continuous M N \<Longrightarrow> A \<in> sets M \<Longrightarrow> emeasure M A = 0 \<Longrightarrow> emeasure N A = 0"
+  by (auto simp: absolutely_continuous_def null_sets_def)
+
+lemma absolutely_continuous_AE:
+  assumes sets_eq: "sets M' = sets M"
+    and "absolutely_continuous M M'" "AE x in M. P x"
+   shows "AE x in M'. P x"
+proof -
+  from \<open>AE x in M. P x\<close> obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
+    unfolding eventually_ae_filter by auto
+  show "AE x in M'. P x"
+  proof (rule AE_I')
+    show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp
+    from \<open>absolutely_continuous M M'\<close> show "N \<in> null_sets M'"
+      using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto
+  qed
+qed
+
+subsection "Existence of the Radon-Nikodym derivative"
+
+lemma (in finite_measure) Radon_Nikodym_finite_measure:
+  assumes "finite_measure N" and sets_eq[simp]: "sets N = sets M"
+  assumes "absolutely_continuous M N"
+  shows "\<exists>f \<in> borel_measurable M. density M f = N"
+proof -
+  interpret N: finite_measure N by fact
+  define G where "G = {g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A}"
+  have [measurable_dest]: "f \<in> G \<Longrightarrow> f \<in> borel_measurable M"
+    and G_D: "\<And>A. f \<in> G \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) \<le> N A" for f
+    by (auto simp: G_def)
+  note this[measurable_dest]
+  have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
+  hence "G \<noteq> {}" by auto
+  { fix f g assume f[measurable]: "f \<in> G" and g[measurable]: "g \<in> G"
+    have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def
+    proof safe
+      let ?A = "{x \<in> space M. f x \<le> g x}"
+      have "?A \<in> sets M" using f g unfolding G_def by auto
+      fix A assume [measurable]: "A \<in> sets M"
+      have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
+        using sets.sets_into_space[OF \<open>A \<in> sets M\<close>] by auto
+      have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
+        g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
+        by (auto simp: indicator_def max_def)
+      hence "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) =
+        (\<integral>\<^sup>+x. g x * indicator (?A \<inter> A) x \<partial>M) +
+        (\<integral>\<^sup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
+        by (auto cong: nn_integral_cong intro!: nn_integral_add)
+      also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)"
+        using f g unfolding G_def by (auto intro!: add_mono)
+      also have "\<dots> = N A"
+        using union by (subst plus_emeasure) auto
+      finally show "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .
+    qed auto }
+  note max_in_G = this
+  { fix f assume  "incseq f" and f: "\<And>i. f i \<in> G"
+    then have [measurable]: "\<And>i. f i \<in> borel_measurable M" by (auto simp: G_def)
+    have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def
+    proof safe
+      show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M" by measurable
+    next
+      fix A assume "A \<in> sets M"
+      have "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) =
+        (\<integral>\<^sup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
+        by (intro nn_integral_cong) (simp split: split_indicator)
+      also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f i x * indicator A x \<partial>M))"
+        using \<open>incseq f\<close> f \<open>A \<in> sets M\<close>
+        by (intro nn_integral_monotone_convergence_SUP)
+           (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
+      finally show "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
+        using f \<open>A \<in> sets M\<close> by (auto intro!: SUP_least simp: G_D)
+    qed }
+  note SUP_in_G = this
+  let ?y = "SUP g : G. integral\<^sup>N M g"
+  have y_le: "?y \<le> N (space M)" unfolding G_def
+  proof (safe intro!: SUP_least)
+    fix g assume "\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A"
+    from this[THEN bspec, OF sets.top] show "integral\<^sup>N M g \<le> N (space M)"
+      by (simp cong: nn_integral_cong)
+  qed
+  from ennreal_SUP_countable_SUP [OF \<open>G \<noteq> {}\<close>, of "integral\<^sup>N M"] guess ys .. note ys = this
+  then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n"
+  proof safe
+    fix n assume "range ys \<subseteq> integral\<^sup>N M ` G"
+    hence "ys n \<in> integral\<^sup>N M ` G" by auto
+    thus "\<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n" by auto
+  qed
+  from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^sup>N M (gs n) = ys n" by auto
+  hence y_eq: "?y = (SUP i. integral\<^sup>N M (gs i))" using ys by auto
+  let ?g = "\<lambda>i x. Max ((\<lambda>n. gs n x) ` {..i})"
+  define f where [abs_def]: "f x = (SUP i. ?g i x)" for x
+  let ?F = "\<lambda>A x. f x * indicator A x"
+  have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
+  { fix i have "?g i \<in> G"
+    proof (induct i)
+      case 0 thus ?case by simp fact
+    next
+      case (Suc i)
+      with Suc gs_not_empty \<open>gs (Suc i) \<in> G\<close> show ?case
+        by (auto simp add: atMost_Suc intro!: max_in_G)
+    qed }
+  note g_in_G = this
+  have "incseq ?g" using gs_not_empty
+    by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc)
+
+  from SUP_in_G[OF this g_in_G] have [measurable]: "f \<in> G" unfolding f_def .
+  then have [measurable]: "f \<in> borel_measurable M" unfolding G_def by auto
+
+  have "integral\<^sup>N M f = (SUP i. integral\<^sup>N M (?g i))" unfolding f_def
+    using g_in_G \<open>incseq ?g\<close> by (auto intro!: nn_integral_monotone_convergence_SUP simp: G_def)
+  also have "\<dots> = ?y"
+  proof (rule antisym)
+    show "(SUP i. integral\<^sup>N M (?g i)) \<le> ?y"
+      using g_in_G by (auto intro: SUP_mono)
+    show "?y \<le> (SUP i. integral\<^sup>N M (?g i))" unfolding y_eq
+      by (auto intro!: SUP_mono nn_integral_mono Max_ge)
+  qed
+  finally have int_f_eq_y: "integral\<^sup>N M f = ?y" .
+
+  have upper_bound: "\<forall>A\<in>sets M. N A \<le> density M f A"
+  proof (rule ccontr)
+    assume "\<not> ?thesis"
+    then obtain A where A[measurable]: "A \<in> sets M" and f_less_N: "density M f A < N A"
+      by (auto simp: not_le)
+    then have pos_A: "0 < M A"
+      using \<open>absolutely_continuous M N\<close>[THEN absolutely_continuousD, OF A]
+      by (auto simp: zero_less_iff_neq_zero)
+
+    define b where "b = (N A - density M f A) / M A / 2"
+    with f_less_N pos_A have "0 < b" "b \<noteq> top"
+      by (auto intro!: diff_gr0_ennreal simp: zero_less_iff_neq_zero diff_eq_0_iff_ennreal ennreal_divide_eq_top_iff)
+
+    let ?f = "\<lambda>x. f x + b"
+    have "nn_integral M f \<noteq> top"
+      using `f \<in> G`[THEN G_D, of "space M"] by (auto simp: top_unique cong: nn_integral_cong)
+    with \<open>b \<noteq> top\<close> interpret Mf: finite_measure "density M ?f"
+      by (intro finite_measureI)
+         (auto simp: field_simps mult_indicator_subset ennreal_mult_eq_top_iff
+                     emeasure_density nn_integral_cmult_indicator nn_integral_add
+               cong: nn_integral_cong)
+
+    from unsigned_Hahn_decomposition[of "density M ?f" N A]
+    obtain Y where [measurable]: "Y \<in> sets M" and [simp]: "Y \<subseteq> A"
+       and Y1: "\<And>C. C \<in> sets M \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> density M ?f C \<le> N C"
+       and Y2: "\<And>C. C \<in> sets M \<Longrightarrow> C \<subseteq> A \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> N C \<le> density M ?f C"
+       by auto
+
+    let ?f' = "\<lambda>x. f x + b * indicator Y x"
+    have "M Y \<noteq> 0"
+    proof
+      assume "M Y = 0"
+      then have "N Y = 0"
+        using \<open>absolutely_continuous M N\<close>[THEN absolutely_continuousD, of Y] by auto
+      then have "N A = N (A - Y)"
+        by (subst emeasure_Diff) auto
+      also have "\<dots> \<le> density M ?f (A - Y)"
+        by (rule Y2) auto
+      also have "\<dots> \<le> density M ?f A - density M ?f Y"
+        by (subst emeasure_Diff) auto
+      also have "\<dots> \<le> density M ?f A - 0"
+        by (intro ennreal_minus_mono) auto
+      also have "density M ?f A = b * M A + density M f A"
+        by (simp add: emeasure_density field_simps mult_indicator_subset nn_integral_add nn_integral_cmult_indicator)
+      also have "\<dots> < N A"
+        using f_less_N pos_A
+        by (cases "density M f A"; cases "M A"; cases "N A")
+           (auto simp: b_def ennreal_less_iff ennreal_minus divide_ennreal ennreal_numeral[symmetric]
+                       ennreal_plus[symmetric] ennreal_mult[symmetric] field_simps
+                    simp del: ennreal_numeral ennreal_plus)
+      finally show False
+        by simp
+    qed
+    then have "nn_integral M f < nn_integral M ?f'"
+      using \<open>0 < b\<close> \<open>nn_integral M f \<noteq> top\<close>
+      by (simp add: nn_integral_add nn_integral_cmult_indicator ennreal_zero_less_mult_iff zero_less_iff_neq_zero)
+    moreover
+    have "?f' \<in> G"
+      unfolding G_def
+    proof safe
+      fix X assume [measurable]: "X \<in> sets M"
+      have "(\<integral>\<^sup>+ x. ?f' x * indicator X x \<partial>M) = density M f (X - Y) + density M ?f (X \<inter> Y)"
+        by (auto simp add: emeasure_density nn_integral_add[symmetric] split: split_indicator intro!: nn_integral_cong)
+      also have "\<dots> \<le> N (X - Y) + N (X \<inter> Y)"
+        using G_D[OF \<open>f \<in> G\<close>] by (intro add_mono Y1) (auto simp: emeasure_density)
+      also have "\<dots> = N X"
+        by (subst plus_emeasure) (auto intro!: arg_cong2[where f=emeasure])
+      finally show "(\<integral>\<^sup>+ x. ?f' x * indicator X x \<partial>M) \<le> N X" .
+    qed simp
+    then have "nn_integral M ?f' \<le> ?y"
+      by (rule SUP_upper)
+    ultimately show False
+      by (simp add: int_f_eq_y)
+  qed
+  show ?thesis
+  proof (intro bexI[of _ f] measure_eqI conjI antisym)
+    fix A assume "A \<in> sets (density M f)" then show "emeasure (density M f) A \<le> emeasure N A"
+      by (auto simp: emeasure_density intro!: G_D[OF \<open>f \<in> G\<close>])
+  next
+    fix A assume A: "A \<in> sets (density M f)" then show "emeasure N A \<le> emeasure (density M f) A"
+      using upper_bound by auto
+  qed auto
+qed
+
+lemma (in finite_measure) split_space_into_finite_sets_and_rest:
+  assumes ac: "absolutely_continuous M N" and sets_eq[simp]: "sets N = sets M"
+  shows "\<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> (\<forall>i. N (B i) \<noteq> \<infinity>) \<and>
+    (\<forall>A\<in>sets M. A \<inter> (\<Union>i. B i) = {} \<longrightarrow> (emeasure M A = 0 \<and> N A = 0) \<or> (emeasure M A > 0 \<and> N A = \<infinity>))"
+proof -
+  let ?Q = "{Q\<in>sets M. N Q \<noteq> \<infinity>}"
+  let ?a = "SUP Q:?Q. emeasure M Q"
+  have "{} \<in> ?Q" by auto
+  then have Q_not_empty: "?Q \<noteq> {}" by blast
+  have "?a \<le> emeasure M (space M)" using sets.sets_into_space
+    by (auto intro!: SUP_least emeasure_mono)
+  then have "?a \<noteq> \<infinity>"
+    using finite_emeasure_space
+    by (auto simp: less_top[symmetric] top_unique simp del: SUP_eq_top_iff Sup_eq_top_iff)
+  from ennreal_SUP_countable_SUP [OF Q_not_empty, of "emeasure M"]
+  obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
+    by auto
+  then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto
+  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = emeasure M (Q' i)" "\<And>i. Q' i \<in> ?Q"
+    by auto
+  then have a_Lim: "?a = (SUP i::nat. emeasure M (Q' i))" using a by simp
+  let ?O = "\<lambda>n. \<Union>i\<le>n. Q' i"
+  have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\<Union>i. ?O i)"
+  proof (rule SUP_emeasure_incseq[of ?O])
+    show "range ?O \<subseteq> sets M" using Q' by auto
+    show "incseq ?O" by (fastforce intro!: incseq_SucI)
+  qed
+  have Q'_sets[measurable]: "\<And>i. Q' i \<in> sets M" using Q' by auto
+  have O_sets: "\<And>i. ?O i \<in> sets M" using Q' by auto
+  then have O_in_G: "\<And>i. ?O i \<in> ?Q"
+  proof (safe del: notI)
+    fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto
+    then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))"
+      by (simp add: emeasure_subadditive_finite)
+    also have "\<dots> < \<infinity>" using Q' by (simp add: less_top)
+    finally show "N (?O i) \<noteq> \<infinity>" by simp
+  qed auto
+  have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce
+  have a_eq: "?a = emeasure M (\<Union>i. ?O i)" unfolding Union[symmetric]
+  proof (rule antisym)
+    show "?a \<le> (SUP i. emeasure M (?O i))" unfolding a_Lim
+      using Q' by (auto intro!: SUP_mono emeasure_mono)
+    show "(SUP i. emeasure M (?O i)) \<le> ?a"
+    proof (safe intro!: Sup_mono, unfold bex_simps)
+      fix i
+      have *: "(\<Union>(Q' ` {..i})) = ?O i" by auto
+      then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and>
+        emeasure M (\<Union>(Q' ` {..i})) \<le> emeasure M x"
+        using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
+    qed
+  qed
+  let ?O_0 = "(\<Union>i. ?O i)"
+  have "?O_0 \<in> sets M" using Q' by auto
+  have "disjointed Q' i \<in> sets M" for i
+    using sets.range_disjointed_sets[of Q' M] using Q'_sets by (auto simp: subset_eq)
+  note Q_sets = this
+  show ?thesis
+  proof (intro bexI exI conjI ballI impI allI)
+    show "disjoint_family (disjointed Q')"
+      by (rule disjoint_family_disjointed)
+    show "range (disjointed Q') \<subseteq> sets M"
+      using Q'_sets by (intro sets.range_disjointed_sets) auto
+    { fix A assume A: "A \<in> sets M" "A \<inter> (\<Union>i. disjointed Q' i) = {}"
+      then have A1: "A \<inter> (\<Union>i. Q' i) = {}"
+        unfolding UN_disjointed_eq by auto
+      show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
+      proof (rule disjCI, simp)
+        assume *: "emeasure M A = 0 \<or> N A \<noteq> top"
+        show "emeasure M A = 0 \<and> N A = 0"
+        proof (cases "emeasure M A = 0")
+          case True
+          with ac A have "N A = 0"
+            unfolding absolutely_continuous_def by auto
+          with True show ?thesis by simp
+        next
+          case False
+          with * have "N A \<noteq> \<infinity>" by auto
+          with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)"
+            using Q' A1 by (auto intro!: plus_emeasure simp: set_eq_iff)
+          also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))"
+          proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
+            show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
+              using \<open>N A \<noteq> \<infinity>\<close> O_sets A by auto
+          qed (fastforce intro!: incseq_SucI)
+          also have "\<dots> \<le> ?a"
+          proof (safe intro!: SUP_least)
+            fix i have "?O i \<union> A \<in> ?Q"
+            proof (safe del: notI)
+              show "?O i \<union> A \<in> sets M" using O_sets A by auto
+              from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A"
+                using emeasure_subadditive[of "?O i" N A] A O_sets by auto
+              with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>"
+                using \<open>N A \<noteq> \<infinity>\<close> by (auto simp: top_unique)
+            qed
+            then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper)
+          qed
+          finally have "emeasure M A = 0"
+            unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure)
+          with \<open>emeasure M A \<noteq> 0\<close> show ?thesis by auto
+        qed
+      qed }
+    { fix i
+      have "N (disjointed Q' i) \<le> N (Q' i)"
+        by (auto intro!: emeasure_mono simp: disjointed_def)
+      then show "N (disjointed Q' i) \<noteq> \<infinity>"
+        using Q'(2)[of i] by (auto simp: top_unique) }
+  qed
+qed
+
+lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
+  assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M"
+  shows "\<exists>f\<in>borel_measurable M. density M f = N"
+proof -
+  from split_space_into_finite_sets_and_rest[OF assms]
+  obtain Q :: "nat \<Rightarrow> 'a set"
+    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
+    and Q_fin: "\<And>i. N (Q i) \<noteq> \<infinity>" by force
+  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
+  let ?N = "\<lambda>i. density N (indicator (Q i))" and ?M = "\<lambda>i. density M (indicator (Q i))"
+  have "\<forall>i. \<exists>f\<in>borel_measurable (?M i). density (?M i) f = ?N i"
+  proof (intro allI finite_measure.Radon_Nikodym_finite_measure)
+    fix i
+    from Q show "finite_measure (?M i)"
+      by (auto intro!: finite_measureI cong: nn_integral_cong
+               simp add: emeasure_density subset_eq sets_eq)
+    from Q have "emeasure (?N i) (space N) = emeasure N (Q i)"
+      by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: nn_integral_cong)
+    with Q_fin show "finite_measure (?N i)"
+      by (auto intro!: finite_measureI)
+    show "sets (?N i) = sets (?M i)" by (simp add: sets_eq)
+    have [measurable]: "\<And>A. A \<in> sets M \<Longrightarrow> A \<in> sets N" by (simp add: sets_eq)
+    show "absolutely_continuous (?M i) (?N i)"
+      using \<open>absolutely_continuous M N\<close> \<open>Q i \<in> sets M\<close>
+      by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq
+               intro!: absolutely_continuous_AE[OF sets_eq])
+  qed
+  from choice[OF this[unfolded Bex_def]]
+  obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
+    and f_density: "\<And>i. density (?M i) (f i) = ?N i"
+    by force
+  { fix A i assume A: "A \<in> sets M"
+    with Q borel have "(\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"
+      by (auto simp add: emeasure_density nn_integral_density subset_eq
+               intro!: nn_integral_cong split: split_indicator)
+    also have "\<dots> = emeasure N (Q i \<inter> A)"
+      using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq)
+    finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }
+  note integral_eq = this
+  let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator (space M - (\<Union>i. Q i)) x"
+  show ?thesis
+  proof (safe intro!: bexI[of _ ?f])
+    show "?f \<in> borel_measurable M" using borel Q_sets
+      by (auto intro!: measurable_If)
+    show "density M ?f = N"
+    proof (rule measure_eqI)
+      fix A assume "A \<in> sets (density M ?f)"
+      then have "A \<in> sets M" by simp
+      have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
+      have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
+        "\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
+        using borel Qi \<open>A \<in> sets M\<close> by auto
+      have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator ((space M - (\<Union>i. Q i)) \<inter> A) x \<partial>M)"
+        using borel by (intro nn_integral_cong) (auto simp: indicator_def)
+      also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)"
+        using borel Qi \<open>A \<in> sets M\<close>
+        by (subst nn_integral_add)
+           (auto simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le)
+      also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)"
+        by (subst integral_eq[OF \<open>A \<in> sets M\<close>], subst nn_integral_suminf) auto
+      finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)" .
+      moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"
+        using Q Q_sets \<open>A \<in> sets M\<close>
+        by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)
+      moreover
+      have "(space M - (\<Union>x. Q x)) \<inter> A \<inter> (\<Union>x. Q x) = {}"
+        by auto
+      then have "\<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A) = N ((space M - (\<Union>i. Q i)) \<inter> A)"
+        using in_Q0[of "(space M - (\<Union>i. Q i)) \<inter> A"] \<open>A \<in> sets M\<close> Q by (auto simp: ennreal_top_mult)
+      moreover have "(space M - (\<Union>i. Q i)) \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
+        using Q_sets \<open>A \<in> sets M\<close> by auto
+      moreover have "((\<Union>i. Q i) \<inter> A) \<union> ((space M - (\<Union>i. Q i)) \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> ((space M - (\<Union>i. Q i)) \<inter> A) = {}"
+        using \<open>A \<in> sets M\<close> sets.sets_into_space by auto
+      ultimately have "N A = (\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M)"
+        using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "(space M - (\<Union>i. Q i)) \<inter> A"] by (simp add: sets_eq)
+      with \<open>A \<in> sets M\<close> borel Q show "emeasure (density M ?f) A = N A"
+        by (auto simp: subset_eq emeasure_density)
+    qed (simp add: sets_eq)
+  qed
+qed
+
+lemma (in sigma_finite_measure) Radon_Nikodym:
+  assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M"
+  shows "\<exists>f \<in> borel_measurable M. density M f = N"
+proof -
+  from Ex_finite_integrable_function
+  obtain h where finite: "integral\<^sup>N M h \<noteq> \<infinity>" and
+    borel: "h \<in> borel_measurable M" and
+    nn: "\<And>x. 0 \<le> h x" and
+    pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
+    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto
+  let ?T = "\<lambda>A. (\<integral>\<^sup>+x. h x * indicator A x \<partial>M)"
+  let ?MT = "density M h"
+  from borel finite nn interpret T: finite_measure ?MT
+    by (auto intro!: finite_measureI cong: nn_integral_cong simp: emeasure_density)
+  have "absolutely_continuous ?MT N" "sets N = sets ?MT"
+  proof (unfold absolutely_continuous_def, safe)
+    fix A assume "A \<in> null_sets ?MT"
+    with borel have "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> h x \<le> 0"
+      by (auto simp add: null_sets_density_iff)
+    with pos sets.sets_into_space have "AE x in M. x \<notin> A"
+      by (elim eventually_mono) (auto simp: not_le[symmetric])
+    then have "A \<in> null_sets M"
+      using \<open>A \<in> sets M\<close> by (simp add: AE_iff_null_sets)
+    with ac show "A \<in> null_sets N"
+      by (auto simp: absolutely_continuous_def)
+  qed (auto simp add: sets_eq)
+  from T.Radon_Nikodym_finite_measure_infinite[OF this]
+  obtain f where f_borel: "f \<in> borel_measurable M" "density ?MT f = N" by auto
+  with nn borel show ?thesis
+    by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)
+qed
+
+subsection \<open>Uniqueness of densities\<close>
+
+lemma finite_density_unique:
+  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
+  and fin: "integral\<^sup>N M f \<noteq> \<infinity>"
+  shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"
+proof (intro iffI ballI)
+  fix A assume eq: "AE x in M. f x = g x"
+  with borel show "density M f = density M g"
+    by (auto intro: density_cong)
+next
+  let ?P = "\<lambda>f A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M"
+  assume "density M f = density M g"
+  with borel have eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+    by (simp add: emeasure_density[symmetric])
+  from this[THEN bspec, OF sets.top] fin
+  have g_fin: "integral\<^sup>N M g \<noteq> \<infinity>" by (simp cong: nn_integral_cong)
+  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+      and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
+      and g_fin: "integral\<^sup>N M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+    let ?N = "{x\<in>space M. g x < f x}"
+    have N: "?N \<in> sets M" using borel by simp
+    have "?P g ?N \<le> integral\<^sup>N M g" using pos
+      by (intro nn_integral_mono_AE) (auto split: split_indicator)
+    then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by (auto simp: top_unique)
+    have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
+      by (auto intro!: nn_integral_cong simp: indicator_def)
+    also have "\<dots> = ?P f ?N - ?P g ?N"
+    proof (rule nn_integral_diff)
+      show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
+        using borel N by auto
+      show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x"
+        using pos by (auto split: split_indicator)
+    qed fact
+    also have "\<dots> = 0"
+      unfolding eq[THEN bspec, OF N] using Pg_fin by auto
+    finally have "AE x in M. f x \<le> g x"
+      using pos borel nn_integral_PInf_AE[OF borel(2) g_fin]
+      by (subst (asm) nn_integral_0_iff_AE)
+         (auto split: split_indicator simp: not_less ennreal_minus_eq_0) }
+  from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
+  show "AE x in M. f x = g x" by auto
+qed
+
+lemma (in finite_measure) density_unique_finite_measure:
+  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
+  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x"
+  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. f' x * indicator A x \<partial>M)"
+    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
+  shows "AE x in M. f x = f' x"
+proof -
+  let ?D = "\<lambda>f. density M f"
+  let ?N = "\<lambda>A. ?P f A" and ?N' = "\<lambda>A. ?P f' A"
+  let ?f = "\<lambda>A x. f x * indicator A x" and ?f' = "\<lambda>A x. f' x * indicator A x"
+
+  have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M"
+    using borel by (auto intro!: absolutely_continuousI_density)
+  from split_space_into_finite_sets_and_rest[OF this]
+  obtain Q :: "nat \<Rightarrow> 'a set"
+    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> ?D f A = 0 \<or> 0 < emeasure M A \<and> ?D f A = \<infinity>"
+    and Q_fin: "\<And>i. ?D f (Q i) \<noteq> \<infinity>" by force
+  with borel pos have in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> ?N A = 0 \<or> 0 < emeasure M A \<and> ?N A = \<infinity>"
+    and Q_fin: "\<And>i. ?N (Q i) \<noteq> \<infinity>" by (auto simp: emeasure_density subset_eq)
+
+  from Q have Q_sets[measurable]: "\<And>i. Q i \<in> sets M" by auto
+  let ?D = "{x\<in>space M. f x \<noteq> f' x}"
+  have "?D \<in> sets M" using borel by auto
+  have *: "\<And>i x A. \<And>y::ennreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
+    unfolding indicator_def by auto
+  have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
+    by (intro finite_density_unique[THEN iffD1] allI)
+       (auto intro!: f measure_eqI simp: emeasure_density * subset_eq)
+  moreover have "AE x in M. ?f (space M - (\<Union>i. Q i)) x = ?f' (space M - (\<Union>i. Q i)) x"
+  proof (rule AE_I')
+    { fix f :: "'a \<Rightarrow> ennreal" assume borel: "f \<in> borel_measurable M"
+        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"
+      let ?A = "\<lambda>i. (space M - (\<Union>i. Q i)) \<inter> {x \<in> space M. f x < (i::nat)}"
+      have "(\<Union>i. ?A i) \<in> null_sets M"
+      proof (rule null_sets_UN)
+        fix i ::nat have "?A i \<in> sets M"
+          using borel by auto
+        have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ennreal) * indicator (?A i) x \<partial>M)"
+          unfolding eq[OF \<open>?A i \<in> sets M\<close>]
+          by (auto intro!: nn_integral_mono simp: indicator_def)
+        also have "\<dots> = i * emeasure M (?A i)"
+          using \<open>?A i \<in> sets M\<close> by (auto intro!: nn_integral_cmult_indicator)
+        also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by (auto simp: ennreal_mult_less_top of_nat_less_top)
+        finally have "?N (?A i) \<noteq> \<infinity>" by simp
+        then show "?A i \<in> null_sets M" using in_Q0[OF \<open>?A i \<in> sets M\<close>] \<open>?A i \<in> sets M\<close> by auto
+      qed
+      also have "(\<Union>i. ?A i) = (space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}"
+        by (auto simp: ennreal_Ex_less_of_nat less_top[symmetric])
+      finally have "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp }
+    from this[OF borel(1) refl] this[OF borel(2) f]
+    have "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all
+    then show "((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets M" by (rule null_sets.Un)
+    show "{x \<in> space M. ?f (space M - (\<Union>i. Q i)) x \<noteq> ?f' (space M - (\<Union>i. Q i)) x} \<subseteq>
+      ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def)
+  qed
+  moreover have "AE x in M. (?f (space M - (\<Union>i. Q i)) x = ?f' (space M - (\<Union>i. Q i)) x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
+    ?f (space M) x = ?f' (space M) x"
+    by (auto simp: indicator_def)
+  ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x"
+    unfolding AE_all_countable[symmetric]
+    by eventually_elim (auto split: if_split_asm simp: indicator_def)
+  then show "AE x in M. f x = f' x" by auto
+qed
+
+lemma (in sigma_finite_measure) density_unique:
+  assumes f: "f \<in> borel_measurable M"
+  assumes f': "f' \<in> borel_measurable M"
+  assumes density_eq: "density M f = density M f'"
+  shows "AE x in M. f x = f' x"
+proof -
+  obtain h where h_borel: "h \<in> borel_measurable M"
+    and fin: "integral\<^sup>N M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
+    using Ex_finite_integrable_function by auto
+  then have h_nn: "AE x in M. 0 \<le> h x" by auto
+  let ?H = "density M h"
+  interpret h: finite_measure ?H
+    using fin h_borel pos
+    by (intro finite_measureI) (simp cong: nn_integral_cong emeasure_density add: fin)
+  let ?fM = "density M f"
+  let ?f'M = "density M f'"
+  { fix A assume "A \<in> sets M"
+    then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
+      using pos(1) sets.sets_into_space by (force simp: indicator_def)
+    then have "(\<integral>\<^sup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
+      using h_borel \<open>A \<in> sets M\<close> h_nn by (subst nn_integral_0_iff) auto }
+  note h_null_sets = this
+  { fix A assume "A \<in> sets M"
+    have "(\<integral>\<^sup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?fM)"
+      using \<open>A \<in> sets M\<close> h_borel h_nn f f'
+      by (intro nn_integral_density[symmetric]) auto
+    also have "\<dots> = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?f'M)"
+      by (simp_all add: density_eq)
+    also have "\<dots> = (\<integral>\<^sup>+x. f' x * (h x * indicator A x) \<partial>M)"
+      using \<open>A \<in> sets M\<close> h_borel h_nn f f'
+      by (intro nn_integral_density) auto
+    finally have "(\<integral>\<^sup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * (f' x * indicator A x) \<partial>M)"
+      by (simp add: ac_simps)
+    then have "(\<integral>\<^sup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^sup>+x. (f' x * indicator A x) \<partial>?H)"
+      using \<open>A \<in> sets M\<close> h_borel h_nn f f'
+      by (subst (asm) (1 2) nn_integral_density[symmetric]) auto }
+  then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
+    by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) auto
+  with AE_space[of M] pos show "AE x in M. f x = f' x"
+    unfolding AE_density[OF h_borel] by auto
+qed
+
+lemma (in sigma_finite_measure) density_unique_iff:
+  assumes f: "f \<in> borel_measurable M" and f': "f' \<in> borel_measurable M"
+  shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)"
+  using density_unique[OF assms] density_cong[OF f f'] by auto
+
+lemma sigma_finite_density_unique:
+  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  and fin: "sigma_finite_measure (density M f)"
+  shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"
+proof
+  assume "AE x in M. f x = g x" with borel show "density M f = density M g"
+    by (auto intro: density_cong)
+next
+  assume eq: "density M f = density M g"
+  interpret f: sigma_finite_measure "density M f" by fact
+  from f.sigma_finite_incseq guess A . note cover = this
+
+  have "AE x in M. \<forall>i. x \<in> A i \<longrightarrow> f x = g x"
+    unfolding AE_all_countable
+  proof
+    fix i
+    have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))"
+      unfolding eq ..
+    moreover have "(\<integral>\<^sup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>"
+      using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq)
+    ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x"
+      using borel cover(1)
+      by (intro finite_density_unique[THEN iffD1]) (auto simp: density_density_eq subset_eq)
+    then show "AE x in M. x \<in> A i \<longrightarrow> f x = g x"
+      by auto
+  qed
+  with AE_space show "AE x in M. f x = g x"
+    apply eventually_elim
+    using cover(2)[symmetric]
+    apply auto
+    done
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_iff_density_finite':
+  assumes f: "f \<in> borel_measurable M"
+  shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
+    (is "sigma_finite_measure ?N \<longleftrightarrow> _")
+proof
+  assume "sigma_finite_measure ?N"
+  then interpret N: sigma_finite_measure ?N .
+  from N.Ex_finite_integrable_function obtain h where
+    h: "h \<in> borel_measurable M" "integral\<^sup>N ?N h \<noteq> \<infinity>" and
+    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>"
+    by auto
+  have "AE x in M. f x * h x \<noteq> \<infinity>"
+  proof (rule AE_I')
+    have "integral\<^sup>N ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)"
+      using f h by (auto intro!: nn_integral_density)
+    then have "(\<integral>\<^sup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
+      using h(2) by simp
+    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
+      using f h(1) by (auto intro!: nn_integral_PInf[unfolded infinity_ennreal_def] borel_measurable_vimage)
+  qed auto
+  then show "AE x in M. f x \<noteq> \<infinity>"
+    using fin by (auto elim!: AE_Ball_mp simp: less_top ennreal_mult_less_top)
+next
+  assume AE: "AE x in M. f x \<noteq> \<infinity>"
+  from sigma_finite guess Q . note Q = this
+  define A where "A i =
+    f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ennreal(of_nat (Suc n))}) \<inter> space M" for i
+  { fix i j have "A i \<inter> Q j \<in> sets M"
+    unfolding A_def using f Q
+    apply (rule_tac sets.Int)
+    by (cases i) (auto intro: measurable_sets[OF f(1)]) }
+  note A_in_sets = this
+
+  show "sigma_finite_measure ?N"
+  proof (standard, intro exI conjI ballI)
+    show "countable (range (\<lambda>(i, j). A i \<inter> Q j))"
+      by auto
+    show "range (\<lambda>(i, j). A i \<inter> Q j) \<subseteq> sets (density M f)"
+      using A_in_sets by auto
+  next
+    have "\<Union>range (\<lambda>(i, j). A i \<inter> Q j) = (\<Union>i j. A i \<inter> Q j)"
+      by auto
+    also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto
+    also have "(\<Union>i. A i) = space M"
+    proof safe
+      fix x assume x: "x \<in> space M"
+      show "x \<in> (\<Union>i. A i)"
+      proof (cases "f x" rule: ennreal_cases)
+        case top with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
+      next
+        case (real r)
+        with ennreal_Ex_less_of_nat[of "f x"] obtain n :: nat where "f x < n"
+          by auto
+        also have "n < (Suc n :: ennreal)"
+          by simp
+        finally show ?thesis
+          using x real by (auto simp: A_def ennreal_of_nat_eq_real_of_nat intro!: exI[of _ "Suc n"])
+      qed
+    qed (auto simp: A_def)
+    finally show "\<Union>range (\<lambda>(i, j). A i \<inter> Q j) = space ?N" by simp
+  next
+    fix X assume "X \<in> range (\<lambda>(i, j). A i \<inter> Q j)"
+    then obtain i j where [simp]:"X = A i \<inter> Q j" by auto
+    have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
+    proof (cases i)
+      case 0
+      have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0"
+        using AE by (auto simp: A_def \<open>i = 0\<close>)
+      from nn_integral_cong_AE[OF this] show ?thesis by simp
+    next
+      case (Suc n)
+      then have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
+        (\<integral>\<^sup>+x. (Suc n :: ennreal) * indicator (Q j) x \<partial>M)"
+        by (auto intro!: nn_integral_mono simp: indicator_def A_def ennreal_of_nat_eq_real_of_nat)
+      also have "\<dots> = Suc n * emeasure M (Q j)"
+        using Q by (auto intro!: nn_integral_cmult_indicator)
+      also have "\<dots> < \<infinity>"
+        using Q by (auto simp: ennreal_mult_less_top less_top of_nat_less_top)
+      finally show ?thesis by simp
+    qed
+    then show "emeasure ?N X \<noteq> \<infinity>"
+      using A_in_sets Q f by (auto simp: emeasure_density)
+  qed
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
+  "f \<in> borel_measurable M \<Longrightarrow> sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
+  by (subst sigma_finite_iff_density_finite')
+     (auto simp: max_def intro!: measurable_If)
+
+subsection \<open>Radon-Nikodym derivative\<close>
+
+definition RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ennreal" where
+  "RN_deriv M N =
+    (if \<exists>f. f \<in> borel_measurable M \<and> density M f = N
+       then SOME f. f \<in> borel_measurable M \<and> density M f = N
+       else (\<lambda>_. 0))"
+
+lemma RN_derivI:
+  assumes "f \<in> borel_measurable M" "density M f = N"
+  shows "density M (RN_deriv M N) = N"
+proof -
+  have *: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
+    using assms by auto
+  then have "density M (SOME f. f \<in> borel_measurable M \<and> density M f = N) = N"
+    by (rule someI2_ex) auto
+  with * show ?thesis
+    by (auto simp: RN_deriv_def)
+qed
+
+lemma borel_measurable_RN_deriv[measurable]: "RN_deriv M N \<in> borel_measurable M"
+proof -
+  { assume ex: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
+    have 1: "(SOME f. f \<in> borel_measurable M \<and> density M f = N) \<in> borel_measurable M"
+      using ex by (rule someI2_ex) auto }
+  from this show ?thesis
+    by (auto simp: RN_deriv_def)
+qed
+
+lemma density_RN_deriv_density:
+  assumes f: "f \<in> borel_measurable M"
+  shows "density M (RN_deriv M (density M f)) = density M f"
+  by (rule RN_derivI[OF f]) simp
+
+lemma (in sigma_finite_measure) density_RN_deriv:
+  "absolutely_continuous M N \<Longrightarrow> sets N = sets M \<Longrightarrow> density M (RN_deriv M N) = N"
+  by (metis RN_derivI Radon_Nikodym)
+
+lemma (in sigma_finite_measure) RN_deriv_nn_integral:
+  assumes N: "absolutely_continuous M N" "sets N = sets M"
+    and f: "f \<in> borel_measurable M"
+  shows "integral\<^sup>N N f = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
+proof -
+  have "integral\<^sup>N N f = integral\<^sup>N (density M (RN_deriv M N)) f"
+    using N by (simp add: density_RN_deriv)
+  also have "\<dots> = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
+    using f by (simp add: nn_integral_density)
+  finally show ?thesis by simp
+qed
+
+lemma null_setsD_AE: "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
+  using AE_iff_null_sets[of N M] by auto
+
+lemma (in sigma_finite_measure) RN_deriv_unique:
+  assumes f: "f \<in> borel_measurable M"
+  and eq: "density M f = N"
+  shows "AE x in M. f x = RN_deriv M N x"
+  unfolding eq[symmetric]
+  by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv
+            density_RN_deriv_density[symmetric])
+
+lemma RN_deriv_unique_sigma_finite:
+  assumes f: "f \<in> borel_measurable M"
+  and eq: "density M f = N" and fin: "sigma_finite_measure N"
+  shows "AE x in M. f x = RN_deriv M N x"
+  using fin unfolding eq[symmetric]
+  by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv
+            density_RN_deriv_density[symmetric])
+
+lemma (in sigma_finite_measure) RN_deriv_distr:
+  fixes T :: "'a \<Rightarrow> 'b"
+  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
+    and inv: "\<forall>x\<in>space M. T' (T x) = x"
+  and ac[simp]: "absolutely_continuous (distr M M' T) (distr N M' T)"
+  and N: "sets N = sets M"
+  shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x"
+proof (rule RN_deriv_unique)
+  have [simp]: "sets N = sets M" by fact
+  note sets_eq_imp_space_eq[OF N, simp]
+  have measurable_N[simp]: "\<And>M'. measurable N M' = measurable M M'" by (auto simp: measurable_def)
+  { fix A assume "A \<in> sets M"
+    with inv T T' sets.sets_into_space[OF this]
+    have "T -` T' -` A \<inter> T -` space M' \<inter> space M = A"
+      by (auto simp: measurable_def) }
+  note eq = this[simp]
+  { fix A assume "A \<in> sets M"
+    with inv T T' sets.sets_into_space[OF this]
+    have "(T' \<circ> T) -` A \<inter> space M = A"
+      by (auto simp: measurable_def) }
+  note eq2 = this[simp]
+  let ?M' = "distr M M' T" and ?N' = "distr N M' T"
+  interpret M': sigma_finite_measure ?M'
+  proof
+    from sigma_finite_countable guess F .. note F = this
+    show "\<exists>A. countable A \<and> A \<subseteq> sets (distr M M' T) \<and> \<Union>A = space (distr M M' T) \<and> (\<forall>a\<in>A. emeasure (distr M M' T) a \<noteq> \<infinity>)"
+    proof (intro exI conjI ballI)
+      show *: "(\<lambda>A. T' -` A \<inter> space ?M') ` F \<subseteq> sets ?M'"
+        using F T' by (auto simp: measurable_def)
+      show "\<Union>((\<lambda>A. T' -` A \<inter> space ?M')`F) = space ?M'"
+        using F T'[THEN measurable_space] by (auto simp: set_eq_iff)
+    next
+      fix X assume "X \<in> (\<lambda>A. T' -` A \<inter> space ?M')`F"
+      then obtain A where [simp]: "X = T' -` A \<inter> space ?M'" and "A \<in> F" by auto
+      have "X \<in> sets M'" using F T' \<open>A\<in>F\<close> by auto
+      moreover
+      have Fi: "A \<in> sets M" using F \<open>A\<in>F\<close> by auto
+      ultimately show "emeasure ?M' X \<noteq> \<infinity>"
+        using F T T' \<open>A\<in>F\<close> by (simp add: emeasure_distr)
+    qed (insert F, auto)
+  qed
+  have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M"
+    using T ac by measurable
+  then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M"
+    by (simp add: comp_def)
+
+  have "N = distr N M (T' \<circ> T)"
+    by (subst measure_of_of_measure[of N, symmetric])
+       (auto simp add: distr_def sets.sigma_sets_eq intro!: measure_of_eq sets.space_closed)
+  also have "\<dots> = distr (distr N M' T) M T'"
+    using T T' by (simp add: distr_distr)
+  also have "\<dots> = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'"
+    using ac by (simp add: M'.density_RN_deriv)
+  also have "\<dots> = density M (RN_deriv (distr M M' T) (distr N M' T) \<circ> T)"
+    by (simp add: distr_density_distr[OF T T', OF inv])
+  finally show "density M (\<lambda>x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N"
+    by (simp add: comp_def)
+qed
+
+lemma (in sigma_finite_measure) RN_deriv_finite:
+  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
+  shows "AE x in M. RN_deriv M N x \<noteq> \<infinity>"
+proof -
+  interpret N: sigma_finite_measure N by fact
+  from N show ?thesis
+    using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N] density_RN_deriv[OF ac]
+    by simp
+qed
+
+lemma (in sigma_finite_measure)
+  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
+    and f: "f \<in> borel_measurable M"
+  shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>
+      integrable M (\<lambda>x. enn2real (RN_deriv M N x) * f x)" (is ?integrable)
+    and RN_deriv_integral: "integral\<^sup>L N f = (\<integral>x. enn2real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
+proof -
+  note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]
+  interpret N: sigma_finite_measure N by fact
+
+  have eq: "density M (RN_deriv M N) = density M (\<lambda>x. enn2real (RN_deriv M N x))"
+  proof (rule density_cong)
+    from RN_deriv_finite[OF assms(1,2,3)]
+    show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
+      by eventually_elim (auto simp: less_top)
+  qed (insert ac, auto)
+
+  show ?integrable
+    apply (subst density_RN_deriv[OF ac, symmetric])
+    unfolding eq
+    apply (intro integrable_real_density f AE_I2 enn2real_nonneg)
+    apply (insert ac, auto)
+    done
+
+  show ?integral
+    apply (subst density_RN_deriv[OF ac, symmetric])
+    unfolding eq
+    apply (intro integral_real_density f AE_I2 enn2real_nonneg)
+    apply (insert ac, auto)
+    done
+qed
+
+lemma (in sigma_finite_measure) real_RN_deriv:
+  assumes "finite_measure N"
+  assumes ac: "absolutely_continuous M N" "sets N = sets M"
+  obtains D where "D \<in> borel_measurable M"
+    and "AE x in M. RN_deriv M N x = ennreal (D x)"
+    and "AE x in N. 0 < D x"
+    and "\<And>x. 0 \<le> D x"
+proof
+  interpret N: finite_measure N by fact
+
+  note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac]
+
+  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}"
+
+  show "(\<lambda>x. enn2real (RN_deriv M N x)) \<in> borel_measurable M"
+    using RN by auto
+
+  have "N (?RN \<infinity>) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
+    using RN(1) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
+  also have "\<dots> = (\<integral>\<^sup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
+    by (intro nn_integral_cong) (auto simp: indicator_def)
+  also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"
+    using RN by (intro nn_integral_cmult_indicator) auto
+  finally have eq: "N (?RN \<infinity>) = \<infinity> * emeasure M (?RN \<infinity>)" .
+  moreover
+  have "emeasure M (?RN \<infinity>) = 0"
+  proof (rule ccontr)
+    assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0"
+    then have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}"
+      by (auto simp: zero_less_iff_neq_zero)
+    with eq have "N (?RN \<infinity>) = \<infinity>" by (simp add: ennreal_mult_eq_top_iff)
+    with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto
+  qed
+  ultimately have "AE x in M. RN_deriv M N x < \<infinity>"
+    using RN by (intro AE_iff_measurable[THEN iffD2]) (auto simp: less_top[symmetric])
+  then show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
+    by auto
+  then have eq: "AE x in N. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
+    using ac absolutely_continuous_AE by auto
+
+
+  have "N (?RN 0) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
+    by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
+  also have "\<dots> = (\<integral>\<^sup>+ x. 0 \<partial>M)"
+    by (intro nn_integral_cong) (auto simp: indicator_def)
+  finally have "AE x in N. RN_deriv M N x \<noteq> 0"
+    using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)
+  with eq show "AE x in N. 0 < enn2real (RN_deriv M N x)"
+    by (auto simp: enn2real_positive_iff less_top[symmetric] zero_less_iff_neq_zero)
+qed (rule enn2real_nonneg)
+
+lemma (in sigma_finite_measure) RN_deriv_singleton:
+  assumes ac: "absolutely_continuous M N" "sets N = sets M"
+  and x: "{x} \<in> sets M"
+  shows "N {x} = RN_deriv M N x * emeasure M {x}"
+proof -
+  from \<open>{x} \<in> sets M\<close>
+  have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
+    by (auto simp: indicator_def emeasure_density intro!: nn_integral_cong)
+  with x density_RN_deriv[OF ac] show ?thesis
+    by (auto simp: max_def)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Regularity.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,383 @@
+(*  Title:      HOL/Probability/Regularity.thy
+    Author:     Fabian Immler, TU München
+*)
+
+section \<open>Regularity of Measures\<close>
+
+theory Regularity
+imports Measure_Space Borel_Space
+begin
+
+lemma
+  fixes M::"'a::{second_countable_topology, complete_space} measure"
+  assumes sb: "sets M = sets borel"
+  assumes "emeasure M (space M) \<noteq> \<infinity>"
+  assumes "B \<in> sets borel"
+  shows inner_regular: "emeasure M B =
+    (SUP K : {K. K \<subseteq> B \<and> compact K}. emeasure M K)" (is "?inner B")
+  and outer_regular: "emeasure M B =
+    (INF U : {U. B \<subseteq> U \<and> open U}. emeasure M U)" (is "?outer B")
+proof -
+  have Us: "UNIV = space M" by (metis assms(1) sets_eq_imp_space_eq space_borel)
+  hence sU: "space M = UNIV" by simp
+  interpret finite_measure M by rule fact
+  have approx_inner: "\<And>A. A \<in> sets M \<Longrightarrow>
+    (\<And>e. e > 0 \<Longrightarrow> \<exists>K. K \<subseteq> A \<and> compact K \<and> emeasure M A \<le> emeasure M K + ennreal e) \<Longrightarrow> ?inner A"
+    by (rule ennreal_approx_SUP)
+      (force intro!: emeasure_mono simp: compact_imp_closed emeasure_eq_measure)+
+  have approx_outer: "\<And>A. A \<in> sets M \<Longrightarrow>
+    (\<And>e. e > 0 \<Longrightarrow> \<exists>B. A \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M A + ennreal e) \<Longrightarrow> ?outer A"
+    by (rule ennreal_approx_INF)
+       (force intro!: emeasure_mono simp: emeasure_eq_measure sb)+
+  from countable_dense_setE guess X::"'a set"  . note X = this
+  {
+    fix r::real assume "r > 0" hence "\<And>y. open (ball y r)" "\<And>y. ball y r \<noteq> {}" by auto
+    with X(2)[OF this]
+    have x: "space M = (\<Union>x\<in>X. cball x r)"
+      by (auto simp add: sU) (metis dist_commute order_less_imp_le)
+    let ?U = "\<Union>k. (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)"
+    have "(\<lambda>k. emeasure M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M ?U"
+      by (rule Lim_emeasure_incseq) (auto intro!: borel_closed bexI simp: incseq_def Us sb)
+    also have "?U = space M"
+    proof safe
+      fix x from X(2)[OF open_ball[of x r]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto
+      show "x \<in> ?U"
+        using X(1) d
+        by simp (auto intro!: exI [where x = "to_nat_on X d"] simp: dist_commute Bex_def)
+    qed (simp add: sU)
+    finally have "(\<lambda>k. M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M (space M)" .
+  } note M_space = this
+  {
+    fix e ::real and n :: nat assume "e > 0" "n > 0"
+    hence "1/n > 0" "e * 2 powr - n > 0" by (auto)
+    from M_space[OF \<open>1/n>0\<close>]
+    have "(\<lambda>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) \<longlonglongrightarrow> measure M (space M)"
+      unfolding emeasure_eq_measure by (auto simp: measure_nonneg)
+    from metric_LIMSEQ_D[OF this \<open>0 < e * 2 powr -n\<close>]
+    obtain k where "dist (measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) <
+      e * 2 powr -n"
+      by auto
+    hence "measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge>
+      measure M (space M) - e * 2 powr -real n"
+      by (auto simp: dist_real_def)
+    hence "\<exists>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge>
+      measure M (space M) - e * 2 powr - real n" ..
+  } note k=this
+  hence "\<forall>e\<in>{0<..}. \<forall>(n::nat)\<in>{0<..}. \<exists>k.
+    measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge> measure M (space M) - e * 2 powr - real n"
+    by blast
+  then obtain k where k: "\<forall>e\<in>{0<..}. \<forall>n\<in>{0<..}. measure M (space M) - e * 2 powr - real (n::nat)
+    \<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
+    by metis
+  hence k: "\<And>e n. e > 0 \<Longrightarrow> n > 0 \<Longrightarrow> measure M (space M) - e * 2 powr - n
+    \<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
+    unfolding Ball_def by blast
+  have approx_space:
+    "\<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ennreal e"
+    (is "?thesis e") if "0 < e" for e :: real
+  proof -
+    define B where [abs_def]:
+      "B n = (\<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n
+    have "\<And>n. closed (B n)" by (auto simp: B_def)
+    hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb)
+    from k[OF \<open>e > 0\<close> zero_less_Suc]
+    have "\<And>n. measure M (space M) - measure M (B n) \<le> e * 2 powr - real (Suc n)"
+      by (simp add: algebra_simps B_def finite_measure_compl)
+    hence B_compl_le: "\<And>n::nat. measure M (space M - B n) \<le> e * 2 powr - real (Suc n)"
+      by (simp add: finite_measure_compl)
+    define K where "K = (\<Inter>n. B n)"
+    from \<open>closed (B _)\<close> have "closed K" by (auto simp: K_def)
+    hence [simp]: "K \<in> sets M" by (simp add: sb)
+    have "measure M (space M) - measure M K = measure M (space M - K)"
+      by (simp add: finite_measure_compl)
+    also have "\<dots> = emeasure M (\<Union>n. space M - B n)" by (auto simp: K_def emeasure_eq_measure)
+    also have "\<dots> \<le> (\<Sum>n. emeasure M (space M - B n))"
+      by (rule emeasure_subadditive_countably) (auto simp: summable_def)
+    also have "\<dots> \<le> (\<Sum>n. ennreal (e*2 powr - real (Suc n)))"
+      using B_compl_le by (intro suminf_le) (simp_all add: measure_nonneg emeasure_eq_measure ennreal_leI)
+    also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
+      by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc)
+    also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))"
+      unfolding ennreal_power[symmetric]
+      using \<open>0 < e\<close>
+      by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
+                    ennreal_power[symmetric])
+    also have "\<dots> = e"
+      by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
+    finally have "measure M (space M) \<le> measure M K + e"
+      using \<open>0 < e\<close> by simp
+    hence "emeasure M (space M) \<le> emeasure M K + e"
+      using \<open>0 < e\<close> by (simp add: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] del: ennreal_plus)
+    moreover have "compact K"
+      unfolding compact_eq_totally_bounded
+    proof safe
+      show "complete K" using \<open>closed K\<close> by (simp add: complete_eq_closed)
+      fix e'::real assume "0 < e'"
+      from nat_approx_posE[OF this] guess n . note n = this
+      let ?k = "from_nat_into X ` {0..k e (Suc n)}"
+      have "finite ?k" by simp
+      moreover have "K \<subseteq> (\<Union>x\<in>?k. ball x e')" unfolding K_def B_def using n by force
+      ultimately show "\<exists>k. finite k \<and> K \<subseteq> (\<Union>x\<in>k. ball x e')" by blast
+    qed
+    ultimately
+    show ?thesis by (auto simp: sU)
+  qed
+  { fix A::"'a set" assume "closed A" hence "A \<in> sets borel" by (simp add: compact_imp_closed)
+    hence [simp]: "A \<in> sets M" by (simp add: sb)
+    have "?inner A"
+    proof (rule approx_inner)
+      fix e::real assume "e > 0"
+      from approx_space[OF this] obtain K where
+        K: "K \<subseteq> space M" "compact K" "emeasure M (space M) \<le> emeasure M K + e"
+        by (auto simp: emeasure_eq_measure)
+      hence [simp]: "K \<in> sets M" by (simp add: sb compact_imp_closed)
+      have "measure M A - measure M (A \<inter> K) = measure M (A - A \<inter> K)"
+        by (subst finite_measure_Diff) auto
+      also have "A - A \<inter> K = A \<union> K - K" by auto
+      also have "measure M \<dots> = measure M (A \<union> K) - measure M K"
+        by (subst finite_measure_Diff) auto
+      also have "\<dots> \<le> measure M (space M) - measure M K"
+        by (simp add: emeasure_eq_measure sU sb finite_measure_mono)
+      also have "\<dots> \<le> e"
+        using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+      finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ennreal e"
+        using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+      moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto
+      ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ennreal e"
+        by blast
+    qed simp
+    have "?outer A"
+    proof cases
+      assume "A \<noteq> {}"
+      let ?G = "\<lambda>d. {x. infdist x A < d}"
+      {
+        fix d
+        have "?G d = (\<lambda>x. infdist x A) -` {..<d}" by auto
+        also have "open \<dots>"
+          by (intro continuous_open_vimage) (auto intro!: continuous_infdist continuous_ident)
+        finally have "open (?G d)" .
+      } note open_G = this
+      from in_closed_iff_infdist_zero[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>]
+      have "A = {x. infdist x A = 0}" by auto
+      also have "\<dots> = (\<Inter>i. ?G (1/real (Suc i)))"
+      proof (auto simp del: of_nat_Suc, rule ccontr)
+        fix x
+        assume "infdist x A \<noteq> 0"
+        hence pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp
+        from nat_approx_posE[OF this] guess n .
+        moreover
+        assume "\<forall>i. infdist x A < 1 / real (Suc i)"
+        hence "infdist x A < 1 / real (Suc n)" by auto
+        ultimately show False by simp
+      qed
+      also have "M \<dots> = (INF n. emeasure M (?G (1 / real (Suc n))))"
+      proof (rule INF_emeasure_decseq[symmetric], safe)
+        fix i::nat
+        from open_G[of "1 / real (Suc i)"]
+        show "?G (1 / real (Suc i)) \<in> sets M" by (simp add: sb borel_open)
+      next
+        show "decseq (\<lambda>i. {x. infdist x A < 1 / real (Suc i)})"
+          by (auto intro: less_trans intro!: divide_strict_left_mono
+            simp: decseq_def le_eq_less_or_eq)
+      qed simp
+      finally
+      have "emeasure M A = (INF n. emeasure M {x. infdist x A < 1 / real (Suc n)})" .
+      moreover
+      have "\<dots> \<ge> (INF U:{U. A \<subseteq> U \<and> open U}. emeasure M U)"
+      proof (intro INF_mono)
+        fix m
+        have "?G (1 / real (Suc m)) \<in> {U. A \<subseteq> U \<and> open U}" using open_G by auto
+        moreover have "M (?G (1 / real (Suc m))) \<le> M (?G (1 / real (Suc m)))" by simp
+        ultimately show "\<exists>U\<in>{U. A \<subseteq> U \<and> open U}.
+          emeasure M U \<le> emeasure M {x. infdist x A < 1 / real (Suc m)}"
+          by blast
+      qed
+      moreover
+      have "emeasure M A \<le> (INF U:{U. A \<subseteq> U \<and> open U}. emeasure M U)"
+        by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb)
+      ultimately show ?thesis by simp
+    qed (auto intro!: INF_eqI)
+    note \<open>?inner A\<close> \<open>?outer A\<close> }
+  note closed_in_D = this
+  from \<open>B \<in> sets borel\<close>
+  have "Int_stable (Collect closed)" "Collect closed \<subseteq> Pow UNIV" "B \<in> sigma_sets UNIV (Collect closed)"
+    by (auto simp: Int_stable_def borel_eq_closed)
+  then show "?inner B" "?outer B"
+  proof (induct B rule: sigma_sets_induct_disjoint)
+    case empty
+    { case 1 show ?case by (intro SUP_eqI[symmetric]) auto }
+    { case 2 show ?case by (intro INF_eqI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) }
+  next
+    case (basic B)
+    { case 1 from basic closed_in_D show ?case by auto }
+    { case 2 from basic closed_in_D show ?case by auto }
+  next
+    case (compl B)
+    note inner = compl(2) and outer = compl(3)
+    from compl have [simp]: "B \<in> sets M" by (auto simp: sb borel_eq_closed)
+    case 2
+    have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
+    also have "\<dots> = (INF K:{K. K \<subseteq> B \<and> compact K}. M (space M) -  M K)"
+      by (subst ennreal_SUP_const_minus) (auto simp: less_top[symmetric] inner)
+    also have "\<dots> = (INF U:{U. U \<subseteq> B \<and> compact U}. M (space M - U))"
+      by (rule INF_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
+    also have "\<dots> \<ge> (INF U:{U. U \<subseteq> B \<and> closed U}. M (space M - U))"
+      by (rule INF_superset_mono) (auto simp add: compact_imp_closed)
+    also have "(INF U:{U. U \<subseteq> B \<and> closed U}. M (space M - U)) =
+        (INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U)"
+      unfolding INF_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
+        by (rule INF_cong) (auto simp add: sU Compl_eq_Diff_UNIV [symmetric, simp])
+    finally have
+      "(INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<le> emeasure M (space M - B)" .
+    moreover have
+      "(INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<ge> emeasure M (space M - B)"
+      by (auto simp: sb sU intro!: INF_greatest emeasure_mono)
+    ultimately show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb])
+
+    case 1
+    have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
+    also have "\<dots> = (SUP U: {U. B \<subseteq> U \<and> open U}. M (space M) -  M U)"
+      unfolding outer by (subst ennreal_INF_const_minus) auto
+    also have "\<dots> = (SUP U:{U. B \<subseteq> U \<and> open U}. M (space M - U))"
+      by (rule SUP_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
+    also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> closed K}. emeasure M K)"
+      unfolding SUP_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
+        by (rule SUP_cong) (auto simp add: sU)
+    also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
+    proof (safe intro!: antisym SUP_least)
+      fix K assume "closed K" "K \<subseteq> space M - B"
+      from closed_in_D[OF \<open>closed K\<close>]
+      have K_inner: "emeasure M K = (SUP K:{Ka. Ka \<subseteq> K \<and> compact Ka}. emeasure M K)" by simp
+      show "emeasure M K \<le> (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
+        unfolding K_inner using \<open>K \<subseteq> space M - B\<close>
+        by (auto intro!: SUP_upper SUP_least)
+    qed (fastforce intro!: SUP_least SUP_upper simp: compact_imp_closed)
+    finally show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb])
+  next
+    case (union D)
+    then have "range D \<subseteq> sets M" by (auto simp: sb borel_eq_closed)
+    with union have M[symmetric]: "(\<Sum>i. M (D i)) = M (\<Union>i. D i)" by (intro suminf_emeasure)
+    also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))"
+      by (intro summable_LIMSEQ) auto
+    finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)"
+      by (simp add: emeasure_eq_measure measure_nonneg setsum_nonneg)
+    have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
+
+    case 1
+    show ?case
+    proof (rule approx_inner)
+      fix e::real assume "e > 0"
+      with measure_LIMSEQ
+      have "\<exists>no. \<forall>n\<ge>no. \<bar>(\<Sum>i<n. measure M (D i)) -measure M (\<Union>x. D x)\<bar> < e/2"
+        by (auto simp: lim_sequentially dist_real_def simp del: less_divide_eq_numeral1)
+      hence "\<exists>n0. \<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto
+      then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
+        unfolding choice_iff by blast
+      have "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
+        by (auto simp add: emeasure_eq_measure setsum_nonneg measure_nonneg)
+      also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule setsum_le_suminf) auto
+      also have "\<dots> = M (\<Union>i. D i)" by (simp add: M)
+      also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure)
+      finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2"
+        using n0 by (auto simp: measure_nonneg setsum_nonneg)
+      have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
+      proof
+        fix i
+        from \<open>0 < e\<close> have "0 < e/(2*Suc n0)" by simp
+        have "emeasure M (D i) = (SUP K:{K. K \<subseteq> (D i) \<and> compact K}. emeasure M K)"
+          using union by blast
+        from SUP_approx_ennreal[OF \<open>0 < e/(2*Suc n0)\<close> _ this]
+        show "\<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
+          by (auto simp: emeasure_eq_measure intro: less_imp_le compact_empty)
+      qed
+      then obtain K where K: "\<And>i. K i \<subseteq> D i" "\<And>i. compact (K i)"
+        "\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)"
+        unfolding choice_iff by blast
+      let ?K = "\<Union>i\<in>{..<n0}. K i"
+      have "disjoint_family_on K {..<n0}" using K \<open>disjoint_family D\<close>
+        unfolding disjoint_family_on_def by blast
+      hence mK: "measure M ?K = (\<Sum>i<n0. measure M (K i))" using K
+        by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed)
+      have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp
+      also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))"
+        using K \<open>0 < e\<close>
+        by (auto intro: setsum_mono simp: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] simp del: ennreal_plus)
+      also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
+        by (simp add: setsum.distrib)
+      also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) +  e / 2" using \<open>0 < e\<close>
+        by (auto simp: field_simps intro!: mult_left_mono)
+      finally
+      have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
+        by auto
+      hence "M (\<Union>i. D i) < M ?K + e"
+        using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] simp del: ennreal_plus)
+      moreover
+      have "?K \<subseteq> (\<Union>i. D i)" using K by auto
+      moreover
+      have "compact ?K" using K by auto
+      ultimately
+      have "?K\<subseteq>(\<Union>i. D i) \<and> compact ?K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M ?K + ennreal e" by simp
+      thus "\<exists>K\<subseteq>\<Union>i. D i. compact K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M K + ennreal e" ..
+    qed fact
+    case 2
+    show ?case
+    proof (rule approx_outer[OF \<open>(\<Union>i. D i) \<in> sets M\<close>])
+      fix e::real assume "e > 0"
+      have "\<forall>i::nat. \<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
+      proof
+        fix i::nat
+        from \<open>0 < e\<close> have "0 < e/(2 powr Suc i)" by simp
+        have "emeasure M (D i) = (INF U:{U. (D i) \<subseteq> U \<and> open U}. emeasure M U)"
+          using union by blast
+        from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
+        show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
+          using \<open>0<e\<close>
+          by (auto simp: emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] ennreal_minus
+                         finite_measure_mono sb
+                   simp del: ennreal_plus)
+      qed
+      then obtain U where U: "\<And>i. D i \<subseteq> U i" "\<And>i. open (U i)"
+        "\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)"
+        unfolding choice_iff by blast
+      let ?U = "\<Union>i. U i"
+      have "ennreal (measure M ?U - measure M (\<Union>i. D i)) = M ?U - M (\<Union>i. D i)"
+        using U(1,2)
+        by (subst ennreal_minus[symmetric])
+           (auto intro!: finite_measure_mono simp: sb measure_nonneg emeasure_eq_measure)
+      also have "\<dots> = M (?U - (\<Union>i. D i))" using U  \<open>(\<Union>i. D i) \<in> sets M\<close>
+        by (subst emeasure_Diff) (auto simp: sb)
+      also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U  \<open>range D \<subseteq> sets M\<close>
+        by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff)
+      also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U  \<open>range D \<subseteq> sets M\<close>
+        by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb)
+      also have "\<dots> \<le> (\<Sum>i. ennreal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
+        using \<open>0<e\<close>
+        by (intro suminf_le, subst emeasure_Diff)
+           (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg ennreal_minus
+                       finite_measure_mono divide_ennreal ennreal_less_iff
+                 intro: less_imp_le)
+      also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
+        using \<open>0<e\<close>
+        by (simp add: powr_minus powr_realpow field_simps divide_ennreal del: of_nat_Suc)
+      also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^  Suc n))"
+        unfolding ennreal_power[symmetric]
+        using \<open>0 < e\<close>
+        by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
+                      ennreal_power[symmetric])
+      also have "\<dots> = ennreal e"
+        by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
+      finally have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e"
+        using \<open>0<e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
+      moreover
+      have "(\<Union>i. D i) \<subseteq> ?U" using U by auto
+      moreover
+      have "open ?U" using U by auto
+      ultimately
+      have "(\<Union>i. D i) \<subseteq> ?U \<and> open ?U \<and> emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e" by simp
+      thus "\<exists>B. (\<Union>i. D i) \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M (\<Union>i. D i) + ennreal e" ..
+    qed
+  qed
+qed
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Set_Integral.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,602 @@
+(*  Title:      HOL/Probability/Set_Integral.thy
+    Author:     Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU)
+
+Notation and useful facts for working with integrals over a set.
+
+TODO: keep all these? Need unicode translations as well.
+*)
+
+theory Set_Integral
+  imports Bochner_Integration Lebesgue_Measure
+begin
+
+(*
+    Notation
+*)
+
+abbreviation "set_borel_measurable M A f \<equiv> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable M"
+
+abbreviation "set_integrable M A f \<equiv> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
+
+abbreviation "set_lebesgue_integral M A f \<equiv> lebesgue_integral M (\<lambda>x. indicator A x *\<^sub>R f x)"
+
+syntax
+"_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
+("(4LINT (_):(_)/|(_)./ _)" [0,60,110,61] 60)
+
+translations
+"LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
+
+abbreviation
+  "set_almost_everywhere A M P \<equiv> AE x in M. x \<in> A \<longrightarrow> P x"
+
+syntax
+  "_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
+("AE _\<in>_ in _./ _" [0,0,0,10] 10)
+
+translations
+  "AE x\<in>A in M. P" == "CONST set_almost_everywhere A M (\<lambda>x. P)"
+
+(*
+    Notation for integration wrt lebesgue measure on the reals:
+
+      LBINT x. f
+      LBINT x : A. f
+
+    TODO: keep all these? Need unicode.
+*)
+
+syntax
+"_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real"
+("(2LBINT _./ _)" [0,60] 60)
+
+translations
+"LBINT x. f" == "CONST lebesgue_integral CONST lborel (\<lambda>x. f)"
+
+syntax
+"_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> real"
+("(3LBINT _:_./ _)" [0,60,61] 60)
+
+translations
+"LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
+
+(*
+    Basic properties
+*)
+
+(*
+lemma indicator_abs_eq: "\<And>A x. \<bar>indicator A x\<bar> = ((indicator A x) :: real)"
+  by (auto simp add: indicator_def)
+*)
+
+lemma set_borel_measurable_sets:
+  fixes f :: "_ \<Rightarrow> _::real_normed_vector"
+  assumes "set_borel_measurable M X f" "B \<in> sets borel" "X \<in> sets M"
+  shows "f -` B \<inter> X \<in> sets M"
+proof -
+  have "f \<in> borel_measurable (restrict_space M X)"
+    using assms by (subst borel_measurable_restrict_space_iff) auto
+  then have "f -` B \<inter> space (restrict_space M X) \<in> sets (restrict_space M X)"
+    by (rule measurable_sets) fact
+  with \<open>X \<in> sets M\<close> show ?thesis
+    by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space)
+qed
+
+lemma set_lebesgue_integral_cong:
+  assumes "A \<in> sets M" and "\<forall>x. x \<in> A \<longrightarrow> f x = g x"
+  shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)"
+  using assms by (auto intro!: integral_cong split: split_indicator simp add: sets.sets_into_space)
+
+lemma set_lebesgue_integral_cong_AE:
+  assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  assumes "AE x \<in> A in M. f x = g x"
+  shows "LINT x:A|M. f x = LINT x:A|M. g x"
+proof-
+  have "AE x in M. indicator A x *\<^sub>R f x = indicator A x *\<^sub>R g x"
+    using assms by auto
+  thus ?thesis by (intro integral_cong_AE) auto
+qed
+
+lemma set_integrable_cong_AE:
+    "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
+    AE x \<in> A in M. f x = g x \<Longrightarrow> A \<in> sets M \<Longrightarrow>
+    set_integrable M A f = set_integrable M A g"
+  by (rule integrable_cong_AE) auto
+
+lemma set_integrable_subset:
+  fixes M A B and f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes "set_integrable M A f" "B \<in> sets M" "B \<subseteq> A"
+  shows "set_integrable M B f"
+proof -
+  have "set_integrable M B (\<lambda>x. indicator A x *\<^sub>R f x)"
+    by (rule integrable_mult_indicator) fact+
+  with \<open>B \<subseteq> A\<close> show ?thesis
+    by (simp add: indicator_inter_arith[symmetric] Int_absorb2)
+qed
+
+(* TODO: integral_cmul_indicator should be named set_integral_const *)
+(* TODO: borel_integrable_atLeastAtMost should be named something like set_integrable_Icc_isCont *)
+
+lemma set_integral_scaleR_right [simp]: "LINT t:A|M. a *\<^sub>R f t = a *\<^sub>R (LINT t:A|M. f t)"
+  by (subst integral_scaleR_right[symmetric]) (auto intro!: integral_cong)
+
+lemma set_integral_mult_right [simp]:
+  fixes a :: "'a::{real_normed_field, second_countable_topology}"
+  shows "LINT t:A|M. a * f t = a * (LINT t:A|M. f t)"
+  by (subst integral_mult_right_zero[symmetric]) (auto intro!: integral_cong)
+
+lemma set_integral_mult_left [simp]:
+  fixes a :: "'a::{real_normed_field, second_countable_topology}"
+  shows "LINT t:A|M. f t * a = (LINT t:A|M. f t) * a"
+  by (subst integral_mult_left_zero[symmetric]) (auto intro!: integral_cong)
+
+lemma set_integral_divide_zero [simp]:
+  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
+  shows "LINT t:A|M. f t / a = (LINT t:A|M. f t) / a"
+  by (subst integral_divide_zero[symmetric], intro integral_cong)
+     (auto split: split_indicator)
+
+lemma set_integrable_scaleR_right [simp, intro]:
+  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a *\<^sub>R f t)"
+  unfolding scaleR_left_commute by (rule integrable_scaleR_right)
+
+lemma set_integrable_scaleR_left [simp, intro]:
+  fixes a :: "_ :: {banach, second_countable_topology}"
+  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t *\<^sub>R a)"
+  using integrable_scaleR_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
+
+lemma set_integrable_mult_right [simp, intro]:
+  fixes a :: "'a::{real_normed_field, second_countable_topology}"
+  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a * f t)"
+  using integrable_mult_right[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
+
+lemma set_integrable_mult_left [simp, intro]:
+  fixes a :: "'a::{real_normed_field, second_countable_topology}"
+  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t * a)"
+  using integrable_mult_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
+
+lemma set_integrable_divide [simp, intro]:
+  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
+  assumes "a \<noteq> 0 \<Longrightarrow> set_integrable M A f"
+  shows "set_integrable M A (\<lambda>t. f t / a)"
+proof -
+  have "integrable M (\<lambda>x. indicator A x *\<^sub>R f x / a)"
+    using assms by (rule integrable_divide_zero)
+  also have "(\<lambda>x. indicator A x *\<^sub>R f x / a) = (\<lambda>x. indicator A x *\<^sub>R (f x / a))"
+    by (auto split: split_indicator)
+  finally show ?thesis .
+qed
+
+lemma set_integral_add [simp, intro]:
+  fixes f g :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes "set_integrable M A f" "set_integrable M A g"
+  shows "set_integrable M A (\<lambda>x. f x + g x)"
+    and "LINT x:A|M. f x + g x = (LINT x:A|M. f x) + (LINT x:A|M. g x)"
+  using assms by (simp_all add: scaleR_add_right)
+
+lemma set_integral_diff [simp, intro]:
+  assumes "set_integrable M A f" "set_integrable M A g"
+  shows "set_integrable M A (\<lambda>x. f x - g x)" and "LINT x:A|M. f x - g x =
+    (LINT x:A|M. f x) - (LINT x:A|M. g x)"
+  using assms by (simp_all add: scaleR_diff_right)
+
+lemma set_integral_reflect:
+  fixes S and f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  shows "(LBINT x : S. f x) = (LBINT x : {x. - x \<in> S}. f (- x))"
+  by (subst lborel_integral_real_affine[where c="-1" and t=0])
+     (auto intro!: integral_cong split: split_indicator)
+
+(* question: why do we have this for negation, but multiplication by a constant
+   requires an integrability assumption? *)
+lemma set_integral_uminus: "set_integrable M A f \<Longrightarrow> LINT x:A|M. - f x = - (LINT x:A|M. f x)"
+  by (subst integral_minus[symmetric]) simp_all
+
+lemma set_integral_complex_of_real:
+  "LINT x:A|M. complex_of_real (f x) = of_real (LINT x:A|M. f x)"
+  by (subst integral_complex_of_real[symmetric])
+     (auto intro!: integral_cong split: split_indicator)
+
+lemma set_integral_mono:
+  fixes f g :: "_ \<Rightarrow> real"
+  assumes "set_integrable M A f" "set_integrable M A g"
+    "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"
+  shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)"
+using assms by (auto intro: integral_mono split: split_indicator)
+
+lemma set_integral_mono_AE:
+  fixes f g :: "_ \<Rightarrow> real"
+  assumes "set_integrable M A f" "set_integrable M A g"
+    "AE x \<in> A in M. f x \<le> g x"
+  shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)"
+using assms by (auto intro: integral_mono_AE split: split_indicator)
+
+lemma set_integrable_abs: "set_integrable M A f \<Longrightarrow> set_integrable M A (\<lambda>x. \<bar>f x\<bar> :: real)"
+  using integrable_abs[of M "\<lambda>x. f x * indicator A x"] by (simp add: abs_mult ac_simps)
+
+lemma set_integrable_abs_iff:
+  fixes f :: "_ \<Rightarrow> real"
+  shows "set_borel_measurable M A f \<Longrightarrow> set_integrable M A (\<lambda>x. \<bar>f x\<bar>) = set_integrable M A f"
+  by (subst (2) integrable_abs_iff[symmetric]) (simp_all add: abs_mult ac_simps)
+
+lemma set_integrable_abs_iff':
+  fixes f :: "_ \<Rightarrow> real"
+  shows "f \<in> borel_measurable M \<Longrightarrow> A \<in> sets M \<Longrightarrow>
+    set_integrable M A (\<lambda>x. \<bar>f x\<bar>) = set_integrable M A f"
+by (intro set_integrable_abs_iff) auto
+
+lemma set_integrable_discrete_difference:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes "countable X"
+  assumes diff: "(A - B) \<union> (B - A) \<subseteq> X"
+  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+  shows "set_integrable M A f \<longleftrightarrow> set_integrable M B f"
+proof (rule integrable_discrete_difference[where X=X])
+  show "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x"
+    using diff by (auto split: split_indicator)
+qed fact+
+
+lemma set_integral_discrete_difference:
+  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
+  assumes "countable X"
+  assumes diff: "(A - B) \<union> (B - A) \<subseteq> X"
+  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+  shows "set_lebesgue_integral M A f = set_lebesgue_integral M B f"
+proof (rule integral_discrete_difference[where X=X])
+  show "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x"
+    using diff by (auto split: split_indicator)
+qed fact+
+
+lemma set_integrable_Un:
+  fixes f g :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes f_A: "set_integrable M A f" and f_B:  "set_integrable M B f"
+    and [measurable]: "A \<in> sets M" "B \<in> sets M"
+  shows "set_integrable M (A \<union> B) f"
+proof -
+  have "set_integrable M (A - B) f"
+    using f_A by (rule set_integrable_subset) auto
+  from integrable_add[OF this f_B] show ?thesis
+    by (rule integrable_cong[THEN iffD1, rotated 2]) (auto split: split_indicator)
+qed
+
+lemma set_integrable_UN:
+  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> set_integrable M (A i) f"
+    "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M"
+  shows "set_integrable M (\<Union>i\<in>I. A i) f"
+using assms by (induct I) (auto intro!: set_integrable_Un)
+
+lemma set_integral_Un:
+  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes "A \<inter> B = {}"
+  and "set_integrable M A f"
+  and "set_integrable M B f"
+  shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
+by (auto simp add: indicator_union_arith indicator_inter_arith[symmetric]
+      scaleR_add_left assms)
+
+lemma set_integral_cong_set:
+  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes [measurable]: "set_borel_measurable M A f" "set_borel_measurable M B f"
+    and ae: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
+  shows "LINT x:B|M. f x = LINT x:A|M. f x"
+proof (rule integral_cong_AE)
+  show "AE x in M. indicator B x *\<^sub>R f x = indicator A x *\<^sub>R f x"
+    using ae by (auto simp: subset_eq split: split_indicator)
+qed fact+
+
+lemma set_borel_measurable_subset:
+  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes [measurable]: "set_borel_measurable M A f" "B \<in> sets M" and "B \<subseteq> A"
+  shows "set_borel_measurable M B f"
+proof -
+  have "set_borel_measurable M B (\<lambda>x. indicator A x *\<^sub>R f x)"
+    by measurable
+  also have "(\<lambda>x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\<lambda>x. indicator B x *\<^sub>R f x)"
+    using \<open>B \<subseteq> A\<close> by (auto simp: fun_eq_iff split: split_indicator)
+  finally show ?thesis .
+qed
+
+lemma set_integral_Un_AE:
+  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes ae: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)" and [measurable]: "A \<in> sets M" "B \<in> sets M"
+  and "set_integrable M A f"
+  and "set_integrable M B f"
+  shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
+proof -
+  have f: "set_integrable M (A \<union> B) f"
+    by (intro set_integrable_Un assms)
+  then have f': "set_borel_measurable M (A \<union> B) f"
+    by (rule borel_measurable_integrable)
+  have "LINT x:A\<union>B|M. f x = LINT x:(A - A \<inter> B) \<union> (B - A \<inter> B)|M. f x"
+  proof (rule set_integral_cong_set)
+    show "AE x in M. (x \<in> A - A \<inter> B \<union> (B - A \<inter> B)) = (x \<in> A \<union> B)"
+      using ae by auto
+    show "set_borel_measurable M (A - A \<inter> B \<union> (B - A \<inter> B)) f"
+      using f' by (rule set_borel_measurable_subset) auto
+  qed fact
+  also have "\<dots> = (LINT x:(A - A \<inter> B)|M. f x) + (LINT x:(B - A \<inter> B)|M. f x)"
+    by (auto intro!: set_integral_Un set_integrable_subset[OF f])
+  also have "\<dots> = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
+    using ae
+    by (intro arg_cong2[where f="op+"] set_integral_cong_set)
+       (auto intro!: set_borel_measurable_subset[OF f'])
+  finally show ?thesis .
+qed
+
+lemma set_integral_finite_Union:
+  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
+  assumes "finite I" "disjoint_family_on A I"
+    and "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
+  shows "(LINT x:(\<Union>i\<in>I. A i)|M. f x) = (\<Sum>i\<in>I. LINT x:A i|M. f x)"
+  using assms
+  apply induct
+  apply (auto intro!: set_integral_Un set_integrable_Un set_integrable_UN simp: disjoint_family_on_def)
+by (subst set_integral_Un, auto intro: set_integrable_UN)
+
+(* TODO: find a better name? *)
+lemma pos_integrable_to_top:
+  fixes l::real
+  assumes "\<And>i. A i \<in> sets M" "mono A"
+  assumes nneg: "\<And>x i. x \<in> A i \<Longrightarrow> 0 \<le> f x"
+  and intgbl: "\<And>i::nat. set_integrable M (A i) f"
+  and lim: "(\<lambda>i::nat. LINT x:A i|M. f x) \<longlonglongrightarrow> l"
+  shows "set_integrable M (\<Union>i. A i) f"
+  apply (rule integrable_monotone_convergence[where f = "\<lambda>i::nat. \<lambda>x. indicator (A i) x *\<^sub>R f x" and x = l])
+  apply (rule intgbl)
+  prefer 3 apply (rule lim)
+  apply (rule AE_I2)
+  using \<open>mono A\<close> apply (auto simp: mono_def nneg split: split_indicator) []
+proof (rule AE_I2)
+  { fix x assume "x \<in> space M"
+    show "(\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Union>i. A i) x *\<^sub>R f x"
+    proof cases
+      assume "\<exists>i. x \<in> A i"
+      then guess i ..
+      then have *: "eventually (\<lambda>i. x \<in> A i) sequentially"
+        using \<open>x \<in> A i\<close> \<open>mono A\<close> by (auto simp: eventually_sequentially mono_def)
+      show ?thesis
+        apply (intro Lim_eventually)
+        using *
+        apply eventually_elim
+        apply (auto split: split_indicator)
+        done
+    qed auto }
+  then show "(\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R f x) \<in> borel_measurable M"
+    apply (rule borel_measurable_LIMSEQ_real)
+    apply assumption
+    apply (intro borel_measurable_integrable intgbl)
+    done
+qed
+
+(* Proof from Royden Real Analysis, p. 91. *)
+lemma lebesgue_integral_countable_add:
+  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  assumes meas[intro]: "\<And>i::nat. A i \<in> sets M"
+    and disj: "\<And>i j. i \<noteq> j \<Longrightarrow> A i \<inter> A j = {}"
+    and intgbl: "set_integrable M (\<Union>i. A i) f"
+  shows "LINT x:(\<Union>i. A i)|M. f x = (\<Sum>i. (LINT x:(A i)|M. f x))"
+proof (subst integral_suminf[symmetric])
+  show int_A: "\<And>i. set_integrable M (A i) f"
+    using intgbl by (rule set_integrable_subset) auto
+  { fix x assume "x \<in> space M"
+    have "(\<lambda>i. indicator (A i) x *\<^sub>R f x) sums (indicator (\<Union>i. A i) x *\<^sub>R f x)"
+      by (intro sums_scaleR_left indicator_sums) fact }
+  note sums = this
+
+  have norm_f: "\<And>i. set_integrable M (A i) (\<lambda>x. norm (f x))"
+    using int_A[THEN integrable_norm] by auto
+
+  show "AE x in M. summable (\<lambda>i. norm (indicator (A i) x *\<^sub>R f x))"
+    using disj by (intro AE_I2) (auto intro!: summable_mult2 sums_summable[OF indicator_sums])
+
+  show "summable (\<lambda>i. LINT x|M. norm (indicator (A i) x *\<^sub>R f x))"
+  proof (rule summableI_nonneg_bounded)
+    fix n
+    show "0 \<le> LINT x|M. norm (indicator (A n) x *\<^sub>R f x)"
+      using norm_f by (auto intro!: integral_nonneg_AE)
+
+    have "(\<Sum>i<n. LINT x|M. norm (indicator (A i) x *\<^sub>R f x)) =
+      (\<Sum>i<n. set_lebesgue_integral M (A i) (\<lambda>x. norm (f x)))"
+      by (simp add: abs_mult)
+    also have "\<dots> = set_lebesgue_integral M (\<Union>i<n. A i) (\<lambda>x. norm (f x))"
+      using norm_f
+      by (subst set_integral_finite_Union) (auto simp: disjoint_family_on_def disj)
+    also have "\<dots> \<le> set_lebesgue_integral M (\<Union>i. A i) (\<lambda>x. norm (f x))"
+      using intgbl[THEN integrable_norm]
+      by (intro integral_mono set_integrable_UN[of "{..<n}"] norm_f)
+         (auto split: split_indicator)
+    finally show "(\<Sum>i<n. LINT x|M. norm (indicator (A i) x *\<^sub>R f x)) \<le>
+      set_lebesgue_integral M (\<Union>i. A i) (\<lambda>x. norm (f x))"
+      by simp
+  qed
+  show "set_lebesgue_integral M (UNION UNIV A) f = LINT x|M. (\<Sum>i. indicator (A i) x *\<^sub>R f x)"
+    apply (rule integral_cong[OF refl])
+    apply (subst suminf_scaleR_left[OF sums_summable[OF indicator_sums, OF disj], symmetric])
+    using sums_unique[OF indicator_sums[OF disj]]
+    apply auto
+    done
+qed
+
+lemma set_integral_cont_up:
+  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  assumes [measurable]: "\<And>i. A i \<in> sets M" and A: "incseq A"
+  and intgbl: "set_integrable M (\<Union>i. A i) f"
+  shows "(\<lambda>i. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Union>i. A i)|M. f x"
+proof (intro integral_dominated_convergence[where w="\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R norm (f x)"])
+  have int_A: "\<And>i. set_integrable M (A i) f"
+    using intgbl by (rule set_integrable_subset) auto
+  then show "\<And>i. set_borel_measurable M (A i) f" "set_borel_measurable M (\<Union>i. A i) f"
+    "set_integrable M (\<Union>i. A i) (\<lambda>x. norm (f x))"
+    using intgbl integrable_norm[OF intgbl] by auto
+
+  { fix x i assume "x \<in> A i"
+    with A have "(\<lambda>xa. indicator (A xa) x::real) \<longlonglongrightarrow> 1 \<longleftrightarrow> (\<lambda>xa. 1::real) \<longlonglongrightarrow> 1"
+      by (intro filterlim_cong refl)
+         (fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) }
+  then show "AE x in M. (\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Union>i. A i) x *\<^sub>R f x"
+    by (intro AE_I2 tendsto_intros) (auto split: split_indicator)
+qed (auto split: split_indicator)
+
+(* Can the int0 hypothesis be dropped? *)
+lemma set_integral_cont_down:
+  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
+  assumes [measurable]: "\<And>i. A i \<in> sets M" and A: "decseq A"
+  and int0: "set_integrable M (A 0) f"
+  shows "(\<lambda>i::nat. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Inter>i. A i)|M. f x"
+proof (rule integral_dominated_convergence)
+  have int_A: "\<And>i. set_integrable M (A i) f"
+    using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
+  show "set_integrable M (A 0) (\<lambda>x. norm (f x))"
+    using int0[THEN integrable_norm] by simp
+  have "set_integrable M (\<Inter>i. A i) f"
+    using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
+  with int_A show "set_borel_measurable M (\<Inter>i. A i) f" "\<And>i. set_borel_measurable M (A i) f"
+    by auto
+  show "\<And>i. AE x in M. norm (indicator (A i) x *\<^sub>R f x) \<le> indicator (A 0) x *\<^sub>R norm (f x)"
+    using A by (auto split: split_indicator simp: decseq_def)
+  { fix x i assume "x \<in> space M" "x \<notin> A i"
+    with A have "(\<lambda>i. indicator (A i) x::real) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<lambda>i. 0::real) \<longlonglongrightarrow> 0"
+      by (intro filterlim_cong refl)
+         (auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) }
+  then show "AE x in M. (\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Inter>i. A i) x *\<^sub>R f x"
+    by (intro AE_I2 tendsto_intros) (auto split: split_indicator)
+qed
+
+lemma set_integral_at_point:
+  fixes a :: real
+  assumes "set_integrable M {a} f"
+  and [simp]: "{a} \<in> sets M" and "(emeasure M) {a} \<noteq> \<infinity>"
+  shows "(LINT x:{a} | M. f x) = f a * measure M {a}"
+proof-
+  have "set_lebesgue_integral M {a} f = set_lebesgue_integral M {a} (%x. f a)"
+    by (intro set_lebesgue_integral_cong) simp_all
+  then show ?thesis using assms by simp
+qed
+
+
+abbreviation complex_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
+  "complex_integrable M f \<equiv> integrable M f"
+
+abbreviation complex_lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" ("integral\<^sup>C") where
+  "integral\<^sup>C M f == integral\<^sup>L M f"
+
+syntax
+  "_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex"
+ ("\<integral>\<^sup>C _. _ \<partial>_" [60,61] 110)
+
+translations
+  "\<integral>\<^sup>Cx. f \<partial>M" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"
+
+syntax
+  "_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
+  ("(3CLINT _|_. _)" [0,110,60] 60)
+
+translations
+  "CLINT x|M. f" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"
+
+lemma complex_integrable_cnj [simp]:
+  "complex_integrable M (\<lambda>x. cnj (f x)) \<longleftrightarrow> complex_integrable M f"
+proof
+  assume "complex_integrable M (\<lambda>x. cnj (f x))"
+  then have "complex_integrable M (\<lambda>x. cnj (cnj (f x)))"
+    by (rule integrable_cnj)
+  then show "complex_integrable M f"
+    by simp
+qed simp
+
+lemma complex_of_real_integrable_eq:
+  "complex_integrable M (\<lambda>x. complex_of_real (f x)) \<longleftrightarrow> integrable M f"
+proof
+  assume "complex_integrable M (\<lambda>x. complex_of_real (f x))"
+  then have "integrable M (\<lambda>x. Re (complex_of_real (f x)))"
+    by (rule integrable_Re)
+  then show "integrable M f"
+    by simp
+qed simp
+
+
+abbreviation complex_set_integrable :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
+  "complex_set_integrable M A f \<equiv> set_integrable M A f"
+
+abbreviation complex_set_lebesgue_integral :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" where
+  "complex_set_lebesgue_integral M A f \<equiv> set_lebesgue_integral M A f"
+
+syntax
+"_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
+("(4CLINT _:_|_. _)" [0,60,110,61] 60)
+
+translations
+"CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\<lambda>x. f)"
+
+(*
+lemma cmod_mult: "cmod ((a :: real) * (x :: complex)) = \<bar>a\<bar> * cmod x"
+  apply (simp add: norm_mult)
+  by (subst norm_mult, auto)
+*)
+
+lemma borel_integrable_atLeastAtMost':
+  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
+  assumes f: "continuous_on {a..b} f"
+  shows "set_integrable lborel {a..b} f" (is "integrable _ ?f")
+  by (intro borel_integrable_compact compact_Icc f)
+
+lemma integral_FTC_atLeastAtMost:
+  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
+  assumes "a \<le> b"
+    and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
+    and f: "continuous_on {a .. b} f"
+  shows "integral\<^sup>L lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) = F b - F a"
+proof -
+  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
+  have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV"
+    using borel_integrable_atLeastAtMost'[OF f] by (rule has_integral_integral_lborel)
+  moreover
+  have "(f has_integral F b - F a) {a .. b}"
+    by (intro fundamental_theorem_of_calculus ballI assms) auto
+  then have "(?f has_integral F b - F a) {a .. b}"
+    by (subst has_integral_cong[where g=f]) auto
+  then have "(?f has_integral F b - F a) UNIV"
+    by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto
+  ultimately show "integral\<^sup>L lborel ?f = F b - F a"
+    by (rule has_integral_unique)
+qed
+
+lemma set_borel_integral_eq_integral:
+  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
+  assumes "set_integrable lborel S f"
+  shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f"
+proof -
+  let ?f = "\<lambda>x. indicator S x *\<^sub>R f x"
+  have "(?f has_integral LINT x : S | lborel. f x) UNIV"
+    by (rule has_integral_integral_lborel) fact
+  hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S"
+    apply (subst has_integral_restrict_univ [symmetric])
+    apply (rule has_integral_eq)
+    by auto
+  thus "f integrable_on S"
+    by (auto simp add: integrable_on_def)
+  with 1 have "(f has_integral (integral S f)) S"
+    by (intro integrable_integral, auto simp add: integrable_on_def)
+  thus "LINT x : S | lborel. f x = integral S f"
+    by (intro has_integral_unique [OF 1])
+qed
+
+lemma set_borel_measurable_continuous:
+  fixes f :: "_ \<Rightarrow> _::real_normed_vector"
+  assumes "S \<in> sets borel" "continuous_on S f"
+  shows "set_borel_measurable borel S f"
+proof -
+  have "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable borel"
+    by (intro assms borel_measurable_continuous_on_if continuous_on_const)
+  also have "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. indicator S x *\<^sub>R f x)"
+    by auto
+  finally show ?thesis .
+qed
+
+lemma set_measurable_continuous_on_ivl:
+  assumes "continuous_on {a..b} (f :: real \<Rightarrow> real)"
+  shows "set_borel_measurable borel {a..b} f"
+  by (rule set_borel_measurable_continuous[OF _ assms]) simp
+
+end
+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Sigma_Algebra.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,2262 @@
+(*  Title:      HOL/Probability/Sigma_Algebra.thy
+    Author:     Stefan Richter, Markus Wenzel, TU München
+    Author:     Johannes Hölzl, TU München
+    Plus material from the Hurd/Coble measure theory development,
+    translated by Lawrence Paulson.
+*)
+
+section \<open>Describing measurable sets\<close>
+
+theory Sigma_Algebra
+imports
+  Complex_Main
+  "~~/src/HOL/Library/Countable_Set"
+  "~~/src/HOL/Library/FuncSet"
+  "~~/src/HOL/Library/Indicator_Function"
+  "~~/src/HOL/Library/Extended_Nonnegative_Real"
+  "~~/src/HOL/Library/Disjoint_Sets"
+begin
+
+text \<open>Sigma algebras are an elementary concept in measure
+  theory. To measure --- that is to integrate --- functions, we first have
+  to measure sets. Unfortunately, when dealing with a large universe,
+  it is often not possible to consistently assign a measure to every
+  subset. Therefore it is necessary to define the set of measurable
+  subsets of the universe. A sigma algebra is such a set that has
+  three very natural and desirable properties.\<close>
+
+subsection \<open>Families of sets\<close>
+
+locale subset_class =
+  fixes \<Omega> :: "'a set" and M :: "'a set set"
+  assumes space_closed: "M \<subseteq> Pow \<Omega>"
+
+lemma (in subset_class) sets_into_space: "x \<in> M \<Longrightarrow> x \<subseteq> \<Omega>"
+  by (metis PowD contra_subsetD space_closed)
+
+subsubsection \<open>Semiring of sets\<close>
+
+locale semiring_of_sets = subset_class +
+  assumes empty_sets[iff]: "{} \<in> M"
+  assumes Int[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"
+  assumes Diff_cover:
+    "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> \<exists>C\<subseteq>M. finite C \<and> disjoint C \<and> a - b = \<Union>C"
+
+lemma (in semiring_of_sets) finite_INT[intro]:
+  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
+  shows "(\<Inter>i\<in>I. A i) \<in> M"
+  using assms by (induct rule: finite_ne_induct) auto
+
+lemma (in semiring_of_sets) Int_space_eq1 [simp]: "x \<in> M \<Longrightarrow> \<Omega> \<inter> x = x"
+  by (metis Int_absorb1 sets_into_space)
+
+lemma (in semiring_of_sets) Int_space_eq2 [simp]: "x \<in> M \<Longrightarrow> x \<inter> \<Omega> = x"
+  by (metis Int_absorb2 sets_into_space)
+
+lemma (in semiring_of_sets) sets_Collect_conj:
+  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
+  shows "{x\<in>\<Omega>. Q x \<and> P x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. Q x \<and> P x} = {x\<in>\<Omega>. Q x} \<inter> {x\<in>\<Omega>. P x}"
+    by auto
+  with assms show ?thesis by auto
+qed
+
+lemma (in semiring_of_sets) sets_Collect_finite_All':
+  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S" "S \<noteq> {}"
+  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (\<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
+    using \<open>S \<noteq> {}\<close> by auto
+  with assms show ?thesis by auto
+qed
+
+locale ring_of_sets = semiring_of_sets +
+  assumes Un [intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"
+
+lemma (in ring_of_sets) finite_Union [intro]:
+  "finite X \<Longrightarrow> X \<subseteq> M \<Longrightarrow> \<Union>X \<in> M"
+  by (induct set: finite) (auto simp add: Un)
+
+lemma (in ring_of_sets) finite_UN[intro]:
+  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
+  shows "(\<Union>i\<in>I. A i) \<in> M"
+  using assms by induct auto
+
+lemma (in ring_of_sets) Diff [intro]:
+  assumes "a \<in> M" "b \<in> M" shows "a - b \<in> M"
+  using Diff_cover[OF assms] by auto
+
+lemma ring_of_setsI:
+  assumes space_closed: "M \<subseteq> Pow \<Omega>"
+  assumes empty_sets[iff]: "{} \<in> M"
+  assumes Un[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"
+  assumes Diff[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a - b \<in> M"
+  shows "ring_of_sets \<Omega> M"
+proof
+  fix a b assume ab: "a \<in> M" "b \<in> M"
+  from ab show "\<exists>C\<subseteq>M. finite C \<and> disjoint C \<and> a - b = \<Union>C"
+    by (intro exI[of _ "{a - b}"]) (auto simp: disjoint_def)
+  have "a \<inter> b = a - (a - b)" by auto
+  also have "\<dots> \<in> M" using ab by auto
+  finally show "a \<inter> b \<in> M" .
+qed fact+
+
+lemma ring_of_sets_iff: "ring_of_sets \<Omega> M \<longleftrightarrow> M \<subseteq> Pow \<Omega> \<and> {} \<in> M \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a \<union> b \<in> M) \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a - b \<in> M)"
+proof
+  assume "ring_of_sets \<Omega> M"
+  then interpret ring_of_sets \<Omega> M .
+  show "M \<subseteq> Pow \<Omega> \<and> {} \<in> M \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a \<union> b \<in> M) \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a - b \<in> M)"
+    using space_closed by auto
+qed (auto intro!: ring_of_setsI)
+
+lemma (in ring_of_sets) insert_in_sets:
+  assumes "{x} \<in> M" "A \<in> M" shows "insert x A \<in> M"
+proof -
+  have "{x} \<union> A \<in> M" using assms by (rule Un)
+  thus ?thesis by auto
+qed
+
+lemma (in ring_of_sets) sets_Collect_disj:
+  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
+  shows "{x\<in>\<Omega>. Q x \<or> P x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. Q x \<or> P x} = {x\<in>\<Omega>. Q x} \<union> {x\<in>\<Omega>. P x}"
+    by auto
+  with assms show ?thesis by auto
+qed
+
+lemma (in ring_of_sets) sets_Collect_finite_Ex:
+  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
+  shows "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} = (\<Union>i\<in>S. {x\<in>\<Omega>. P i x})"
+    by auto
+  with assms show ?thesis by auto
+qed
+
+locale algebra = ring_of_sets +
+  assumes top [iff]: "\<Omega> \<in> M"
+
+lemma (in algebra) compl_sets [intro]:
+  "a \<in> M \<Longrightarrow> \<Omega> - a \<in> M"
+  by auto
+
+lemma algebra_iff_Un:
+  "algebra \<Omega> M \<longleftrightarrow>
+    M \<subseteq> Pow \<Omega> \<and>
+    {} \<in> M \<and>
+    (\<forall>a \<in> M. \<Omega> - a \<in> M) \<and>
+    (\<forall>a \<in> M. \<forall> b \<in> M. a \<union> b \<in> M)" (is "_ \<longleftrightarrow> ?Un")
+proof
+  assume "algebra \<Omega> M"
+  then interpret algebra \<Omega> M .
+  show ?Un using sets_into_space by auto
+next
+  assume ?Un
+  then have "\<Omega> \<in> M" by auto
+  interpret ring_of_sets \<Omega> M
+  proof (rule ring_of_setsI)
+    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
+      using \<open>?Un\<close> by auto
+    fix a b assume a: "a \<in> M" and b: "b \<in> M"
+    then show "a \<union> b \<in> M" using \<open>?Un\<close> by auto
+    have "a - b = \<Omega> - ((\<Omega> - a) \<union> b)"
+      using \<Omega> a b by auto
+    then show "a - b \<in> M"
+      using a b  \<open>?Un\<close> by auto
+  qed
+  show "algebra \<Omega> M" proof qed fact
+qed
+
+lemma algebra_iff_Int:
+     "algebra \<Omega> M \<longleftrightarrow>
+       M \<subseteq> Pow \<Omega> & {} \<in> M &
+       (\<forall>a \<in> M. \<Omega> - a \<in> M) &
+       (\<forall>a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)" (is "_ \<longleftrightarrow> ?Int")
+proof
+  assume "algebra \<Omega> M"
+  then interpret algebra \<Omega> M .
+  show ?Int using sets_into_space by auto
+next
+  assume ?Int
+  show "algebra \<Omega> M"
+  proof (unfold algebra_iff_Un, intro conjI ballI)
+    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
+      using \<open>?Int\<close> by auto
+    from \<open>?Int\<close> show "\<And>a. a \<in> M \<Longrightarrow> \<Omega> - a \<in> M" by auto
+    fix a b assume M: "a \<in> M" "b \<in> M"
+    hence "a \<union> b = \<Omega> - ((\<Omega> - a) \<inter> (\<Omega> - b))"
+      using \<Omega> by blast
+    also have "... \<in> M"
+      using M \<open>?Int\<close> by auto
+    finally show "a \<union> b \<in> M" .
+  qed
+qed
+
+lemma (in algebra) sets_Collect_neg:
+  assumes "{x\<in>\<Omega>. P x} \<in> M"
+  shows "{x\<in>\<Omega>. \<not> P x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<not> P x} = \<Omega> - {x\<in>\<Omega>. P x}" by auto
+  with assms show ?thesis by auto
+qed
+
+lemma (in algebra) sets_Collect_imp:
+  "{x\<in>\<Omega>. P x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x \<longrightarrow> P x} \<in> M"
+  unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg)
+
+lemma (in algebra) sets_Collect_const:
+  "{x\<in>\<Omega>. P} \<in> M"
+  by (cases P) auto
+
+lemma algebra_single_set:
+  "X \<subseteq> S \<Longrightarrow> algebra S { {}, X, S - X, S }"
+  by (auto simp: algebra_iff_Int)
+
+subsubsection \<open>Restricted algebras\<close>
+
+abbreviation (in algebra)
+  "restricted_space A \<equiv> (op \<inter> A) ` M"
+
+lemma (in algebra) restricted_algebra:
+  assumes "A \<in> M" shows "algebra A (restricted_space A)"
+  using assms by (auto simp: algebra_iff_Int)
+
+subsubsection \<open>Sigma Algebras\<close>
+
+locale sigma_algebra = algebra +
+  assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+
+lemma (in algebra) is_sigma_algebra:
+  assumes "finite M"
+  shows "sigma_algebra \<Omega> M"
+proof
+  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> M"
+  then have "(\<Union>i. A i) = (\<Union>s\<in>M \<inter> range A. s)"
+    by auto
+  also have "(\<Union>s\<in>M \<inter> range A. s) \<in> M"
+    using \<open>finite M\<close> by auto
+  finally show "(\<Union>i. A i) \<in> M" .
+qed
+
+lemma countable_UN_eq:
+  fixes A :: "'i::countable \<Rightarrow> 'a set"
+  shows "(range A \<subseteq> M \<longrightarrow> (\<Union>i. A i) \<in> M) \<longleftrightarrow>
+    (range (A \<circ> from_nat) \<subseteq> M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> M)"
+proof -
+  let ?A' = "A \<circ> from_nat"
+  have *: "(\<Union>i. ?A' i) = (\<Union>i. A i)" (is "?l = ?r")
+  proof safe
+    fix x i assume "x \<in> A i" thus "x \<in> ?l"
+      by (auto intro!: exI[of _ "to_nat i"])
+  next
+    fix x i assume "x \<in> ?A' i" thus "x \<in> ?r"
+      by (auto intro!: exI[of _ "from_nat i"])
+  qed
+  have **: "range ?A' = range A"
+    using surj_from_nat
+    by (auto simp: image_comp [symmetric] intro!: imageI)
+  show ?thesis unfolding * ** ..
+qed
+
+lemma (in sigma_algebra) countable_Union [intro]:
+  assumes "countable X" "X \<subseteq> M" shows "\<Union>X \<in> M"
+proof cases
+  assume "X \<noteq> {}"
+  hence "\<Union>X = (\<Union>n. from_nat_into X n)"
+    using assms by (auto intro: from_nat_into) (metis from_nat_into_surj)
+  also have "\<dots> \<in> M" using assms
+    by (auto intro!: countable_nat_UN) (metis \<open>X \<noteq> {}\<close> from_nat_into set_mp)
+  finally show ?thesis .
+qed simp
+
+lemma (in sigma_algebra) countable_UN[intro]:
+  fixes A :: "'i::countable \<Rightarrow> 'a set"
+  assumes "A`X \<subseteq> M"
+  shows  "(\<Union>x\<in>X. A x) \<in> M"
+proof -
+  let ?A = "\<lambda>i. if i \<in> X then A i else {}"
+  from assms have "range ?A \<subseteq> M" by auto
+  with countable_nat_UN[of "?A \<circ> from_nat"] countable_UN_eq[of ?A M]
+  have "(\<Union>x. ?A x) \<in> M" by auto
+  moreover have "(\<Union>x. ?A x) = (\<Union>x\<in>X. A x)" by (auto split: if_split_asm)
+  ultimately show ?thesis by simp
+qed
+
+lemma (in sigma_algebra) countable_UN':
+  fixes A :: "'i \<Rightarrow> 'a set"
+  assumes X: "countable X"
+  assumes A: "A`X \<subseteq> M"
+  shows  "(\<Union>x\<in>X. A x) \<in> M"
+proof -
+  have "(\<Union>x\<in>X. A x) = (\<Union>i\<in>to_nat_on X ` X. A (from_nat_into X i))"
+    using X by auto
+  also have "\<dots> \<in> M"
+    using A X
+    by (intro countable_UN) auto
+  finally show ?thesis .
+qed
+
+lemma (in sigma_algebra) countable_UN'':
+  "\<lbrakk> countable X; \<And>x y. x \<in> X \<Longrightarrow> A x \<in> M \<rbrakk> \<Longrightarrow> (\<Union>x\<in>X. A x) \<in> M"
+by(erule countable_UN')(auto)
+
+lemma (in sigma_algebra) countable_INT [intro]:
+  fixes A :: "'i::countable \<Rightarrow> 'a set"
+  assumes A: "A`X \<subseteq> M" "X \<noteq> {}"
+  shows "(\<Inter>i\<in>X. A i) \<in> M"
+proof -
+  from A have "\<forall>i\<in>X. A i \<in> M" by fast
+  hence "\<Omega> - (\<Union>i\<in>X. \<Omega> - A i) \<in> M" by blast
+  moreover
+  have "(\<Inter>i\<in>X. A i) = \<Omega> - (\<Union>i\<in>X. \<Omega> - A i)" using space_closed A
+    by blast
+  ultimately show ?thesis by metis
+qed
+
+lemma (in sigma_algebra) countable_INT':
+  fixes A :: "'i \<Rightarrow> 'a set"
+  assumes X: "countable X" "X \<noteq> {}"
+  assumes A: "A`X \<subseteq> M"
+  shows  "(\<Inter>x\<in>X. A x) \<in> M"
+proof -
+  have "(\<Inter>x\<in>X. A x) = (\<Inter>i\<in>to_nat_on X ` X. A (from_nat_into X i))"
+    using X by auto
+  also have "\<dots> \<in> M"
+    using A X
+    by (intro countable_INT) auto
+  finally show ?thesis .
+qed
+
+lemma (in sigma_algebra) countable_INT'':
+  "UNIV \<in> M \<Longrightarrow> countable I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> M) \<Longrightarrow> (\<Inter>i\<in>I. F i) \<in> M"
+  by (cases "I = {}") (auto intro: countable_INT')
+
+lemma (in sigma_algebra) countable:
+  assumes "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> M" "countable A"
+  shows "A \<in> M"
+proof -
+  have "(\<Union>a\<in>A. {a}) \<in> M"
+    using assms by (intro countable_UN') auto
+  also have "(\<Union>a\<in>A. {a}) = A" by auto
+  finally show ?thesis by auto
+qed
+
+lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)"
+  by (auto simp: ring_of_sets_iff)
+
+lemma algebra_Pow: "algebra sp (Pow sp)"
+  by (auto simp: algebra_iff_Un)
+
+lemma sigma_algebra_iff:
+  "sigma_algebra \<Omega> M \<longleftrightarrow>
+    algebra \<Omega> M \<and> (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
+  by (simp add: sigma_algebra_def sigma_algebra_axioms_def)
+
+lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)"
+  by (auto simp: sigma_algebra_iff algebra_iff_Int)
+
+lemma (in sigma_algebra) sets_Collect_countable_All:
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} = (\<Inter>i. {x\<in>\<Omega>. P i x})" by auto
+  with assms show ?thesis by auto
+qed
+
+lemma (in sigma_algebra) sets_Collect_countable_Ex:
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} = (\<Union>i. {x\<in>\<Omega>. P i x})" by auto
+  with assms show ?thesis by auto
+qed
+
+lemma (in sigma_algebra) sets_Collect_countable_Ex':
+  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
+  assumes "countable I"
+  shows "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} = (\<Union>i\<in>I. {x\<in>\<Omega>. P i x})" by auto
+  with assms show ?thesis
+    by (auto intro!: countable_UN')
+qed
+
+lemma (in sigma_algebra) sets_Collect_countable_All':
+  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
+  assumes "countable I"
+  shows "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} = (\<Inter>i\<in>I. {x\<in>\<Omega>. P i x}) \<inter> \<Omega>" by auto
+  with assms show ?thesis
+    by (cases "I = {}") (auto intro!: countable_INT')
+qed
+
+lemma (in sigma_algebra) sets_Collect_countable_Ex1':
+  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
+  assumes "countable I"
+  shows "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} \<in> M"
+proof -
+  have "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} = {x\<in>\<Omega>. \<exists>i\<in>I. P i x \<and> (\<forall>j\<in>I. P j x \<longrightarrow> i = j)}"
+    by auto
+  with assms show ?thesis
+    by (auto intro!: sets_Collect_countable_All' sets_Collect_countable_Ex' sets_Collect_conj sets_Collect_imp sets_Collect_const)
+qed
+
+lemmas (in sigma_algebra) sets_Collect =
+  sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const
+  sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All
+
+lemma (in sigma_algebra) sets_Collect_countable_Ball:
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<forall>i::'i::countable\<in>X. P i x} \<in> M"
+  unfolding Ball_def by (intro sets_Collect assms)
+
+lemma (in sigma_algebra) sets_Collect_countable_Bex:
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<exists>i::'i::countable\<in>X. P i x} \<in> M"
+  unfolding Bex_def by (intro sets_Collect assms)
+
+lemma sigma_algebra_single_set:
+  assumes "X \<subseteq> S"
+  shows "sigma_algebra S { {}, X, S - X, S }"
+  using algebra.is_sigma_algebra[OF algebra_single_set[OF \<open>X \<subseteq> S\<close>]] by simp
+
+subsubsection \<open>Binary Unions\<close>
+
+definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
+  where "binary a b =  (\<lambda>x. b)(0 := a)"
+
+lemma range_binary_eq: "range(binary a b) = {a,b}"
+  by (auto simp add: binary_def)
+
+lemma Un_range_binary: "a \<union> b = (\<Union>i::nat. binary a b i)"
+  by (simp add: range_binary_eq cong del: strong_SUP_cong)
+
+lemma Int_range_binary: "a \<inter> b = (\<Inter>i::nat. binary a b i)"
+  by (simp add: range_binary_eq cong del: strong_INF_cong)
+
+lemma sigma_algebra_iff2:
+     "sigma_algebra \<Omega> M \<longleftrightarrow>
+       M \<subseteq> Pow \<Omega> \<and>
+       {} \<in> M \<and> (\<forall>s \<in> M. \<Omega> - s \<in> M) \<and>
+       (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
+  by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
+         algebra_iff_Un Un_range_binary)
+
+subsubsection \<open>Initial Sigma Algebra\<close>
+
+text \<open>Sigma algebras can naturally be created as the closure of any set of
+  M with regard to the properties just postulated.\<close>
+
+inductive_set sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
+  for sp :: "'a set" and A :: "'a set set"
+  where
+    Basic[intro, simp]: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"
+  | Empty: "{} \<in> sigma_sets sp A"
+  | Compl: "a \<in> sigma_sets sp A \<Longrightarrow> sp - a \<in> sigma_sets sp A"
+  | Union: "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Union>i. a i) \<in> sigma_sets sp A"
+
+lemma (in sigma_algebra) sigma_sets_subset:
+  assumes a: "a \<subseteq> M"
+  shows "sigma_sets \<Omega> a \<subseteq> M"
+proof
+  fix x
+  assume "x \<in> sigma_sets \<Omega> a"
+  from this show "x \<in> M"
+    by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
+qed
+
+lemma sigma_sets_into_sp: "A \<subseteq> Pow sp \<Longrightarrow> x \<in> sigma_sets sp A \<Longrightarrow> x \<subseteq> sp"
+  by (erule sigma_sets.induct, auto)
+
+lemma sigma_algebra_sigma_sets:
+     "a \<subseteq> Pow \<Omega> \<Longrightarrow> sigma_algebra \<Omega> (sigma_sets \<Omega> a)"
+  by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp
+           intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl)
+
+lemma sigma_sets_least_sigma_algebra:
+  assumes "A \<subseteq> Pow S"
+  shows "sigma_sets S A = \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
+proof safe
+  fix B X assume "A \<subseteq> B" and sa: "sigma_algebra S B"
+    and X: "X \<in> sigma_sets S A"
+  from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF \<open>A \<subseteq> B\<close>] X
+  show "X \<in> B" by auto
+next
+  fix X assume "X \<in> \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
+  then have [intro!]: "\<And>B. A \<subseteq> B \<Longrightarrow> sigma_algebra S B \<Longrightarrow> X \<in> B"
+     by simp
+  have "A \<subseteq> sigma_sets S A" using assms by auto
+  moreover have "sigma_algebra S (sigma_sets S A)"
+    using assms by (intro sigma_algebra_sigma_sets[of A]) auto
+  ultimately show "X \<in> sigma_sets S A" by auto
+qed
+
+lemma sigma_sets_top: "sp \<in> sigma_sets sp A"
+  by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty)
+
+lemma sigma_sets_Un:
+  "a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"
+apply (simp add: Un_range_binary range_binary_eq)
+apply (rule Union, simp add: binary_def)
+done
+
+lemma sigma_sets_Inter:
+  assumes Asb: "A \<subseteq> Pow sp"
+  shows "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Inter>i. a i) \<in> sigma_sets sp A"
+proof -
+  assume ai: "\<And>i::nat. a i \<in> sigma_sets sp A"
+  hence "\<And>i::nat. sp-(a i) \<in> sigma_sets sp A"
+    by (rule sigma_sets.Compl)
+  hence "(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
+    by (rule sigma_sets.Union)
+  hence "sp-(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
+    by (rule sigma_sets.Compl)
+  also have "sp-(\<Union>i. sp-(a i)) = sp Int (\<Inter>i. a i)"
+    by auto
+  also have "... = (\<Inter>i. a i)" using ai
+    by (blast dest: sigma_sets_into_sp [OF Asb])
+  finally show ?thesis .
+qed
+
+lemma sigma_sets_INTER:
+  assumes Asb: "A \<subseteq> Pow sp"
+      and ai: "\<And>i::nat. i \<in> S \<Longrightarrow> a i \<in> sigma_sets sp A" and non: "S \<noteq> {}"
+  shows "(\<Inter>i\<in>S. a i) \<in> sigma_sets sp A"
+proof -
+  from ai have "\<And>i. (if i\<in>S then a i else sp) \<in> sigma_sets sp A"
+    by (simp add: sigma_sets.intros(2-) sigma_sets_top)
+  hence "(\<Inter>i. (if i\<in>S then a i else sp)) \<in> sigma_sets sp A"
+    by (rule sigma_sets_Inter [OF Asb])
+  also have "(\<Inter>i. (if i\<in>S then a i else sp)) = (\<Inter>i\<in>S. a i)"
+    by auto (metis ai non sigma_sets_into_sp subset_empty subset_iff Asb)+
+  finally show ?thesis .
+qed
+
+lemma sigma_sets_UNION:
+  "countable B \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets X A) \<Longrightarrow> (\<Union>B) \<in> sigma_sets X A"
+  apply (cases "B = {}")
+  apply (simp add: sigma_sets.Empty)
+  using from_nat_into [of B] range_from_nat_into [of B] sigma_sets.Union [of "from_nat_into B" X A]
+  apply simp
+  apply auto
+  apply (metis Sup_bot_conv(1) Union_empty \<open>\<lbrakk>B \<noteq> {}; countable B\<rbrakk> \<Longrightarrow> range (from_nat_into B) = B\<close>)
+  done
+
+lemma (in sigma_algebra) sigma_sets_eq:
+     "sigma_sets \<Omega> M = M"
+proof
+  show "M \<subseteq> sigma_sets \<Omega> M"
+    by (metis Set.subsetI sigma_sets.Basic)
+  next
+  show "sigma_sets \<Omega> M \<subseteq> M"
+    by (metis sigma_sets_subset subset_refl)
+qed
+
+lemma sigma_sets_eqI:
+  assumes A: "\<And>a. a \<in> A \<Longrightarrow> a \<in> sigma_sets M B"
+  assumes B: "\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets M A"
+  shows "sigma_sets M A = sigma_sets M B"
+proof (intro set_eqI iffI)
+  fix a assume "a \<in> sigma_sets M A"
+  from this A show "a \<in> sigma_sets M B"
+    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
+next
+  fix b assume "b \<in> sigma_sets M B"
+  from this B show "b \<in> sigma_sets M A"
+    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
+qed
+
+lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+    by induct (insert \<open>A \<subseteq> B\<close>, auto intro: sigma_sets.intros(2-))
+qed
+
+lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+    by induct (insert \<open>A \<subseteq> sigma_sets X B\<close>, auto intro: sigma_sets.intros(2-))
+qed
+
+lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+    by induct (insert \<open>A \<subseteq> B\<close>, auto intro: sigma_sets.intros(2-))
+qed
+
+lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"
+  by (auto intro: sigma_sets.Basic)
+
+lemma (in sigma_algebra) restriction_in_sets:
+  fixes A :: "nat \<Rightarrow> 'a set"
+  assumes "S \<in> M"
+  and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` M" (is "_ \<subseteq> ?r")
+  shows "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
+proof -
+  { fix i have "A i \<in> ?r" using * by auto
+    hence "\<exists>B. A i = B \<inter> S \<and> B \<in> M" by auto
+    hence "A i \<subseteq> S" "A i \<in> M" using \<open>S \<in> M\<close> by auto }
+  thus "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
+    by (auto intro!: image_eqI[of _ _ "(\<Union>i. A i)"])
+qed
+
+lemma (in sigma_algebra) restricted_sigma_algebra:
+  assumes "S \<in> M"
+  shows "sigma_algebra S (restricted_space S)"
+  unfolding sigma_algebra_def sigma_algebra_axioms_def
+proof safe
+  show "algebra S (restricted_space S)" using restricted_algebra[OF assms] .
+next
+  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> restricted_space S"
+  from restriction_in_sets[OF assms this[simplified]]
+  show "(\<Union>i. A i) \<in> restricted_space S" by simp
+qed
+
+lemma sigma_sets_Int:
+  assumes "A \<in> sigma_sets sp st" "A \<subseteq> sp"
+  shows "op \<inter> A ` sigma_sets sp st = sigma_sets A (op \<inter> A ` st)"
+proof (intro equalityI subsetI)
+  fix x assume "x \<in> op \<inter> A ` sigma_sets sp st"
+  then obtain y where "y \<in> sigma_sets sp st" "x = y \<inter> A" by auto
+  then have "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+  proof (induct arbitrary: x)
+    case (Compl a)
+    then show ?case
+      by (force intro!: sigma_sets.Compl simp: Diff_Int_distrib ac_simps)
+  next
+    case (Union a)
+    then show ?case
+      by (auto intro!: sigma_sets.Union
+               simp add: UN_extend_simps simp del: UN_simps)
+  qed (auto intro!: sigma_sets.intros(2-))
+  then show "x \<in> sigma_sets A (op \<inter> A ` st)"
+    using \<open>A \<subseteq> sp\<close> by (simp add: Int_absorb2)
+next
+  fix x assume "x \<in> sigma_sets A (op \<inter> A ` st)"
+  then show "x \<in> op \<inter> A ` sigma_sets sp st"
+  proof induct
+    case (Compl a)
+    then obtain x where "a = A \<inter> x" "x \<in> sigma_sets sp st" by auto
+    then show ?case using \<open>A \<subseteq> sp\<close>
+      by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)
+  next
+    case (Union a)
+    then have "\<forall>i. \<exists>x. x \<in> sigma_sets sp st \<and> a i = A \<inter> x"
+      by (auto simp: image_iff Bex_def)
+    from choice[OF this] guess f ..
+    then show ?case
+      by (auto intro!: bexI[of _ "(\<Union>x. f x)"] sigma_sets.Union
+               simp add: image_iff)
+  qed (auto intro!: sigma_sets.intros(2-))
+qed
+
+lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}"
+proof (intro set_eqI iffI)
+  fix a assume "a \<in> sigma_sets A {}" then show "a \<in> {{}, A}"
+    by induct blast+
+qed (auto intro: sigma_sets.Empty sigma_sets_top)
+
+lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}"
+proof (intro set_eqI iffI)
+  fix x assume "x \<in> sigma_sets A {A}"
+  then show "x \<in> {{}, A}"
+    by induct blast+
+next
+  fix x assume "x \<in> {{}, A}"
+  then show "x \<in> sigma_sets A {A}"
+    by (auto intro: sigma_sets.Empty sigma_sets_top)
+qed
+
+lemma sigma_sets_sigma_sets_eq:
+  "M \<subseteq> Pow S \<Longrightarrow> sigma_sets S (sigma_sets S M) = sigma_sets S M"
+  by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto
+
+lemma sigma_sets_singleton:
+  assumes "X \<subseteq> S"
+  shows "sigma_sets S { X } = { {}, X, S - X, S }"
+proof -
+  interpret sigma_algebra S "{ {}, X, S - X, S }"
+    by (rule sigma_algebra_single_set) fact
+  have "sigma_sets S { X } \<subseteq> sigma_sets S { {}, X, S - X, S }"
+    by (rule sigma_sets_subseteq) simp
+  moreover have "\<dots> = { {}, X, S - X, S }"
+    using sigma_sets_eq by simp
+  moreover
+  { fix A assume "A \<in> { {}, X, S - X, S }"
+    then have "A \<in> sigma_sets S { X }"
+      by (auto intro: sigma_sets.intros(2-) sigma_sets_top) }
+  ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }"
+    by (intro antisym) auto
+  with sigma_sets_eq show ?thesis by simp
+qed
+
+lemma restricted_sigma:
+  assumes S: "S \<in> sigma_sets \<Omega> M" and M: "M \<subseteq> Pow \<Omega>"
+  shows "algebra.restricted_space (sigma_sets \<Omega> M) S =
+    sigma_sets S (algebra.restricted_space M S)"
+proof -
+  from S sigma_sets_into_sp[OF M]
+  have "S \<in> sigma_sets \<Omega> M" "S \<subseteq> \<Omega>" by auto
+  from sigma_sets_Int[OF this]
+  show ?thesis by simp
+qed
+
+lemma sigma_sets_vimage_commute:
+  assumes X: "X \<in> \<Omega> \<rightarrow> \<Omega>'"
+  shows "{X -` A \<inter> \<Omega> |A. A \<in> sigma_sets \<Omega>' M'}
+       = sigma_sets \<Omega> {X -` A \<inter> \<Omega> |A. A \<in> M'}" (is "?L = ?R")
+proof
+  show "?L \<subseteq> ?R"
+  proof clarify
+    fix A assume "A \<in> sigma_sets \<Omega>' M'"
+    then show "X -` A \<inter> \<Omega> \<in> ?R"
+    proof induct
+      case Empty then show ?case
+        by (auto intro!: sigma_sets.Empty)
+    next
+      case (Compl B)
+      have [simp]: "X -` (\<Omega>' - B) \<inter> \<Omega> = \<Omega> - (X -` B \<inter> \<Omega>)"
+        by (auto simp add: funcset_mem [OF X])
+      with Compl show ?case
+        by (auto intro!: sigma_sets.Compl)
+    next
+      case (Union F)
+      then show ?case
+        by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps
+                 intro!: sigma_sets.Union)
+    qed auto
+  qed
+  show "?R \<subseteq> ?L"
+  proof clarify
+    fix A assume "A \<in> ?R"
+    then show "\<exists>B. A = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'"
+    proof induct
+      case (Basic B) then show ?case by auto
+    next
+      case Empty then show ?case
+        by (auto intro!: sigma_sets.Empty exI[of _ "{}"])
+    next
+      case (Compl B)
+      then obtain A where A: "B = X -` A \<inter> \<Omega>" "A \<in> sigma_sets \<Omega>' M'" by auto
+      then have [simp]: "\<Omega> - B = X -` (\<Omega>' - A) \<inter> \<Omega>"
+        by (auto simp add: funcset_mem [OF X])
+      with A(2) show ?case
+        by (auto intro: sigma_sets.Compl)
+    next
+      case (Union F)
+      then have "\<forall>i. \<exists>B. F i = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'" by auto
+      from choice[OF this] guess A .. note A = this
+      with A show ?case
+        by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union)
+    qed
+  qed
+qed
+
+lemma (in ring_of_sets) UNION_in_sets:
+  fixes A:: "nat \<Rightarrow> 'a set"
+  assumes A: "range A \<subseteq> M"
+  shows  "(\<Union>i\<in>{0..<n}. A i) \<in> M"
+proof (induct n)
+  case 0 show ?case by simp
+next
+  case (Suc n)
+  thus ?case
+    by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)
+qed
+
+lemma (in ring_of_sets) range_disjointed_sets:
+  assumes A: "range A \<subseteq> M"
+  shows  "range (disjointed A) \<subseteq> M"
+proof (auto simp add: disjointed_def)
+  fix n
+  show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> M" using UNION_in_sets
+    by (metis A Diff UNIV_I image_subset_iff)
+qed
+
+lemma (in algebra) range_disjointed_sets':
+  "range A \<subseteq> M \<Longrightarrow> range (disjointed A) \<subseteq> M"
+  using range_disjointed_sets .
+
+lemma sigma_algebra_disjoint_iff:
+  "sigma_algebra \<Omega> M \<longleftrightarrow> algebra \<Omega> M \<and>
+    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
+proof (auto simp add: sigma_algebra_iff)
+  fix A :: "nat \<Rightarrow> 'a set"
+  assume M: "algebra \<Omega> M"
+     and A: "range A \<subseteq> M"
+     and UnA: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  hence "range (disjointed A) \<subseteq> M \<longrightarrow>
+         disjoint_family (disjointed A) \<longrightarrow>
+         (\<Union>i. disjointed A i) \<in> M" by blast
+  hence "(\<Union>i. disjointed A i) \<in> M"
+    by (simp add: algebra.range_disjointed_sets'[of \<Omega>] M A disjoint_family_disjointed)
+  thus "(\<Union>i::nat. A i) \<in> M" by (simp add: UN_disjointed_eq)
+qed
+
+subsubsection \<open>Ring generated by a semiring\<close>
+
+definition (in semiring_of_sets)
+  "generated_ring = { \<Union>C | C. C \<subseteq> M \<and> finite C \<and> disjoint C }"
+
+lemma (in semiring_of_sets) generated_ringE[elim?]:
+  assumes "a \<in> generated_ring"
+  obtains C where "finite C" "disjoint C" "C \<subseteq> M" "a = \<Union>C"
+  using assms unfolding generated_ring_def by auto
+
+lemma (in semiring_of_sets) generated_ringI[intro?]:
+  assumes "finite C" "disjoint C" "C \<subseteq> M" "a = \<Union>C"
+  shows "a \<in> generated_ring"
+  using assms unfolding generated_ring_def by auto
+
+lemma (in semiring_of_sets) generated_ringI_Basic:
+  "A \<in> M \<Longrightarrow> A \<in> generated_ring"
+  by (rule generated_ringI[of "{A}"]) (auto simp: disjoint_def)
+
+lemma (in semiring_of_sets) generated_ring_disjoint_Un[intro]:
+  assumes a: "a \<in> generated_ring" and b: "b \<in> generated_ring"
+  and "a \<inter> b = {}"
+  shows "a \<union> b \<in> generated_ring"
+proof -
+  from a guess Ca .. note Ca = this
+  from b guess Cb .. note Cb = this
+  show ?thesis
+  proof
+    show "disjoint (Ca \<union> Cb)"
+      using \<open>a \<inter> b = {}\<close> Ca Cb by (auto intro!: disjoint_union)
+  qed (insert Ca Cb, auto)
+qed
+
+lemma (in semiring_of_sets) generated_ring_empty: "{} \<in> generated_ring"
+  by (auto simp: generated_ring_def disjoint_def)
+
+lemma (in semiring_of_sets) generated_ring_disjoint_Union:
+  assumes "finite A" shows "A \<subseteq> generated_ring \<Longrightarrow> disjoint A \<Longrightarrow> \<Union>A \<in> generated_ring"
+  using assms by (induct A) (auto simp: disjoint_def intro!: generated_ring_disjoint_Un generated_ring_empty)
+
+lemma (in semiring_of_sets) generated_ring_disjoint_UNION:
+  "finite I \<Longrightarrow> disjoint (A ` I) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> UNION I A \<in> generated_ring"
+  by (intro generated_ring_disjoint_Union) auto
+
+lemma (in semiring_of_sets) generated_ring_Int:
+  assumes a: "a \<in> generated_ring" and b: "b \<in> generated_ring"
+  shows "a \<inter> b \<in> generated_ring"
+proof -
+  from a guess Ca .. note Ca = this
+  from b guess Cb .. note Cb = this
+  define C where "C = (\<lambda>(a,b). a \<inter> b)` (Ca\<times>Cb)"
+  show ?thesis
+  proof
+    show "disjoint C"
+    proof (simp add: disjoint_def C_def, intro ballI impI)
+      fix a1 b1 a2 b2 assume sets: "a1 \<in> Ca" "b1 \<in> Cb" "a2 \<in> Ca" "b2 \<in> Cb"
+      assume "a1 \<inter> b1 \<noteq> a2 \<inter> b2"
+      then have "a1 \<noteq> a2 \<or> b1 \<noteq> b2" by auto
+      then show "(a1 \<inter> b1) \<inter> (a2 \<inter> b2) = {}"
+      proof
+        assume "a1 \<noteq> a2"
+        with sets Ca have "a1 \<inter> a2 = {}"
+          by (auto simp: disjoint_def)
+        then show ?thesis by auto
+      next
+        assume "b1 \<noteq> b2"
+        with sets Cb have "b1 \<inter> b2 = {}"
+          by (auto simp: disjoint_def)
+        then show ?thesis by auto
+      qed
+    qed
+  qed (insert Ca Cb, auto simp: C_def)
+qed
+
+lemma (in semiring_of_sets) generated_ring_Inter:
+  assumes "finite A" "A \<noteq> {}" shows "A \<subseteq> generated_ring \<Longrightarrow> \<Inter>A \<in> generated_ring"
+  using assms by (induct A rule: finite_ne_induct) (auto intro: generated_ring_Int)
+
+lemma (in semiring_of_sets) generated_ring_INTER:
+  "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> INTER I A \<in> generated_ring"
+  by (intro generated_ring_Inter) auto
+
+lemma (in semiring_of_sets) generating_ring:
+  "ring_of_sets \<Omega> generated_ring"
+proof (rule ring_of_setsI)
+  let ?R = generated_ring
+  show "?R \<subseteq> Pow \<Omega>"
+    using sets_into_space by (auto simp: generated_ring_def generated_ring_empty)
+  show "{} \<in> ?R" by (rule generated_ring_empty)
+
+  { fix a assume a: "a \<in> ?R" then guess Ca .. note Ca = this
+    fix b assume b: "b \<in> ?R" then guess Cb .. note Cb = this
+
+    show "a - b \<in> ?R"
+    proof cases
+      assume "Cb = {}" with Cb \<open>a \<in> ?R\<close> show ?thesis
+        by simp
+    next
+      assume "Cb \<noteq> {}"
+      with Ca Cb have "a - b = (\<Union>a'\<in>Ca. \<Inter>b'\<in>Cb. a' - b')" by auto
+      also have "\<dots> \<in> ?R"
+      proof (intro generated_ring_INTER generated_ring_disjoint_UNION)
+        fix a b assume "a \<in> Ca" "b \<in> Cb"
+        with Ca Cb Diff_cover[of a b] show "a - b \<in> ?R"
+          by (auto simp add: generated_ring_def)
+            (metis DiffI Diff_eq_empty_iff empty_iff)
+      next
+        show "disjoint ((\<lambda>a'. \<Inter>b'\<in>Cb. a' - b')`Ca)"
+          using Ca by (auto simp add: disjoint_def \<open>Cb \<noteq> {}\<close>)
+      next
+        show "finite Ca" "finite Cb" "Cb \<noteq> {}" by fact+
+      qed
+      finally show "a - b \<in> ?R" .
+    qed }
+  note Diff = this
+
+  fix a b assume sets: "a \<in> ?R" "b \<in> ?R"
+  have "a \<union> b = (a - b) \<union> (a \<inter> b) \<union> (b - a)" by auto
+  also have "\<dots> \<in> ?R"
+    by (intro sets generated_ring_disjoint_Un generated_ring_Int Diff) auto
+  finally show "a \<union> b \<in> ?R" .
+qed
+
+lemma (in semiring_of_sets) sigma_sets_generated_ring_eq: "sigma_sets \<Omega> generated_ring = sigma_sets \<Omega> M"
+proof
+  interpret M: sigma_algebra \<Omega> "sigma_sets \<Omega> M"
+    using space_closed by (rule sigma_algebra_sigma_sets)
+  show "sigma_sets \<Omega> generated_ring \<subseteq> sigma_sets \<Omega> M"
+    by (blast intro!: sigma_sets_mono elim: generated_ringE)
+qed (auto intro!: generated_ringI_Basic sigma_sets_mono)
+
+subsubsection \<open>A Two-Element Series\<close>
+
+definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set"
+  where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
+
+lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
+  apply (simp add: binaryset_def)
+  apply (rule set_eqI)
+  apply (auto simp add: image_iff)
+  done
+
+lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
+  by (simp add: range_binaryset_eq cong del: strong_SUP_cong)
+
+subsubsection \<open>Closed CDI\<close>
+
+definition closed_cdi where
+  "closed_cdi \<Omega> M \<longleftrightarrow>
+   M \<subseteq> Pow \<Omega> &
+   (\<forall>s \<in> M. \<Omega> - s \<in> M) &
+   (\<forall>A. (range A \<subseteq> M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
+        (\<Union>i. A i) \<in> M) &
+   (\<forall>A. (range A \<subseteq> M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
+
+inductive_set
+  smallest_ccdi_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
+  for \<Omega> M
+  where
+    Basic [intro]:
+      "a \<in> M \<Longrightarrow> a \<in> smallest_ccdi_sets \<Omega> M"
+  | Compl [intro]:
+      "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> \<Omega> - a \<in> smallest_ccdi_sets \<Omega> M"
+  | Inc:
+      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
+       \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets \<Omega> M"
+  | Disj:
+      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> disjoint_family A
+       \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets \<Omega> M"
+
+lemma (in subset_class) smallest_closed_cdi1: "M \<subseteq> smallest_ccdi_sets \<Omega> M"
+  by auto
+
+lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets \<Omega> M \<subseteq> Pow \<Omega>"
+  apply (rule subsetI)
+  apply (erule smallest_ccdi_sets.induct)
+  apply (auto intro: range_subsetD dest: sets_into_space)
+  done
+
+lemma (in subset_class) smallest_closed_cdi2: "closed_cdi \<Omega> (smallest_ccdi_sets \<Omega> M)"
+  apply (auto simp add: closed_cdi_def smallest_ccdi_sets)
+  apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
+  done
+
+lemma closed_cdi_subset: "closed_cdi \<Omega> M \<Longrightarrow> M \<subseteq> Pow \<Omega>"
+  by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Compl: "closed_cdi \<Omega> M \<Longrightarrow> s \<in> M \<Longrightarrow> \<Omega> - s \<in> M"
+  by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Inc:
+  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow> (\<Union>i. A i) \<in> M"
+  by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Disj:
+  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  by (simp add: closed_cdi_def)
+
+lemma closed_cdi_Un:
+  assumes cdi: "closed_cdi \<Omega> M" and empty: "{} \<in> M"
+      and A: "A \<in> M" and B: "B \<in> M"
+      and disj: "A \<inter> B = {}"
+    shows "A \<union> B \<in> M"
+proof -
+  have ra: "range (binaryset A B) \<subseteq> M"
+   by (simp add: range_binaryset_eq empty A B)
+ have di:  "disjoint_family (binaryset A B)" using disj
+   by (simp add: disjoint_family_on_def binaryset_def Int_commute)
+ from closed_cdi_Disj [OF cdi ra di]
+ show ?thesis
+   by (simp add: UN_binaryset_eq)
+qed
+
+lemma (in algebra) smallest_ccdi_sets_Un:
+  assumes A: "A \<in> smallest_ccdi_sets \<Omega> M" and B: "B \<in> smallest_ccdi_sets \<Omega> M"
+      and disj: "A \<inter> B = {}"
+    shows "A \<union> B \<in> smallest_ccdi_sets \<Omega> M"
+proof -
+  have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets \<Omega> M)"
+    by (simp add: range_binaryset_eq  A B smallest_ccdi_sets.Basic)
+  have di:  "disjoint_family (binaryset A B)" using disj
+    by (simp add: disjoint_family_on_def binaryset_def Int_commute)
+  from Disj [OF ra di]
+  show ?thesis
+    by (simp add: UN_binaryset_eq)
+qed
+
+lemma (in algebra) smallest_ccdi_sets_Int1:
+  assumes a: "a \<in> M"
+  shows "b \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"
+proof (induct rule: smallest_ccdi_sets.induct)
+  case (Basic x)
+  thus ?case
+    by (metis a Int smallest_ccdi_sets.Basic)
+next
+  case (Compl x)
+  have "a \<inter> (\<Omega> - x) = \<Omega> - ((\<Omega> - a) \<union> (a \<inter> x))"
+    by blast
+  also have "... \<in> smallest_ccdi_sets \<Omega> M"
+    by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
+           Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un
+           smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl)
+  finally show ?case .
+next
+  case (Inc A)
+  have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
+    by blast
+  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc
+    by blast
+  moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
+    by (simp add: Inc)
+  moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc
+    by blast
+  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"
+    by (rule smallest_ccdi_sets.Inc)
+  show ?case
+    by (metis 1 2)
+next
+  case (Disj A)
+  have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
+    by blast
+  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj
+    by blast
+  moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj
+    by (auto simp add: disjoint_family_on_def)
+  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"
+    by (rule smallest_ccdi_sets.Disj)
+  show ?case
+    by (metis 1 2)
+qed
+
+
+lemma (in algebra) smallest_ccdi_sets_Int:
+  assumes b: "b \<in> smallest_ccdi_sets \<Omega> M"
+  shows "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"
+proof (induct rule: smallest_ccdi_sets.induct)
+  case (Basic x)
+  thus ?case
+    by (metis b smallest_ccdi_sets_Int1)
+next
+  case (Compl x)
+  have "(\<Omega> - x) \<inter> b = \<Omega> - (x \<inter> b \<union> (\<Omega> - b))"
+    by blast
+  also have "... \<in> smallest_ccdi_sets \<Omega> M"
+    by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
+           smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
+  finally show ?case .
+next
+  case (Inc A)
+  have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
+    by blast
+  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc
+    by blast
+  moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
+    by (simp add: Inc)
+  moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc
+    by blast
+  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"
+    by (rule smallest_ccdi_sets.Inc)
+  show ?case
+    by (metis 1 2)
+next
+  case (Disj A)
+  have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
+    by blast
+  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj
+    by blast
+  moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj
+    by (auto simp add: disjoint_family_on_def)
+  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"
+    by (rule smallest_ccdi_sets.Disj)
+  show ?case
+    by (metis 1 2)
+qed
+
+lemma (in algebra) sigma_property_disjoint_lemma:
+  assumes sbC: "M \<subseteq> C"
+      and ccdi: "closed_cdi \<Omega> C"
+  shows "sigma_sets \<Omega> M \<subseteq> C"
+proof -
+  have "smallest_ccdi_sets \<Omega> M \<in> {B . M \<subseteq> B \<and> sigma_algebra \<Omega> B}"
+    apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
+            smallest_ccdi_sets_Int)
+    apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
+    apply (blast intro: smallest_ccdi_sets.Disj)
+    done
+  hence "sigma_sets (\<Omega>) (M) \<subseteq> smallest_ccdi_sets \<Omega> M"
+    by clarsimp
+       (drule sigma_algebra.sigma_sets_subset [where a="M"], auto)
+  also have "...  \<subseteq> C"
+    proof
+      fix x
+      assume x: "x \<in> smallest_ccdi_sets \<Omega> M"
+      thus "x \<in> C"
+        proof (induct rule: smallest_ccdi_sets.induct)
+          case (Basic x)
+          thus ?case
+            by (metis Basic subsetD sbC)
+        next
+          case (Compl x)
+          thus ?case
+            by (blast intro: closed_cdi_Compl [OF ccdi, simplified])
+        next
+          case (Inc A)
+          thus ?case
+               by (auto intro: closed_cdi_Inc [OF ccdi, simplified])
+        next
+          case (Disj A)
+          thus ?case
+               by (auto intro: closed_cdi_Disj [OF ccdi, simplified])
+        qed
+    qed
+  finally show ?thesis .
+qed
+
+lemma (in algebra) sigma_property_disjoint:
+  assumes sbC: "M \<subseteq> C"
+      and compl: "!!s. s \<in> C \<inter> sigma_sets (\<Omega>) (M) \<Longrightarrow> \<Omega> - s \<in> C"
+      and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)
+                     \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
+                     \<Longrightarrow> (\<Union>i. A i) \<in> C"
+      and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)
+                      \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"
+  shows "sigma_sets (\<Omega>) (M) \<subseteq> C"
+proof -
+  have "sigma_sets (\<Omega>) (M) \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"
+    proof (rule sigma_property_disjoint_lemma)
+      show "M \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"
+        by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
+    next
+      show "closed_cdi \<Omega> (C \<inter> sigma_sets (\<Omega>) (M))"
+        by (simp add: closed_cdi_def compl inc disj)
+           (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
+             IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
+    qed
+  thus ?thesis
+    by blast
+qed
+
+subsubsection \<open>Dynkin systems\<close>
+
+locale dynkin_system = subset_class +
+  assumes space: "\<Omega> \<in> M"
+    and   compl[intro!]: "\<And>A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
+    and   UN[intro!]: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
+                           \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+
+lemma (in dynkin_system) empty[intro, simp]: "{} \<in> M"
+  using space compl[of "\<Omega>"] by simp
+
+lemma (in dynkin_system) diff:
+  assumes sets: "D \<in> M" "E \<in> M" and "D \<subseteq> E"
+  shows "E - D \<in> M"
+proof -
+  let ?f = "\<lambda>x. if x = 0 then D else if x = Suc 0 then \<Omega> - E else {}"
+  have "range ?f = {D, \<Omega> - E, {}}"
+    by (auto simp: image_iff)
+  moreover have "D \<union> (\<Omega> - E) = (\<Union>i. ?f i)"
+    by (auto simp: image_iff split: if_split_asm)
+  moreover
+  have "disjoint_family ?f" unfolding disjoint_family_on_def
+    using \<open>D \<in> M\<close>[THEN sets_into_space] \<open>D \<subseteq> E\<close> by auto
+  ultimately have "\<Omega> - (D \<union> (\<Omega> - E)) \<in> M"
+    using sets by auto
+  also have "\<Omega> - (D \<union> (\<Omega> - E)) = E - D"
+    using assms sets_into_space by auto
+  finally show ?thesis .
+qed
+
+lemma dynkin_systemI:
+  assumes "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>" "\<Omega> \<in> M"
+  assumes "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
+  assumes "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
+          \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  shows "dynkin_system \<Omega> M"
+  using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def)
+
+lemma dynkin_systemI':
+  assumes 1: "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>"
+  assumes empty: "{} \<in> M"
+  assumes Diff: "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
+  assumes 2: "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
+          \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  shows "dynkin_system \<Omega> M"
+proof -
+  from Diff[OF empty] have "\<Omega> \<in> M" by auto
+  from 1 this Diff 2 show ?thesis
+    by (intro dynkin_systemI) auto
+qed
+
+lemma dynkin_system_trivial:
+  shows "dynkin_system A (Pow A)"
+  by (rule dynkin_systemI) auto
+
+lemma sigma_algebra_imp_dynkin_system:
+  assumes "sigma_algebra \<Omega> M" shows "dynkin_system \<Omega> M"
+proof -
+  interpret sigma_algebra \<Omega> M by fact
+  show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI)
+qed
+
+subsubsection "Intersection sets systems"
+
+definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)"
+
+lemma (in algebra) Int_stable: "Int_stable M"
+  unfolding Int_stable_def by auto
+
+lemma Int_stableI:
+  "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A) \<Longrightarrow> Int_stable A"
+  unfolding Int_stable_def by auto
+
+lemma Int_stableD:
+  "Int_stable M \<Longrightarrow> a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"
+  unfolding Int_stable_def by auto
+
+lemma (in dynkin_system) sigma_algebra_eq_Int_stable:
+  "sigma_algebra \<Omega> M \<longleftrightarrow> Int_stable M"
+proof
+  assume "sigma_algebra \<Omega> M" then show "Int_stable M"
+    unfolding sigma_algebra_def using algebra.Int_stable by auto
+next
+  assume "Int_stable M"
+  show "sigma_algebra \<Omega> M"
+    unfolding sigma_algebra_disjoint_iff algebra_iff_Un
+  proof (intro conjI ballI allI impI)
+    show "M \<subseteq> Pow (\<Omega>)" using sets_into_space by auto
+  next
+    fix A B assume "A \<in> M" "B \<in> M"
+    then have "A \<union> B = \<Omega> - ((\<Omega> - A) \<inter> (\<Omega> - B))"
+              "\<Omega> - A \<in> M" "\<Omega> - B \<in> M"
+      using sets_into_space by auto
+    then show "A \<union> B \<in> M"
+      using \<open>Int_stable M\<close> unfolding Int_stable_def by auto
+  qed auto
+qed
+
+subsubsection "Smallest Dynkin systems"
+
+definition dynkin where
+  "dynkin \<Omega> M =  (\<Inter>{D. dynkin_system \<Omega> D \<and> M \<subseteq> D})"
+
+lemma dynkin_system_dynkin:
+  assumes "M \<subseteq> Pow (\<Omega>)"
+  shows "dynkin_system \<Omega> (dynkin \<Omega> M)"
+proof (rule dynkin_systemI)
+  fix A assume "A \<in> dynkin \<Omega> M"
+  moreover
+  { fix D assume "A \<in> D" and d: "dynkin_system \<Omega> D"
+    then have "A \<subseteq> \<Omega>" by (auto simp: dynkin_system_def subset_class_def) }
+  moreover have "{D. dynkin_system \<Omega> D \<and> M \<subseteq> D} \<noteq> {}"
+    using assms dynkin_system_trivial by fastforce
+  ultimately show "A \<subseteq> \<Omega>"
+    unfolding dynkin_def using assms
+    by auto
+next
+  show "\<Omega> \<in> dynkin \<Omega> M"
+    unfolding dynkin_def using dynkin_system.space by fastforce
+next
+  fix A assume "A \<in> dynkin \<Omega> M"
+  then show "\<Omega> - A \<in> dynkin \<Omega> M"
+    unfolding dynkin_def using dynkin_system.compl by force
+next
+  fix A :: "nat \<Rightarrow> 'a set"
+  assume A: "disjoint_family A" "range A \<subseteq> dynkin \<Omega> M"
+  show "(\<Union>i. A i) \<in> dynkin \<Omega> M" unfolding dynkin_def
+  proof (simp, safe)
+    fix D assume "dynkin_system \<Omega> D" "M \<subseteq> D"
+    with A have "(\<Union>i. A i) \<in> D"
+      by (intro dynkin_system.UN) (auto simp: dynkin_def)
+    then show "(\<Union>i. A i) \<in> D" by auto
+  qed
+qed
+
+lemma dynkin_Basic[intro]: "A \<in> M \<Longrightarrow> A \<in> dynkin \<Omega> M"
+  unfolding dynkin_def by auto
+
+lemma (in dynkin_system) restricted_dynkin_system:
+  assumes "D \<in> M"
+  shows "dynkin_system \<Omega> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
+proof (rule dynkin_systemI, simp_all)
+  have "\<Omega> \<inter> D = D"
+    using \<open>D \<in> M\<close> sets_into_space by auto
+  then show "\<Omega> \<inter> D \<in> M"
+    using \<open>D \<in> M\<close> by auto
+next
+  fix A assume "A \<subseteq> \<Omega> \<and> A \<inter> D \<in> M"
+  moreover have "(\<Omega> - A) \<inter> D = (\<Omega> - (A \<inter> D)) - (\<Omega> - D)"
+    by auto
+  ultimately show "\<Omega> - A \<subseteq> \<Omega> \<and> (\<Omega> - A) \<inter> D \<in> M"
+    using  \<open>D \<in> M\<close> by (auto intro: diff)
+next
+  fix A :: "nat \<Rightarrow> 'a set"
+  assume "disjoint_family A" "range A \<subseteq> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
+  then have "\<And>i. A i \<subseteq> \<Omega>" "disjoint_family (\<lambda>i. A i \<inter> D)"
+    "range (\<lambda>i. A i \<inter> D) \<subseteq> M" "(\<Union>x. A x) \<inter> D = (\<Union>x. A x \<inter> D)"
+    by ((fastforce simp: disjoint_family_on_def)+)
+  then show "(\<Union>x. A x) \<subseteq> \<Omega> \<and> (\<Union>x. A x) \<inter> D \<in> M"
+    by (auto simp del: UN_simps)
+qed
+
+lemma (in dynkin_system) dynkin_subset:
+  assumes "N \<subseteq> M"
+  shows "dynkin \<Omega> N \<subseteq> M"
+proof -
+  have "dynkin_system \<Omega> M" ..
+  then have "dynkin_system \<Omega> M"
+    using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp
+  with \<open>N \<subseteq> M\<close> show ?thesis by (auto simp add: dynkin_def)
+qed
+
+lemma sigma_eq_dynkin:
+  assumes sets: "M \<subseteq> Pow \<Omega>"
+  assumes "Int_stable M"
+  shows "sigma_sets \<Omega> M = dynkin \<Omega> M"
+proof -
+  have "dynkin \<Omega> M \<subseteq> sigma_sets (\<Omega>) (M)"
+    using sigma_algebra_imp_dynkin_system
+    unfolding dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto
+  moreover
+  interpret dynkin_system \<Omega> "dynkin \<Omega> M"
+    using dynkin_system_dynkin[OF sets] .
+  have "sigma_algebra \<Omega> (dynkin \<Omega> M)"
+    unfolding sigma_algebra_eq_Int_stable Int_stable_def
+  proof (intro ballI)
+    fix A B assume "A \<in> dynkin \<Omega> M" "B \<in> dynkin \<Omega> M"
+    let ?D = "\<lambda>E. {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> E \<in> dynkin \<Omega> M}"
+    have "M \<subseteq> ?D B"
+    proof
+      fix E assume "E \<in> M"
+      then have "M \<subseteq> ?D E" "E \<in> dynkin \<Omega> M"
+        using sets_into_space \<open>Int_stable M\<close> by (auto simp: Int_stable_def)
+      then have "dynkin \<Omega> M \<subseteq> ?D E"
+        using restricted_dynkin_system \<open>E \<in> dynkin \<Omega> M\<close>
+        by (intro dynkin_system.dynkin_subset) simp_all
+      then have "B \<in> ?D E"
+        using \<open>B \<in> dynkin \<Omega> M\<close> by auto
+      then have "E \<inter> B \<in> dynkin \<Omega> M"
+        by (subst Int_commute) simp
+      then show "E \<in> ?D B"
+        using sets \<open>E \<in> M\<close> by auto
+    qed
+    then have "dynkin \<Omega> M \<subseteq> ?D B"
+      using restricted_dynkin_system \<open>B \<in> dynkin \<Omega> M\<close>
+      by (intro dynkin_system.dynkin_subset) simp_all
+    then show "A \<inter> B \<in> dynkin \<Omega> M"
+      using \<open>A \<in> dynkin \<Omega> M\<close> sets_into_space by auto
+  qed
+  from sigma_algebra.sigma_sets_subset[OF this, of "M"]
+  have "sigma_sets (\<Omega>) (M) \<subseteq> dynkin \<Omega> M" by auto
+  ultimately have "sigma_sets (\<Omega>) (M) = dynkin \<Omega> M" by auto
+  then show ?thesis
+    by (auto simp: dynkin_def)
+qed
+
+lemma (in dynkin_system) dynkin_idem:
+  "dynkin \<Omega> M = M"
+proof -
+  have "dynkin \<Omega> M = M"
+  proof
+    show "M \<subseteq> dynkin \<Omega> M"
+      using dynkin_Basic by auto
+    show "dynkin \<Omega> M \<subseteq> M"
+      by (intro dynkin_subset) auto
+  qed
+  then show ?thesis
+    by (auto simp: dynkin_def)
+qed
+
+lemma (in dynkin_system) dynkin_lemma:
+  assumes "Int_stable E"
+  and E: "E \<subseteq> M" "M \<subseteq> sigma_sets \<Omega> E"
+  shows "sigma_sets \<Omega> E = M"
+proof -
+  have "E \<subseteq> Pow \<Omega>"
+    using E sets_into_space by force
+  then have *: "sigma_sets \<Omega> E = dynkin \<Omega> E"
+    using \<open>Int_stable E\<close> by (rule sigma_eq_dynkin)
+  then have "dynkin \<Omega> E = M"
+    using assms dynkin_subset[OF E(1)] by simp
+  with * show ?thesis
+    using assms by (auto simp: dynkin_def)
+qed
+
+subsubsection \<open>Induction rule for intersection-stable generators\<close>
+
+text \<open>The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras
+generated by a generator closed under intersection.\<close>
+
+lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:
+  assumes "Int_stable G"
+    and closed: "G \<subseteq> Pow \<Omega>"
+    and A: "A \<in> sigma_sets \<Omega> G"
+  assumes basic: "\<And>A. A \<in> G \<Longrightarrow> P A"
+    and empty: "P {}"
+    and compl: "\<And>A. A \<in> sigma_sets \<Omega> G \<Longrightarrow> P A \<Longrightarrow> P (\<Omega> - A)"
+    and union: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> sigma_sets \<Omega> G \<Longrightarrow> (\<And>i. P (A i)) \<Longrightarrow> P (\<Union>i::nat. A i)"
+  shows "P A"
+proof -
+  let ?D = "{ A \<in> sigma_sets \<Omega> G. P A }"
+  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> G"
+    using closed by (rule sigma_algebra_sigma_sets)
+  from compl[OF _ empty] closed have space: "P \<Omega>" by simp
+  interpret dynkin_system \<Omega> ?D
+    by standard (auto dest: sets_into_space intro!: space compl union)
+  have "sigma_sets \<Omega> G = ?D"
+    by (rule dynkin_lemma) (auto simp: basic \<open>Int_stable G\<close>)
+  with A show ?thesis by auto
+qed
+
+subsection \<open>Measure type\<close>
+
+definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+  "positive M \<mu> \<longleftrightarrow> \<mu> {} = 0"
+
+definition countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+  "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
+    (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
+
+definition measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
+  "measure_space \<Omega> A \<mu> \<longleftrightarrow> sigma_algebra \<Omega> A \<and> positive A \<mu> \<and> countably_additive A \<mu>"
+
+typedef 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
+proof
+  have "sigma_algebra UNIV {{}, UNIV}"
+    by (auto simp: sigma_algebra_iff2)
+  then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
+    by (auto simp: measure_space_def positive_def countably_additive_def)
+qed
+
+definition space :: "'a measure \<Rightarrow> 'a set" where
+  "space M = fst (Rep_measure M)"
+
+definition sets :: "'a measure \<Rightarrow> 'a set set" where
+  "sets M = fst (snd (Rep_measure M))"
+
+definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal" where
+  "emeasure M = snd (snd (Rep_measure M))"
+
+definition measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where
+  "measure M A = enn2real (emeasure M A)"
+
+declare [[coercion sets]]
+
+declare [[coercion measure]]
+
+declare [[coercion emeasure]]
+
+lemma measure_space: "measure_space (space M) (sets M) (emeasure M)"
+  by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse)
+
+interpretation sets: sigma_algebra "space M" "sets M" for M :: "'a measure"
+  using measure_space[of M] by (auto simp: measure_space_def)
+
+definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
+  "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
+    \<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"
+
+abbreviation "sigma \<Omega> A \<equiv> measure_of \<Omega> A (\<lambda>x. 0)"
+
+lemma measure_space_0: "A \<subseteq> Pow \<Omega> \<Longrightarrow> measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>x. 0)"
+  unfolding measure_space_def
+  by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)
+
+lemma sigma_algebra_trivial: "sigma_algebra \<Omega> {{}, \<Omega>}"
+by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\<Omega>}"])+
+
+lemma measure_space_0': "measure_space \<Omega> {{}, \<Omega>} (\<lambda>x. 0)"
+by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial)
+
+lemma measure_space_closed:
+  assumes "measure_space \<Omega> M \<mu>"
+  shows "M \<subseteq> Pow \<Omega>"
+proof -
+  interpret sigma_algebra \<Omega> M using assms by(simp add: measure_space_def)
+  show ?thesis by(rule space_closed)
+qed
+
+lemma (in ring_of_sets) positive_cong_eq:
+  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> positive M \<mu>' = positive M \<mu>"
+  by (auto simp add: positive_def)
+
+lemma (in sigma_algebra) countably_additive_eq:
+  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> countably_additive M \<mu>' = countably_additive M \<mu>"
+  unfolding countably_additive_def
+  by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)
+
+lemma measure_space_eq:
+  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a"
+  shows "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
+proof -
+  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" using closed by (rule sigma_algebra_sigma_sets)
+  from positive_cong_eq[OF eq, of "\<lambda>i. i"] countably_additive_eq[OF eq, of "\<lambda>i. i"] show ?thesis
+    by (auto simp: measure_space_def)
+qed
+
+lemma measure_of_eq:
+  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "(\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a)"
+  shows "measure_of \<Omega> A \<mu> = measure_of \<Omega> A \<mu>'"
+proof -
+  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
+    using assms by (rule measure_space_eq)
+  with eq show ?thesis
+    by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure])
+qed
+
+lemma
+  shows space_measure_of_conv: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)
+  and sets_measure_of_conv:
+  "sets (measure_of \<Omega> A \<mu>) = (if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>})" (is ?sets)
+  and emeasure_measure_of_conv:
+  "emeasure (measure_of \<Omega> A \<mu>) =
+  (\<lambda>B. if B \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> B else 0)" (is ?emeasure)
+proof -
+  have "?space \<and> ?sets \<and> ?emeasure"
+  proof(cases "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>")
+    case True
+    from measure_space_closed[OF this] sigma_sets_superset_generator[of A \<Omega>]
+    have "A \<subseteq> Pow \<Omega>" by simp
+    hence "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)
+      (\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"
+      by(rule measure_space_eq) auto
+    with True \<open>A \<subseteq> Pow \<Omega>\<close> show ?thesis
+      by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse)
+  next
+    case False thus ?thesis
+      by(cases "A \<subseteq> Pow \<Omega>")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0')
+  qed
+  thus ?space ?sets ?emeasure by simp_all
+qed
+
+lemma [simp]:
+  assumes A: "A \<subseteq> Pow \<Omega>"
+  shows sets_measure_of: "sets (measure_of \<Omega> A \<mu>) = sigma_sets \<Omega> A"
+    and space_measure_of: "space (measure_of \<Omega> A \<mu>) = \<Omega>"
+using assms
+by(simp_all add: sets_measure_of_conv space_measure_of_conv)
+
+lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \<Omega> M \<mu>) = M"
+  using space_closed by (auto intro!: sigma_sets_eq)
+
+lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \<Omega> M \<mu>) = \<Omega>"
+  by (rule space_measure_of_conv)
+
+lemma measure_of_subset: "M \<subseteq> Pow \<Omega> \<Longrightarrow> M' \<subseteq> M \<Longrightarrow> sets (measure_of \<Omega> M' \<mu>) \<subseteq> sets (measure_of \<Omega> M \<mu>')"
+  by (auto intro!: sigma_sets_subseteq)
+
+lemma emeasure_sigma: "emeasure (sigma \<Omega> A) = (\<lambda>x. 0)"
+  unfolding measure_of_def emeasure_def
+  by (subst Abs_measure_inverse)
+     (auto simp: measure_space_def positive_def countably_additive_def
+           intro!: sigma_algebra_sigma_sets sigma_algebra_trivial)
+
+lemma sigma_sets_mono'':
+  assumes "A \<in> sigma_sets C D"
+  assumes "B \<subseteq> D"
+  assumes "D \<subseteq> Pow C"
+  shows "sigma_sets A B \<subseteq> sigma_sets C D"
+proof
+  fix x assume "x \<in> sigma_sets A B"
+  thus "x \<in> sigma_sets C D"
+  proof induct
+    case (Basic a) with assms have "a \<in> D" by auto
+    thus ?case ..
+  next
+    case Empty show ?case by (rule sigma_sets.Empty)
+  next
+    from assms have "A \<in> sets (sigma C D)" by (subst sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
+    moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
+    ultimately have "A - a \<in> sets (sigma C D)" ..
+    thus ?case by (subst (asm) sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
+  next
+    case (Union a)
+    thus ?case by (intro sigma_sets.Union)
+  qed
+qed
+
+lemma in_measure_of[intro, simp]: "M \<subseteq> Pow \<Omega> \<Longrightarrow> A \<in> M \<Longrightarrow> A \<in> sets (measure_of \<Omega> M \<mu>)"
+  by auto
+
+lemma space_empty_iff: "space N = {} \<longleftrightarrow> sets N = {{}}"
+  by (metis Pow_empty Sup_bot_conv(1) cSup_singleton empty_iff
+            sets.sigma_sets_eq sets.space_closed sigma_sets_top subset_singletonD)
+
+subsubsection \<open>Constructing simple @{typ "'a measure"}\<close>
+
+lemma emeasure_measure_of:
+  assumes M: "M = measure_of \<Omega> A \<mu>"
+  assumes ms: "A \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>" "countably_additive (sets M) \<mu>"
+  assumes X: "X \<in> sets M"
+  shows "emeasure M X = \<mu> X"
+proof -
+  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" by (rule sigma_algebra_sigma_sets) fact
+  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"
+    using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets)
+  thus ?thesis using X ms
+    by(simp add: M emeasure_measure_of_conv sets_measure_of_conv)
+qed
+
+lemma emeasure_measure_of_sigma:
+  assumes ms: "sigma_algebra \<Omega> M" "positive M \<mu>" "countably_additive M \<mu>"
+  assumes A: "A \<in> M"
+  shows "emeasure (measure_of \<Omega> M \<mu>) A = \<mu> A"
+proof -
+  interpret sigma_algebra \<Omega> M by fact
+  have "measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+    using ms sigma_sets_eq by (simp add: measure_space_def)
+  thus ?thesis by(simp add: emeasure_measure_of_conv A)
+qed
+
+lemma measure_cases[cases type: measure]:
+  obtains (measure) \<Omega> A \<mu> where "x = Abs_measure (\<Omega>, A, \<mu>)" "\<forall>a\<in>-A. \<mu> a = 0" "measure_space \<Omega> A \<mu>"
+  by atomize_elim (cases x, auto)
+
+lemma sets_le_imp_space_le: "sets A \<subseteq> sets B \<Longrightarrow> space A \<subseteq> space B"
+  by (auto dest: sets.sets_into_space)
+
+lemma sets_eq_imp_space_eq: "sets M = sets M' \<Longrightarrow> space M = space M'"
+  by (auto intro!: antisym sets_le_imp_space_le)
+
+lemma emeasure_notin_sets: "A \<notin> sets M \<Longrightarrow> emeasure M A = 0"
+  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
+
+lemma emeasure_neq_0_sets: "emeasure M A \<noteq> 0 \<Longrightarrow> A \<in> sets M"
+  using emeasure_notin_sets[of A M] by blast
+
+lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"
+  by (simp add: measure_def emeasure_notin_sets zero_ennreal.rep_eq)
+
+lemma measure_eqI:
+  fixes M N :: "'a measure"
+  assumes "sets M = sets N" and eq: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A = emeasure N A"
+  shows "M = N"
+proof (cases M N rule: measure_cases[case_product measure_cases])
+  case (measure_measure \<Omega> A \<mu> \<Omega>' A' \<mu>')
+  interpret M: sigma_algebra \<Omega> A using measure_measure by (auto simp: measure_space_def)
+  interpret N: sigma_algebra \<Omega>' A' using measure_measure by (auto simp: measure_space_def)
+  have "A = sets M" "A' = sets N"
+    using measure_measure by (simp_all add: sets_def Abs_measure_inverse)
+  with \<open>sets M = sets N\<close> have AA': "A = A'" by simp
+  moreover from M.top N.top M.space_closed N.space_closed AA' have "\<Omega> = \<Omega>'" by auto
+  moreover { fix B have "\<mu> B = \<mu>' B"
+    proof cases
+      assume "B \<in> A"
+      with eq \<open>A = sets M\<close> have "emeasure M B = emeasure N B" by simp
+      with measure_measure show "\<mu> B = \<mu>' B"
+        by (simp add: emeasure_def Abs_measure_inverse)
+    next
+      assume "B \<notin> A"
+      with \<open>A = sets M\<close> \<open>A' = sets N\<close> \<open>A = A'\<close> have "B \<notin> sets M" "B \<notin> sets N"
+        by auto
+      then have "emeasure M B = 0" "emeasure N B = 0"
+        by (simp_all add: emeasure_notin_sets)
+      with measure_measure show "\<mu> B = \<mu>' B"
+        by (simp add: emeasure_def Abs_measure_inverse)
+    qed }
+  then have "\<mu> = \<mu>'" by auto
+  ultimately show "M = N"
+    by (simp add: measure_measure)
+qed
+
+lemma sigma_eqI:
+  assumes [simp]: "M \<subseteq> Pow \<Omega>" "N \<subseteq> Pow \<Omega>" "sigma_sets \<Omega> M = sigma_sets \<Omega> N"
+  shows "sigma \<Omega> M = sigma \<Omega> N"
+  by (rule measure_eqI) (simp_all add: emeasure_sigma)
+
+subsubsection \<open>Measurable functions\<close>
+
+definition measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr "\<rightarrow>\<^sub>M" 60) where
+  "measurable A B = {f \<in> space A \<rightarrow> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
+
+lemma measurableI:
+  "(\<And>x. x \<in> space M \<Longrightarrow> f x \<in> space N) \<Longrightarrow> (\<And>A. A \<in> sets N \<Longrightarrow> f -` A \<inter> space M \<in> sets M) \<Longrightarrow>
+    f \<in> measurable M N"
+  by (auto simp: measurable_def)
+
+lemma measurable_space:
+  "f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"
+   unfolding measurable_def by auto
+
+lemma measurable_sets:
+  "f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+   unfolding measurable_def by auto
+
+lemma measurable_sets_Collect:
+  assumes f: "f \<in> measurable M N" and P: "{x\<in>space N. P x} \<in> sets N" shows "{x\<in>space M. P (f x)} \<in> sets M"
+proof -
+  have "f -` {x \<in> space N. P x} \<inter> space M = {x\<in>space M. P (f x)}"
+    using measurable_space[OF f] by auto
+  with measurable_sets[OF f P] show ?thesis
+    by simp
+qed
+
+lemma measurable_sigma_sets:
+  assumes B: "sets N = sigma_sets \<Omega> A" "A \<subseteq> Pow \<Omega>"
+      and f: "f \<in> space M \<rightarrow> \<Omega>"
+      and ba: "\<And>y. y \<in> A \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M N"
+proof -
+  interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)
+  from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force
+
+  { fix X assume "X \<in> sigma_sets \<Omega> A"
+    then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"
+      proof induct
+        case (Basic a) then show ?case
+          by (auto simp add: ba) (metis B(2) subsetD PowD)
+      next
+        case (Compl a)
+        have [simp]: "f -` \<Omega> \<inter> space M = space M"
+          by (auto simp add: funcset_mem [OF f])
+        then show ?case
+          by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl)
+      next
+        case (Union a)
+        then show ?case
+          by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast
+      qed auto }
+  with f show ?thesis
+    by (auto simp add: measurable_def B \<Omega>)
+qed
+
+lemma measurable_measure_of:
+  assumes B: "N \<subseteq> Pow \<Omega>"
+      and f: "f \<in> space M \<rightarrow> \<Omega>"
+      and ba: "\<And>y. y \<in> N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>)"
+proof -
+  have "sets (measure_of \<Omega> N \<mu>) = sigma_sets \<Omega> N"
+    using B by (rule sets_measure_of)
+  from this assms show ?thesis by (rule measurable_sigma_sets)
+qed
+
+lemma measurable_iff_measure_of:
+  assumes "N \<subseteq> Pow \<Omega>" "f \<in> space M \<rightarrow> \<Omega>"
+  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>) \<longleftrightarrow> (\<forall>A\<in>N. f -` A \<inter> space M \<in> sets M)"
+  by (metis assms in_measure_of measurable_measure_of assms measurable_sets)
+
+lemma measurable_cong_sets:
+  assumes sets: "sets M = sets M'" "sets N = sets N'"
+  shows "measurable M N = measurable M' N'"
+  using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)
+
+lemma measurable_cong:
+  assumes "\<And>w. w \<in> space M \<Longrightarrow> f w = g w"
+  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
+  unfolding measurable_def using assms
+  by (simp cong: vimage_inter_cong Pi_cong)
+
+lemma measurable_cong':
+  assumes "\<And>w. w \<in> space M =simp=> f w = g w"
+  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
+  unfolding measurable_def using assms
+  by (simp cong: vimage_inter_cong Pi_cong add: simp_implies_def)
+
+lemma measurable_cong_strong:
+  "M = N \<Longrightarrow> M' = N' \<Longrightarrow> (\<And>w. w \<in> space M \<Longrightarrow> f w = g w) \<Longrightarrow>
+    f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable N N'"
+  by (metis measurable_cong)
+
+lemma measurable_compose:
+  assumes f: "f \<in> measurable M N" and g: "g \<in> measurable N L"
+  shows "(\<lambda>x. g (f x)) \<in> measurable M L"
+proof -
+  have "\<And>A. (\<lambda>x. g (f x)) -` A \<inter> space M = f -` (g -` A \<inter> space N) \<inter> space M"
+    using measurable_space[OF f] by auto
+  with measurable_space[OF f] measurable_space[OF g] show ?thesis
+    by (auto intro: measurable_sets[OF f] measurable_sets[OF g]
+             simp del: vimage_Int simp add: measurable_def)
+qed
+
+lemma measurable_comp:
+  "f \<in> measurable M N \<Longrightarrow> g \<in> measurable N L \<Longrightarrow> g \<circ> f \<in> measurable M L"
+  using measurable_compose[of f M N g L] by (simp add: comp_def)
+
+lemma measurable_const:
+  "c \<in> space M' \<Longrightarrow> (\<lambda>x. c) \<in> measurable M M'"
+  by (auto simp add: measurable_def)
+
+lemma measurable_ident: "id \<in> measurable M M"
+  by (auto simp add: measurable_def)
+
+lemma measurable_id: "(\<lambda>x. x) \<in> measurable M M"
+  by (simp add: measurable_def)
+
+lemma measurable_ident_sets:
+  assumes eq: "sets M = sets M'" shows "(\<lambda>x. x) \<in> measurable M M'"
+  using measurable_ident[of M]
+  unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] .
+
+lemma sets_Least:
+  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> M"
+  shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"
+proof -
+  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
+    proof cases
+      assume i: "(LEAST j. False) = i"
+      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
+        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
+        by (simp add: set_eq_iff, safe)
+           (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)
+      with meas show ?thesis
+        by (auto intro!: sets.Int)
+    next
+      assume i: "(LEAST j. False) \<noteq> i"
+      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
+        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
+      proof (simp add: set_eq_iff, safe)
+        fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"
+        have "\<exists>j. P j x"
+          by (rule ccontr) (insert neq, auto)
+        then show "P (LEAST j. P j x) x" by (rule LeastI_ex)
+      qed (auto dest: Least_le intro!: Least_equality)
+      with meas show ?thesis
+        by auto
+    qed }
+  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
+    by (intro sets.countable_UN) auto
+  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
+    (\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto
+  ultimately show ?thesis by auto
+qed
+
+lemma measurable_mono1:
+  "M' \<subseteq> Pow \<Omega> \<Longrightarrow> M \<subseteq> M' \<Longrightarrow>
+    measurable (measure_of \<Omega> M \<mu>) N \<subseteq> measurable (measure_of \<Omega> M' \<mu>') N"
+  using measure_of_subset[of M' \<Omega> M] by (auto simp add: measurable_def)
+
+subsubsection \<open>Counting space\<close>
+
+definition count_space :: "'a set \<Rightarrow> 'a measure" where
+  "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then of_nat (card A) else \<infinity>)"
+
+lemma
+  shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"
+    and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"
+  using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]
+  by (auto simp: count_space_def)
+
+lemma measurable_count_space_eq1[simp]:
+  "f \<in> measurable (count_space A) M \<longleftrightarrow> f \<in> A \<rightarrow> space M"
+ unfolding measurable_def by simp
+
+lemma measurable_compose_countable':
+  assumes f: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. f i x) \<in> measurable M N"
+  and g: "g \<in> measurable M (count_space I)" and I: "countable I"
+  shows "(\<lambda>x. f (g x) x) \<in> measurable M N"
+  unfolding measurable_def
+proof safe
+  fix x assume "x \<in> space M" then show "f (g x) x \<in> space N"
+    using measurable_space[OF f] g[THEN measurable_space] by auto
+next
+  fix A assume A: "A \<in> sets N"
+  have "(\<lambda>x. f (g x) x) -` A \<inter> space M = (\<Union>i\<in>I. (g -` {i} \<inter> space M) \<inter> (f i -` A \<inter> space M))"
+    using measurable_space[OF g] by auto
+  also have "\<dots> \<in> sets M"
+    using f[THEN measurable_sets, OF _ A] g[THEN measurable_sets]
+    by (auto intro!: sets.countable_UN' I intro: sets.Int[OF measurable_sets measurable_sets])
+  finally show "(\<lambda>x. f (g x) x) -` A \<inter> space M \<in> sets M" .
+qed
+
+lemma measurable_count_space_eq_countable:
+  assumes "countable A"
+  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
+proof -
+  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
+    with \<open>countable A\<close> have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
+      by (auto dest: countable_subset)
+    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
+    ultimately have "f -` X \<inter> space M \<in> sets M"
+      using \<open>X \<subseteq> A\<close> by (auto intro!: sets.countable_UN' simp del: UN_simps) }
+  then show ?thesis
+    unfolding measurable_def by auto
+qed
+
+lemma measurable_count_space_eq2:
+  "finite A \<Longrightarrow> f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
+  by (intro measurable_count_space_eq_countable countable_finite)
+
+lemma measurable_count_space_eq2_countable:
+  fixes f :: "'a => 'c::countable"
+  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
+  by (intro measurable_count_space_eq_countable countableI_type)
+
+lemma measurable_compose_countable:
+  assumes f: "\<And>i::'i::countable. (\<lambda>x. f i x) \<in> measurable M N" and g: "g \<in> measurable M (count_space UNIV)"
+  shows "(\<lambda>x. f (g x) x) \<in> measurable M N"
+  by (rule measurable_compose_countable'[OF assms]) auto
+
+lemma measurable_count_space_const:
+  "(\<lambda>x. c) \<in> measurable M (count_space UNIV)"
+  by (simp add: measurable_const)
+
+lemma measurable_count_space:
+  "f \<in> measurable (count_space A) (count_space UNIV)"
+  by simp
+
+lemma measurable_compose_rev:
+  assumes f: "f \<in> measurable L N" and g: "g \<in> measurable M L"
+  shows "(\<lambda>x. f (g x)) \<in> measurable M N"
+  using measurable_compose[OF g f] .
+
+lemma measurable_empty_iff:
+  "space N = {} \<Longrightarrow> f \<in> measurable M N \<longleftrightarrow> space M = {}"
+  by (auto simp add: measurable_def Pi_iff)
+
+subsubsection \<open>Extend measure\<close>
+
+definition "extend_measure \<Omega> I G \<mu> =
+  (if (\<exists>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>') \<and> \<not> (\<forall>i\<in>I. \<mu> i = 0)
+      then measure_of \<Omega> (G`I) (SOME \<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>')
+      else measure_of \<Omega> (G`I) (\<lambda>_. 0))"
+
+lemma space_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> space (extend_measure \<Omega> I G \<mu>) = \<Omega>"
+  unfolding extend_measure_def by simp
+
+lemma sets_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> sets (extend_measure \<Omega> I G \<mu>) = sigma_sets \<Omega> (G`I)"
+  unfolding extend_measure_def by simp
+
+lemma emeasure_extend_measure:
+  assumes M: "M = extend_measure \<Omega> I G \<mu>"
+    and eq: "\<And>i. i \<in> I \<Longrightarrow> \<mu>' (G i) = \<mu> i"
+    and ms: "G ` I \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
+    and "i \<in> I"
+  shows "emeasure M (G i) = \<mu> i"
+proof cases
+  assume *: "(\<forall>i\<in>I. \<mu> i = 0)"
+  with M have M_eq: "M = measure_of \<Omega> (G`I) (\<lambda>_. 0)"
+   by (simp add: extend_measure_def)
+  from measure_space_0[OF ms(1)] ms \<open>i\<in>I\<close>
+  have "emeasure M (G i) = 0"
+    by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure)
+  with \<open>i\<in>I\<close> * show ?thesis
+    by simp
+next
+  define P where "P \<mu>' \<longleftrightarrow> (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'" for \<mu>'
+  assume "\<not> (\<forall>i\<in>I. \<mu> i = 0)"
+  moreover
+  have "measure_space (space M) (sets M) \<mu>'"
+    using ms unfolding measure_space_def by auto standard
+  with ms eq have "\<exists>\<mu>'. P \<mu>'"
+    unfolding P_def
+    by (intro exI[of _ \<mu>']) (auto simp add: M space_extend_measure sets_extend_measure)
+  ultimately have M_eq: "M = measure_of \<Omega> (G`I) (Eps P)"
+    by (simp add: M extend_measure_def P_def[symmetric])
+
+  from \<open>\<exists>\<mu>'. P \<mu>'\<close> have P: "P (Eps P)" by (rule someI_ex)
+  show "emeasure M (G i) = \<mu> i"
+  proof (subst emeasure_measure_of[OF M_eq])
+    have sets_M: "sets M = sigma_sets \<Omega> (G`I)"
+      using M_eq ms by (auto simp: sets_extend_measure)
+    then show "G i \<in> sets M" using \<open>i \<in> I\<close> by auto
+    show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \<mu> i"
+      using P \<open>i\<in>I\<close> by (auto simp add: sets_M measure_space_def P_def)
+  qed fact
+qed
+
+lemma emeasure_extend_measure_Pair:
+  assumes M: "M = extend_measure \<Omega> {(i, j). I i j} (\<lambda>(i, j). G i j) (\<lambda>(i, j). \<mu> i j)"
+    and eq: "\<And>i j. I i j \<Longrightarrow> \<mu>' (G i j) = \<mu> i j"
+    and ms: "\<And>i j. I i j \<Longrightarrow> G i j \<in> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
+    and "I i j"
+  shows "emeasure M (G i j) = \<mu> i j"
+  using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) \<open>I i j\<close>
+  by (auto simp: subset_eq)
+
+subsection \<open>The smallest $\sigma$-algebra regarding a function\<close>
+
+definition
+  "vimage_algebra X f M = sigma X {f -` A \<inter> X | A. A \<in> sets M}"
+
+lemma space_vimage_algebra[simp]: "space (vimage_algebra X f M) = X"
+  unfolding vimage_algebra_def by (rule space_measure_of) auto
+
+lemma sets_vimage_algebra: "sets (vimage_algebra X f M) = sigma_sets X {f -` A \<inter> X | A. A \<in> sets M}"
+  unfolding vimage_algebra_def by (rule sets_measure_of) auto
+
+lemma sets_vimage_algebra2:
+  "f \<in> X \<rightarrow> space M \<Longrightarrow> sets (vimage_algebra X f M) = {f -` A \<inter> X | A. A \<in> sets M}"
+  using sigma_sets_vimage_commute[of f X "space M" "sets M"]
+  unfolding sets_vimage_algebra sets.sigma_sets_eq by simp
+
+lemma sets_vimage_algebra_cong: "sets M = sets N \<Longrightarrow> sets (vimage_algebra X f M) = sets (vimage_algebra X f N)"
+  by (simp add: sets_vimage_algebra)
+
+lemma vimage_algebra_cong:
+  assumes "X = Y"
+  assumes "\<And>x. x \<in> Y \<Longrightarrow> f x = g x"
+  assumes "sets M = sets N"
+  shows "vimage_algebra X f M = vimage_algebra Y g N"
+  by (auto simp: vimage_algebra_def assms intro!: arg_cong2[where f=sigma])
+
+lemma in_vimage_algebra: "A \<in> sets M \<Longrightarrow> f -` A \<inter> X \<in> sets (vimage_algebra X f M)"
+  by (auto simp: vimage_algebra_def)
+
+lemma sets_image_in_sets:
+  assumes N: "space N = X"
+  assumes f: "f \<in> measurable N M"
+  shows "sets (vimage_algebra X f M) \<subseteq> sets N"
+  unfolding sets_vimage_algebra N[symmetric]
+  by (rule sets.sigma_sets_subset) (auto intro!: measurable_sets f)
+
+lemma measurable_vimage_algebra1: "f \<in> X \<rightarrow> space M \<Longrightarrow> f \<in> measurable (vimage_algebra X f M) M"
+  unfolding measurable_def by (auto intro: in_vimage_algebra)
+
+lemma measurable_vimage_algebra2:
+  assumes g: "g \<in> space N \<rightarrow> X" and f: "(\<lambda>x. f (g x)) \<in> measurable N M"
+  shows "g \<in> measurable N (vimage_algebra X f M)"
+  unfolding vimage_algebra_def
+proof (rule measurable_measure_of)
+  fix A assume "A \<in> {f -` A \<inter> X | A. A \<in> sets M}"
+  then obtain Y where Y: "Y \<in> sets M" and A: "A = f -` Y \<inter> X"
+    by auto
+  then have "g -` A \<inter> space N = (\<lambda>x. f (g x)) -` Y \<inter> space N"
+    using g by auto
+  also have "\<dots> \<in> sets N"
+    using f Y by (rule measurable_sets)
+  finally show "g -` A \<inter> space N \<in> sets N" .
+qed (insert g, auto)
+
+lemma vimage_algebra_sigma:
+  assumes X: "X \<subseteq> Pow \<Omega>'" and f: "f \<in> \<Omega> \<rightarrow> \<Omega>'"
+  shows "vimage_algebra \<Omega> f (sigma \<Omega>' X) = sigma \<Omega> {f -` A \<inter> \<Omega> | A. A \<in> X }" (is "?V = ?S")
+proof (rule measure_eqI)
+  have \<Omega>: "{f -` A \<inter> \<Omega> |A. A \<in> X} \<subseteq> Pow \<Omega>" by auto
+  show "sets ?V = sets ?S"
+    using sigma_sets_vimage_commute[OF f, of X]
+    by (simp add: space_measure_of_conv f sets_vimage_algebra2 \<Omega> X)
+qed (simp add: vimage_algebra_def emeasure_sigma)
+
+lemma vimage_algebra_vimage_algebra_eq:
+  assumes *: "f \<in> X \<rightarrow> Y" "g \<in> Y \<rightarrow> space M"
+  shows "vimage_algebra X f (vimage_algebra Y g M) = vimage_algebra X (\<lambda>x. g (f x)) M"
+    (is "?VV = ?V")
+proof (rule measure_eqI)
+  have "(\<lambda>x. g (f x)) \<in> X \<rightarrow> space M" "\<And>A. A \<inter> f -` Y \<inter> X = A \<inter> X"
+    using * by auto
+  with * show "sets ?VV = sets ?V"
+    by (simp add: sets_vimage_algebra2 ex_simps[symmetric] vimage_comp comp_def del: ex_simps)
+qed (simp add: vimage_algebra_def emeasure_sigma)
+
+subsubsection \<open>Restricted Space Sigma Algebra\<close>
+
+definition restrict_space where
+  "restrict_space M \<Omega> = measure_of (\<Omega> \<inter> space M) ((op \<inter> \<Omega>) ` sets M) (emeasure M)"
+
+lemma space_restrict_space: "space (restrict_space M \<Omega>) = \<Omega> \<inter> space M"
+  using sets.sets_into_space unfolding restrict_space_def by (subst space_measure_of) auto
+
+lemma space_restrict_space2: "\<Omega> \<in> sets M \<Longrightarrow> space (restrict_space M \<Omega>) = \<Omega>"
+  by (simp add: space_restrict_space sets.sets_into_space)
+
+lemma sets_restrict_space: "sets (restrict_space M \<Omega>) = (op \<inter> \<Omega>) ` sets M"
+  unfolding restrict_space_def
+proof (subst sets_measure_of)
+  show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)"
+    by (auto dest: sets.sets_into_space)
+  have "sigma_sets (\<Omega> \<inter> space M) {((\<lambda>x. x) -` X) \<inter> (\<Omega> \<inter> space M) | X. X \<in> sets M} =
+    (\<lambda>X. X \<inter> (\<Omega> \<inter> space M)) ` sets M"
+    by (subst sigma_sets_vimage_commute[symmetric, where \<Omega>' = "space M"])
+       (auto simp add: sets.sigma_sets_eq)
+  moreover have "{((\<lambda>x. x) -` X) \<inter> (\<Omega> \<inter> space M) | X. X \<in> sets M} = (\<lambda>X. X \<inter> (\<Omega> \<inter> space M)) `  sets M"
+    by auto
+  moreover have "(\<lambda>X. X \<inter> (\<Omega> \<inter> space M)) `  sets M = (op \<inter> \<Omega>) ` sets M"
+    by (intro image_cong) (auto dest: sets.sets_into_space)
+  ultimately show "sigma_sets (\<Omega> \<inter> space M) (op \<inter> \<Omega> ` sets M) = op \<inter> \<Omega> ` sets M"
+    by simp
+qed
+
+lemma restrict_space_sets_cong:
+  "A = B \<Longrightarrow> sets M = sets N \<Longrightarrow> sets (restrict_space M A) = sets (restrict_space N B)"
+  by (auto simp: sets_restrict_space)
+
+lemma sets_restrict_space_count_space :
+  "sets (restrict_space (count_space A) B) = sets (count_space (A \<inter> B))"
+by(auto simp add: sets_restrict_space)
+
+lemma sets_restrict_UNIV[simp]: "sets (restrict_space M UNIV) = sets M"
+  by (auto simp add: sets_restrict_space)
+
+lemma sets_restrict_restrict_space:
+  "sets (restrict_space (restrict_space M A) B) = sets (restrict_space M (A \<inter> B))"
+  unfolding sets_restrict_space image_comp by (intro image_cong) auto
+
+lemma sets_restrict_space_iff:
+  "\<Omega> \<inter> space M \<in> sets M \<Longrightarrow> A \<in> sets (restrict_space M \<Omega>) \<longleftrightarrow> (A \<subseteq> \<Omega> \<and> A \<in> sets M)"
+proof (subst sets_restrict_space, safe)
+  fix A assume "\<Omega> \<inter> space M \<in> sets M" and A: "A \<in> sets M"
+  then have "(\<Omega> \<inter> space M) \<inter> A \<in> sets M"
+    by rule
+  also have "(\<Omega> \<inter> space M) \<inter> A = \<Omega> \<inter> A"
+    using sets.sets_into_space[OF A] by auto
+  finally show "\<Omega> \<inter> A \<in> sets M"
+    by auto
+qed auto
+
+lemma sets_restrict_space_cong: "sets M = sets N \<Longrightarrow> sets (restrict_space M \<Omega>) = sets (restrict_space N \<Omega>)"
+  by (simp add: sets_restrict_space)
+
+lemma restrict_space_eq_vimage_algebra:
+  "\<Omega> \<subseteq> space M \<Longrightarrow> sets (restrict_space M \<Omega>) = sets (vimage_algebra \<Omega> (\<lambda>x. x) M)"
+  unfolding restrict_space_def
+  apply (subst sets_measure_of)
+  apply (auto simp add: image_subset_iff dest: sets.sets_into_space) []
+  apply (auto simp add: sets_vimage_algebra intro!: arg_cong2[where f=sigma_sets])
+  done
+
+lemma sets_Collect_restrict_space_iff:
+  assumes "S \<in> sets M"
+  shows "{x\<in>space (restrict_space M S). P x} \<in> sets (restrict_space M S) \<longleftrightarrow> {x\<in>space M. x \<in> S \<and> P x} \<in> sets M"
+proof -
+  have "{x\<in>S. P x} = {x\<in>space M. x \<in> S \<and> P x}"
+    using sets.sets_into_space[OF assms] by auto
+  then show ?thesis
+    by (subst sets_restrict_space_iff) (auto simp add: space_restrict_space assms)
+qed
+
+lemma measurable_restrict_space1:
+  assumes f: "f \<in> measurable M N"
+  shows "f \<in> measurable (restrict_space M \<Omega>) N"
+  unfolding measurable_def
+proof (intro CollectI conjI ballI)
+  show sp: "f \<in> space (restrict_space M \<Omega>) \<rightarrow> space N"
+    using measurable_space[OF f] by (auto simp: space_restrict_space)
+
+  fix A assume "A \<in> sets N"
+  have "f -` A \<inter> space (restrict_space M \<Omega>) = (f -` A \<inter> space M) \<inter> (\<Omega> \<inter> space M)"
+    by (auto simp: space_restrict_space)
+  also have "\<dots> \<in> sets (restrict_space M \<Omega>)"
+    unfolding sets_restrict_space
+    using measurable_sets[OF f \<open>A \<in> sets N\<close>] by blast
+  finally show "f -` A \<inter> space (restrict_space M \<Omega>) \<in> sets (restrict_space M \<Omega>)" .
+qed
+
+lemma measurable_restrict_space2_iff:
+  "f \<in> measurable M (restrict_space N \<Omega>) \<longleftrightarrow> (f \<in> measurable M N \<and> f \<in> space M \<rightarrow> \<Omega>)"
+proof -
+  have "\<And>A. f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f -` \<Omega> \<inter> f -` A \<inter> space M = f -` A \<inter> space M"
+    by auto
+  then show ?thesis
+    by (auto simp: measurable_def space_restrict_space Pi_Int[symmetric] sets_restrict_space)
+qed
+
+lemma measurable_restrict_space2:
+  "f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> f \<in> measurable M (restrict_space N \<Omega>)"
+  by (simp add: measurable_restrict_space2_iff)
+
+lemma measurable_piecewise_restrict:
+  assumes I: "countable C"
+    and X: "\<And>\<Omega>. \<Omega> \<in> C \<Longrightarrow> \<Omega> \<inter> space M \<in> sets M" "space M \<subseteq> \<Union>C"
+    and f: "\<And>\<Omega>. \<Omega> \<in> C \<Longrightarrow> f \<in> measurable (restrict_space M \<Omega>) N"
+  shows "f \<in> measurable M N"
+proof (rule measurableI)
+  fix x assume "x \<in> space M"
+  with X obtain \<Omega> where "\<Omega> \<in> C" "x \<in> \<Omega>" "x \<in> space M" by auto
+  then show "f x \<in> space N"
+    by (auto simp: space_restrict_space intro: f measurable_space)
+next
+  fix A assume A: "A \<in> sets N"
+  have "f -` A \<inter> space M = (\<Union>\<Omega>\<in>C. (f -` A \<inter> (\<Omega> \<inter> space M)))"
+    using X by (auto simp: subset_eq)
+  also have "\<dots> \<in> sets M"
+    using measurable_sets[OF f A] X I
+    by (intro sets.countable_UN') (auto simp: sets_restrict_space_iff space_restrict_space)
+  finally show "f -` A \<inter> space M \<in> sets M" .
+qed
+
+lemma measurable_piecewise_restrict_iff:
+  "countable C \<Longrightarrow> (\<And>\<Omega>. \<Omega> \<in> C \<Longrightarrow> \<Omega> \<inter> space M \<in> sets M) \<Longrightarrow> space M \<subseteq> (\<Union>C) \<Longrightarrow>
+    f \<in> measurable M N \<longleftrightarrow> (\<forall>\<Omega>\<in>C. f \<in> measurable (restrict_space M \<Omega>) N)"
+  by (auto intro: measurable_piecewise_restrict measurable_restrict_space1)
+
+lemma measurable_If_restrict_space_iff:
+  "{x\<in>space M. P x} \<in> sets M \<Longrightarrow>
+    (\<lambda>x. if P x then f x else g x) \<in> measurable M N \<longleftrightarrow>
+    (f \<in> measurable (restrict_space M {x. P x}) N \<and> g \<in> measurable (restrict_space M {x. \<not> P x}) N)"
+  by (subst measurable_piecewise_restrict_iff[where C="{{x. P x}, {x. \<not> P x}}"])
+     (auto simp: Int_def sets.sets_Collect_neg space_restrict_space conj_commute[of _ "x \<in> space M" for x]
+           cong: measurable_cong')
+
+lemma measurable_If:
+  "f \<in> measurable M M' \<Longrightarrow> g \<in> measurable M M' \<Longrightarrow> {x\<in>space M. P x} \<in> sets M \<Longrightarrow>
+    (\<lambda>x. if P x then f x else g x) \<in> measurable M M'"
+  unfolding measurable_If_restrict_space_iff by (auto intro: measurable_restrict_space1)
+
+lemma measurable_If_set:
+  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
+  assumes P: "A \<inter> space M \<in> sets M"
+  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
+proof (rule measurable_If[OF measure])
+  have "{x \<in> space M. x \<in> A} = A \<inter> space M" by auto
+  thus "{x \<in> space M. x \<in> A} \<in> sets M" using \<open>A \<inter> space M \<in> sets M\<close> by auto
+qed
+
+lemma measurable_restrict_space_iff:
+  "\<Omega> \<inter> space M \<in> sets M \<Longrightarrow> c \<in> space N \<Longrightarrow>
+    f \<in> measurable (restrict_space M \<Omega>) N \<longleftrightarrow> (\<lambda>x. if x \<in> \<Omega> then f x else c) \<in> measurable M N"
+  by (subst measurable_If_restrict_space_iff)
+     (simp_all add: Int_def conj_commute measurable_const)
+
+lemma restrict_space_singleton: "{x} \<in> sets M \<Longrightarrow> sets (restrict_space M {x}) = sets (count_space {x})"
+  using sets_restrict_space_iff[of "{x}" M]
+  by (auto simp add: sets_restrict_space_iff dest!: subset_singletonD)
+
+lemma measurable_restrict_countable:
+  assumes X[intro]: "countable X"
+  assumes sets[simp]: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
+  assumes space[simp]: "\<And>x. x \<in> X \<Longrightarrow> f x \<in> space N"
+  assumes f: "f \<in> measurable (restrict_space M (- X)) N"
+  shows "f \<in> measurable M N"
+  using f sets.countable[OF sets X]
+  by (intro measurable_piecewise_restrict[where M=M and C="{- X} \<union> ((\<lambda>x. {x}) ` X)"])
+     (auto simp: Diff_Int_distrib2 Compl_eq_Diff_UNIV Int_insert_left sets.Diff restrict_space_singleton
+           simp del: sets_count_space  cong: measurable_cong_sets)
+
+lemma measurable_discrete_difference:
+  assumes f: "f \<in> measurable M N"
+  assumes X: "countable X" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" "\<And>x. x \<in> X \<Longrightarrow> g x \<in> space N"
+  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
+  shows "g \<in> measurable M N"
+  by (rule measurable_restrict_countable[OF X])
+     (auto simp: eq[symmetric] space_restrict_space cong: measurable_cong' intro: f measurable_restrict_space1)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/measurable.ML	Fri Aug 05 18:34:57 2016 +0200
@@ -0,0 +1,280 @@
+(*  Title:      HOL/Probability/measurable.ML
+    Author:     Johannes Hölzl <hoelzl@in.tum.de>
+
+Measurability prover.
+*)
+
+signature MEASURABLE = 
+sig
+  type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
+
+  datatype level = Concrete | Generic
+
+  val dest_thm_attr : attribute context_parser
+  val cong_thm_attr : attribute context_parser
+  val measurable_thm_attr : bool * (bool * level) -> attribute
+
+  val add_del_cong_thm : bool -> thm -> Context.generic -> Context.generic ;
+
+  val get_all : Context.generic -> thm list
+  val get_dest : Context.generic -> thm list
+  val get_cong : Context.generic -> thm list
+
+  val measurable_tac : Proof.context -> thm list -> tactic
+
+  val simproc : Proof.context -> cterm -> thm option
+
+  val add_preprocessor : string -> preprocessor -> Context.generic -> Context.generic
+  val del_preprocessor : string -> Context.generic -> Context.generic
+  val add_local_cong : thm -> Proof.context -> Proof.context
+
+  val prepare_facts : Proof.context -> thm list -> (thm list * Proof.context)
+end ;
+
+structure Measurable : MEASURABLE =
+struct
+
+type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
+
+datatype level = Concrete | Generic;
+
+fun eq_measurable_thms ((th1, d1), (th2, d2)) = 
+  d1 = d2 andalso Thm.eq_thm_prop (th1, th2) ;
+
+fun merge_dups (xs:(string * preprocessor) list) ys =
+  xs @ (filter (fn (name, _) => is_none (find_first (fn (name', _) => name' = name) xs)) ys) 
+
+structure Data = Generic_Data
+(
+  type T = {
+    measurable_thms : (thm * (bool * level)) Item_Net.T,
+    dest_thms : thm Item_Net.T,
+    cong_thms : thm Item_Net.T,
+    preprocessors : (string * preprocessor) list }
+  val empty: T = {
+    measurable_thms = Item_Net.init eq_measurable_thms (single o Thm.prop_of o fst),
+    dest_thms = Thm.full_rules,
+    cong_thms = Thm.full_rules,
+    preprocessors = [] };
+  val extend = I;
+  fun merge ({measurable_thms = t1, dest_thms = dt1, cong_thms = ct1, preprocessors = i1 },
+      {measurable_thms = t2, dest_thms = dt2, cong_thms = ct2, preprocessors = i2 }) : T = {
+    measurable_thms = Item_Net.merge (t1, t2),
+    dest_thms = Item_Net.merge (dt1, dt2),
+    cong_thms = Item_Net.merge (ct1, ct2),
+    preprocessors = merge_dups i1 i2 
+    };
+);
+
+val debug =
+  Attrib.setup_config_bool @{binding measurable_debug} (K false)
+
+val split =
+  Attrib.setup_config_bool @{binding measurable_split} (K true)
+
+fun map_data f1 f2 f3 f4
+  {measurable_thms = t1,    dest_thms = t2,    cong_thms = t3,    preprocessors = t4 } =
+  {measurable_thms = f1 t1, dest_thms = f2 t2, cong_thms = f3 t3, preprocessors = f4 t4}
+
+fun map_measurable_thms f = map_data f I I I
+fun map_dest_thms f = map_data I f I I
+fun map_cong_thms f = map_data I I f I
+fun map_preprocessors f = map_data I I I f
+
+fun generic_add_del map : attribute context_parser =
+  Scan.lift
+    (Args.add >> K Item_Net.update || Args.del >> K Item_Net.remove || Scan.succeed Item_Net.update) >>
+    (fn f => Thm.declaration_attribute (Data.map o map o f))
+
+val dest_thm_attr = generic_add_del map_dest_thms
+
+val cong_thm_attr = generic_add_del map_cong_thms
+
+fun del_thm th net =
+  let
+    val thms = net |> Item_Net.content |> filter (fn (th', _) => Thm.eq_thm (th, th'))
+  in fold Item_Net.remove thms net end ;
+
+fun measurable_thm_attr (do_add, d) = Thm.declaration_attribute
+  (Data.map o map_measurable_thms o (if do_add then Item_Net.update o rpair d else del_thm))
+
+val get_dest = Item_Net.content o #dest_thms o Data.get;
+
+val get_cong = Item_Net.content o #cong_thms o Data.get;
+val add_cong = Data.map o map_cong_thms o Item_Net.update;
+val del_cong = Data.map o map_cong_thms o Item_Net.remove;
+fun add_del_cong_thm true = add_cong
+  | add_del_cong_thm false = del_cong
+
+fun add_preprocessor name f = Data.map (map_preprocessors (fn xs => xs @ [(name, f)]))
+fun del_preprocessor name = Data.map (map_preprocessors (filter (fn (n, _) => n <> name)))
+val add_local_cong = Context.proof_map o add_cong
+
+val get_preprocessors = Context.Proof #> Data.get #> #preprocessors ;
+
+fun is_too_generic thm =
+  let 
+    val concl = Thm.concl_of thm
+    val concl' = HOLogic.dest_Trueprop concl handle TERM _ => concl
+  in is_Var (head_of concl') end
+
+val get_thms = Data.get #> #measurable_thms #> Item_Net.content ;
+
+val get_all = get_thms #> map fst ;
+
+fun debug_tac ctxt msg f =
+  if Config.get ctxt debug then print_tac ctxt (msg ()) THEN f else f
+
+fun nth_hol_goal thm i =
+  HOLogic.dest_Trueprop (Logic.strip_imp_concl (strip_all_body (nth (Thm.prems_of thm) (i - 1))))
+
+fun dest_measurable_fun t =
+  (case t of
+    (Const (@{const_name "Set.member"}, _) $ f $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => f
+  | _ => raise (TERM ("not a measurability predicate", [t])))
+
+fun not_measurable_prop n thm =
+  if length (Thm.prems_of thm) < n then false
+  else
+    (case nth_hol_goal thm n of
+      (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "sets"}, _) $ _)) => false
+    | (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => false
+    | _ => true)
+    handle TERM _ => true;
+
+fun indep (Bound i) t b = i < b orelse t <= i
+  | indep (f $ t) top bot = indep f top bot andalso indep t top bot
+  | indep (Abs (_,_,t)) top bot = indep t (top + 1) (bot + 1)
+  | indep _ _ _ = true;
+
+fun cnt_prefixes ctxt (Abs (n, T, t)) =
+    let
+      fun is_countable ty = Sign.of_sort (Proof_Context.theory_of ctxt) (ty, @{sort countable})
+      fun cnt_walk (Abs (ns, T, t)) Ts =
+          map (fn (t', t'') => (Abs (ns, T, t'), t'')) (cnt_walk t (T::Ts))
+        | cnt_walk (f $ g) Ts = let
+            val n = length Ts - 1
+          in
+            map (fn (f', t) => (f' $ g, t)) (cnt_walk f Ts) @
+            map (fn (g', t) => (f $ g', t)) (cnt_walk g Ts) @
+            (if is_countable (type_of1 (Ts, g)) andalso loose_bvar1 (g, n)
+                andalso indep g n 0 andalso g <> Bound n
+              then [(f $ Bound (n + 1), incr_boundvars (~ n) g)]
+              else [])
+          end
+        | cnt_walk _ _ = []
+    in map (fn (t1, t2) => let
+        val T1 = type_of1 ([T], t2)
+        val T2 = type_of1 ([T], t)
+      in ([SOME (Abs (n, T1, Abs (n, T, t1))), NONE, NONE, SOME (Abs (n, T, t2))],
+        [SOME T1, SOME T, SOME T2])
+      end) (cnt_walk t [T])
+    end
+  | cnt_prefixes _ _ = []
+
+fun apply_dests thm dests =
+  let
+    fun apply thm th' =
+      let
+        val th'' = thm RS th'
+      in [th''] @ loop th'' end
+      handle (THM _) => []
+    and loop thm =
+      flat (map (apply thm) dests)
+  in
+    [thm] @ ([thm RS @{thm measurable_compose_rev}] handle (THM _) => []) @ loop thm
+  end
+
+fun prepare_facts ctxt facts = 
+  let
+    val dests = get_dest (Context.Proof ctxt)
+    fun prep_dest thm =
+      (if is_too_generic thm then [] else apply_dests thm dests) ;
+    val preprocessors = (("std", prep_dest #> pair) :: get_preprocessors ctxt) ;
+    fun preprocess_thm (thm, raw) =
+      if raw then pair [thm] else fold_map (fn (_, proc) => proc thm) preprocessors #>> flat
+    
+    fun sel lv (th, (raw, lv')) = if lv = lv' then SOME (th, raw) else NONE ;
+    fun get lv = ctxt |> Context.Proof |> get_thms |> rev |> map_filter (sel lv) ;
+    val pre_thms = map (Simplifier.norm_hhf ctxt #> rpair false) facts @ get Concrete @ get Generic
+
+    val (thms, ctxt) = fold_map preprocess_thm pre_thms ctxt |>> flat
+  in (thms, ctxt) end
+
+fun measurable_tac ctxt facts =
+  let
+    fun debug_fact msg thm () =
+      msg ^ " " ^ Pretty.unformatted_string_of (Syntax.pretty_term ctxt (Thm.prop_of thm))
+
+    fun IF' c t i = COND (c i) (t i) no_tac
+
+    fun r_tac msg =
+      if Config.get ctxt debug
+      then FIRST' o
+        map (fn thm => resolve_tac ctxt [thm]
+          THEN' K (debug_tac ctxt (debug_fact (msg ^ " resolved using") thm) all_tac))
+      else resolve_tac ctxt
+
+    val elem_congI = @{lemma "A = B \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A" by simp}
+
+    val (thms, ctxt) = prepare_facts ctxt facts
+
+    fun is_sets_eq (Const (@{const_name "HOL.eq"}, _) $
+          (Const (@{const_name "sets"}, _) $ _) $
+          (Const (@{const_name "sets"}, _) $ _)) = true
+      | is_sets_eq (Const (@{const_name "HOL.eq"}, _) $
+          (Const (@{const_name "measurable"}, _) $ _ $ _) $
+          (Const (@{const_name "measurable"}, _) $ _ $ _)) = true
+      | is_sets_eq _ = false
+
+    val cong_thms = get_cong (Context.Proof ctxt) @
+      filter (fn thm => Thm.concl_of thm |> HOLogic.dest_Trueprop |> is_sets_eq handle TERM _ => false) facts
+
+    fun sets_cong_tac i =
+      Subgoal.FOCUS (fn {context = ctxt', prems = prems, ...} => (
+        let
+          val ctxt'' = Simplifier.add_prems prems ctxt'
+        in
+          r_tac "cong intro" [elem_congI]
+          THEN' SOLVED' (fn i => REPEAT_DETERM (
+              ((r_tac "cong solve" (cong_thms @ [@{thm refl}])
+                ORELSE' IF' (fn i => fn thm => Thm.nprems_of thm > i)
+                  (SOLVED' (asm_full_simp_tac ctxt''))) i)))
+        end) 1) ctxt i
+        THEN flexflex_tac ctxt
+
+    val simp_solver_tac = 
+      IF' not_measurable_prop (debug_tac ctxt (K "simp ") o SOLVED' (asm_full_simp_tac ctxt))
+
+    val split_countable_tac =
+      Subgoal.FOCUS (fn {context = ctxt, ...} => SUBGOAL (fn (t, i) =>
+        let
+          val f = dest_measurable_fun (HOLogic.dest_Trueprop t)
+          fun inst (ts, Ts) =
+            Thm.instantiate'
+              (map (Option.map (Thm.ctyp_of ctxt)) Ts)
+              (map (Option.map (Thm.cterm_of ctxt)) ts)
+              @{thm measurable_compose_countable}
+        in r_tac "case_prod countable" (cnt_prefixes ctxt f |> map inst) i end
+        handle TERM _ => no_tac) 1)
+
+    val splitter = if Config.get ctxt split then split_countable_tac ctxt else K no_tac
+
+    val single_step_tac =
+      simp_solver_tac
+      ORELSE' r_tac "step" thms
+      ORELSE' splitter
+      ORELSE' (CHANGED o sets_cong_tac)
+      ORELSE' (K (debug_tac ctxt (K "backtrack") no_tac))
+
+  in debug_tac ctxt (K "start") (REPEAT (single_step_tac 1)) end;
+
+fun simproc ctxt redex =
+  let
+    val t = HOLogic.mk_Trueprop (Thm.term_of redex);
+    fun tac {context = ctxt, prems = _ } =
+      SOLVE (measurable_tac ctxt (Simplifier.prems_of ctxt));
+  in try (fn () => Goal.prove ctxt [] [] t tac RS @{thm Eq_TrueI}) () end;
+
+end
+
--- a/src/HOL/Probability/Binary_Product_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1110 +0,0 @@
-(*  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
--- a/src/HOL/Probability/Bochner_Integration.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3066 +0,0 @@
-(*  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
--- a/src/HOL/Probability/Borel_Space.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1916 +0,0 @@
-(*  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
-  "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
-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
--- a/src/HOL/Probability/Caratheodory.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,891 +0,0 @@
-(*  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
--- a/src/HOL/Probability/Characteristic_Functions.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Characteristic_Functions.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -5,7 +5,7 @@
 section \<open>Characteristic Functions\<close>
 
 theory Characteristic_Functions
-  imports Weak_Convergence Interval_Integral Independent_Family Distributions
+  imports Weak_Convergence Independent_Family Distributions
 begin
 
 lemma mult_min_right: "a \<ge> 0 \<Longrightarrow> (a :: real) * min b c = min (a * b) (a * c)"
--- a/src/HOL/Probability/Complete_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,310 +0,0 @@
-(*  Title:      HOL/Probability/Complete_Measure.thy
-    Author:     Robert Himmelmann, Johannes Hoelzl, TU Muenchen
-*)
-
-theory Complete_Measure
-  imports Bochner_Integration Probability_Measure
-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 (in prob_space) prob_space_completion: "prob_space (completion M)"
-  by (rule prob_spaceI) (simp add: emeasure_space_1)
-
-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
--- a/src/HOL/Probability/Embed_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,398 +0,0 @@
-(*  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)"
-    by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1)
-  also have "\<dots> = \<integral>\<^sup>+ x. (SUP i:Y. ?f i x) \<partial>count_space UNIV"
-    by(simp add: nn_integral_count_space_indicator ennreal_indicator[symmetric] SUP_mult_right_ennreal nonempty)
-  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV)"
-  proof(rule nn_integral_monotone_convergence_SUP_nat)
-    show "Complete_Partial_Order.chain op \<le> (?f ` Y)"
-      by(rule chain_imageI[OF chain, unfolded image_image])(auto intro!: le_funI split: split_indicator dest: le_funD)
-  qed fact
-  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B x) \<partial>count_space (to_nat_on B ` B))"
-    by(simp add: nn_integral_count_space_indicator)
-  also have "\<dots> = (SUP i:Y. \<integral>\<^sup>+ x. f i (from_nat_into B (to_nat_on B x)) \<partial>count_space B)"
-    by(simp add: embed_measure_count_space'[symmetric] inj_on_to_nat_on countable nn_integral_embed_measure' measurable_embed_measure1)
-  also have "\<dots> = ?rhs"
-    by(intro arg_cong2[where f="SUPREMUM"] ext nn_integral_cong_AE)(simp_all add: AE_count_space countable)
-  finally show ?thesis .
-qed
-
-end
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1199 +0,0 @@
-(*  Title:      HOL/Probability/Finite_Product_Measure.thy
-    Author:     Johannes Hölzl, TU München
-*)
-
-section \<open>Finite product measures\<close>
-
-theory Finite_Product_Measure
-imports Binary_Product_Measure
-begin
-
-lemma PiE_choice: "(\<exists>f\<in>PiE I F. \<forall>i\<in>I. P i (f i)) \<longleftrightarrow> (\<forall>i\<in>I. \<exists>x\<in>F i. P i x)"
-  by (auto simp: Bex_def PiE_iff Ball_def dest!: choice_iff'[THEN iffD1])
-     (force intro: exI[of _ "restrict f I" for f])
-
-lemma case_prod_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
-  by auto
-
-subsubsection \<open>More about Function restricted by @{const extensional}\<close>
-
-definition
-  "merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
-
-lemma merge_apply[simp]:
-  "I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i"
-  "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i"
-  "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I J (x, y) i = x i"
-  "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I J (x, y) i = y i"
-  "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I J (x, y) i = undefined"
-  unfolding merge_def by auto
-
-lemma merge_commute:
-  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)"
-  by (force simp: merge_def)
-
-lemma Pi_cancel_merge_range[simp]:
-  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A"
-  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A"
-  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I J (A, B)) \<longleftrightarrow> x \<in> Pi I A"
-  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J I (B, A)) \<longleftrightarrow> x \<in> Pi I A"
-  by (auto simp: Pi_def)
-
-lemma Pi_cancel_merge[simp]:
-  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
-  "J \<inter> I = {} \<Longrightarrow> merge I J (x, y) \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
-  by (auto simp: Pi_def)
-
-lemma extensional_merge[simp]: "merge I J (x, y) \<in> extensional (I \<union> J)"
-  by (auto simp: extensional_def)
-
-lemma restrict_merge[simp]:
-  "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I"
-  "I \<inter> J = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
-  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) I = restrict x I"
-  "J \<inter> I = {} \<Longrightarrow> restrict (merge I J (x, y)) J = restrict y J"
-  by (auto simp: restrict_def)
-
-lemma split_merge: "P (merge I J (x,y) i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)"
-  unfolding merge_def by auto
-
-lemma PiE_cancel_merge[simp]:
-  "I \<inter> J = {} \<Longrightarrow>
-    merge I J (x, y) \<in> PiE (I \<union> J) B \<longleftrightarrow> x \<in> Pi I B \<and> y \<in> Pi J B"
-  by (auto simp: PiE_def restrict_Pi_cancel)
-
-lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I {i} (x,y) = restrict (x(i := y i)) (insert i I)"
-  unfolding merge_def by (auto simp: fun_eq_iff)
-
-lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I J (x, y) \<in> extensional K"
-  unfolding merge_def extensional_def by auto
-
-lemma merge_restrict[simp]:
-  "merge I J (restrict x I, y) = merge I J (x, y)"
-  "merge I J (x, restrict y J) = merge I J (x, y)"
-  unfolding merge_def by auto
-
-lemma merge_x_x_eq_restrict[simp]:
-  "merge I J (x, x) = restrict x (I \<union> J)"
-  unfolding merge_def by auto
-
-lemma injective_vimage_restrict:
-  assumes J: "J \<subseteq> I"
-  and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}"
-  and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
-  shows "A = B"
-proof  (intro set_eqI)
-  fix x
-  from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto
-  have "J \<inter> (I - J) = {}" by auto
-  show "x \<in> A \<longleftrightarrow> x \<in> B"
-  proof cases
-    assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
-    have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
-      using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
-      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
-    then show "x \<in> A \<longleftrightarrow> x \<in> B"
-      using y x \<open>J \<subseteq> I\<close> PiE_cancel_merge[of "J" "I - J" x y S]
-      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1 eq)
-  qed (insert sets, auto)
-qed
-
-lemma restrict_vimage:
-  "I \<inter> J = {} \<Longrightarrow>
-    (\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^sub>E I E \<times> Pi\<^sub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
-  by (auto simp: restrict_Pi_cancel PiE_def)
-
-lemma merge_vimage:
-  "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
-  by (auto simp: restrict_Pi_cancel PiE_def)
-
-subsection \<open>Finite product spaces\<close>
-
-subsubsection \<open>Products\<close>
-
-definition prod_emb where
-  "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
-
-lemma prod_emb_iff:
-  "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
-  unfolding prod_emb_def PiE_def by auto
-
-lemma
-  shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
-    and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B"
-    and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B"
-    and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))"
-    and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))"
-    and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
-  by (auto simp: prod_emb_def)
-
-lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
-    prod_emb I M J (\<Pi>\<^sub>E i\<in>J. E i) = (\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i))"
-  by (force simp: prod_emb_def PiE_iff if_split_mem2)
-
-lemma prod_emb_PiE_same_index[simp]:
-    "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^sub>E I E) = Pi\<^sub>E I E"
-  by (auto simp: prod_emb_def PiE_iff)
-
-lemma prod_emb_trans[simp]:
-  "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
-  by (auto simp add: Int_absorb1 prod_emb_def PiE_def)
-
-lemma prod_emb_Pi:
-  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
-  shows "prod_emb K M J (Pi\<^sub>E J X) = (\<Pi>\<^sub>E i\<in>K. if i \<in> J then X i else space (M i))"
-  using assms sets.space_closed
-  by (auto simp: prod_emb_def PiE_iff split: if_split_asm) blast+
-
-lemma prod_emb_id:
-  "B \<subseteq> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
-  by (auto simp: prod_emb_def subset_eq extensional_restrict)
-
-lemma prod_emb_mono:
-  "F \<subseteq> G \<Longrightarrow> prod_emb A M B F \<subseteq> prod_emb A M B G"
-  by (auto simp: prod_emb_def)
-
-definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
-  "PiM I M = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
-    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
-    (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
-    (\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
-
-definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
-  "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j)) `
-    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
-
-abbreviation
-  "Pi\<^sub>M I M \<equiv> PiM I M"
-
-syntax
-  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^sub>M _\<in>_./ _)"  10)
-translations
-  "\<Pi>\<^sub>M x\<in>I. M" == "CONST PiM I (%x. M)"
-
-lemma extend_measure_cong:
-  assumes "\<Omega> = \<Omega>'" "I = I'" "G = G'" "\<And>i. i \<in> I' \<Longrightarrow> \<mu> i = \<mu>' i"
-  shows "extend_measure \<Omega> I G \<mu> = extend_measure \<Omega>' I' G' \<mu>'"
-  unfolding extend_measure_def by (auto simp add: assms)
-
-lemma Pi_cong_sets:
-    "\<lbrakk>I = J; \<And>x. x \<in> I \<Longrightarrow> M x = N x\<rbrakk> \<Longrightarrow> Pi I M = Pi J N"
-  unfolding Pi_def by auto
-
-lemma PiM_cong:
-  assumes "I = J" "\<And>x. x \<in> I \<Longrightarrow> M x = N x"
-  shows "PiM I M = PiM J N"
-  unfolding PiM_def
-proof (rule extend_measure_cong, goal_cases)
-  case 1
-  show ?case using assms
-    by (subst assms(1), intro PiE_cong[of J "\<lambda>i. space (M i)" "\<lambda>i. space (N i)"]) simp_all
-next
-  case 2
-  have "\<And>K. K \<subseteq> J \<Longrightarrow> (\<Pi> j\<in>K. sets (M j)) = (\<Pi> j\<in>K. sets (N j))"
-    using assms by (intro Pi_cong_sets) auto
-  thus ?case by (auto simp: assms)
-next
-  case 3
-  show ?case using assms
-    by (intro ext) (auto simp: prod_emb_def dest: PiE_mem)
-next
-  case (4 x)
-  thus ?case using assms
-    by (auto intro!: setprod.cong split: if_split_asm)
-qed
-
-
-lemma prod_algebra_sets_into_space:
-  "prod_algebra I M \<subseteq> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  by (auto simp: prod_emb_def prod_algebra_def)
-
-lemma prod_algebra_eq_finite:
-  assumes I: "finite I"
-  shows "prod_algebra I M = {(\<Pi>\<^sub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
-proof (intro iffI set_eqI)
-  fix A assume "A \<in> ?L"
-  then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
-    and A: "A = prod_emb I M J (PIE j:J. E j)"
-    by (auto simp: prod_algebra_def)
-  let ?A = "\<Pi>\<^sub>E i\<in>I. if i \<in> J then E i else space (M i)"
-  have A: "A = ?A"
-    unfolding A using J by (intro prod_emb_PiE sets.sets_into_space) auto
-  show "A \<in> ?R" unfolding A using J sets.top
-    by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp
-next
-  fix A assume "A \<in> ?R"
-  then obtain X where A: "A = (\<Pi>\<^sub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
-  then have A: "A = prod_emb I M I (\<Pi>\<^sub>E i\<in>I. X i)"
-    by (simp add: prod_emb_PiE_same_index[OF sets.sets_into_space] Pi_iff)
-  from X I show "A \<in> ?L" unfolding A
-    by (auto simp: prod_algebra_def)
-qed
-
-lemma prod_algebraI:
-  "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
-    \<Longrightarrow> prod_emb I M J (PIE j:J. E j) \<in> prod_algebra I M"
-  by (auto simp: prod_algebra_def)
-
-lemma prod_algebraI_finite:
-  "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>E I E) \<in> prod_algebra I M"
-  using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets.sets_into_space] by simp
-
-lemma Int_stable_PiE: "Int_stable {Pi\<^sub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
-proof (safe intro!: Int_stableI)
-  fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
-  then show "\<exists>G. Pi\<^sub>E J E \<inter> Pi\<^sub>E J F = Pi\<^sub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
-    by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"] simp: PiE_Int)
-qed
-
-lemma prod_algebraE:
-  assumes A: "A \<in> prod_algebra I M"
-  obtains J E where "A = prod_emb I M J (PIE j:J. E j)"
-    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)"
-  using A by (auto simp: prod_algebra_def)
-
-lemma prod_algebraE_all:
-  assumes A: "A \<in> prod_algebra I M"
-  obtains E where "A = Pi\<^sub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
-proof -
-  from A obtain E J where A: "A = prod_emb I M J (Pi\<^sub>E J E)"
-    and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
-    by (auto simp: prod_algebra_def)
-  from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)"
-    using sets.sets_into_space by auto
-  then have "A = (\<Pi>\<^sub>E i\<in>I. if i\<in>J then E i else space (M i))"
-    using A J by (auto simp: prod_emb_PiE)
-  moreover have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))"
-    using sets.top E by auto
-  ultimately show ?thesis using that by auto
-qed
-
-lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
-proof (unfold Int_stable_def, safe)
-  fix A assume "A \<in> prod_algebra I M"
-  from prod_algebraE[OF this] guess J E . note A = this
-  fix B assume "B \<in> prod_algebra I M"
-  from prod_algebraE[OF this] guess K F . note B = this
-  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^sub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter>
-      (if i \<in> K then F i else space (M i)))"
-    unfolding A B using A(2,3,4) A(5)[THEN sets.sets_into_space] B(2,3,4)
-      B(5)[THEN sets.sets_into_space]
-    apply (subst (1 2 3) prod_emb_PiE)
-    apply (simp_all add: subset_eq PiE_Int)
-    apply blast
-    apply (intro PiE_cong)
-    apply auto
-    done
-  also have "\<dots> \<in> prod_algebra I M"
-    using A B by (auto intro!: prod_algebraI)
-  finally show "A \<inter> B \<in> prod_algebra I M" .
-qed
-
-lemma prod_algebra_mono:
-  assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
-  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
-  shows "prod_algebra I E \<subseteq> prod_algebra I F"
-proof
-  fix A assume "A \<in> prod_algebra I E"
-  then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
-    and A: "A = prod_emb I E J (\<Pi>\<^sub>E i\<in>J. G i)"
-    and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)"
-    by (auto simp: prod_algebra_def)
-  moreover
-  from space have "(\<Pi>\<^sub>E i\<in>I. space (E i)) = (\<Pi>\<^sub>E i\<in>I. space (F i))"
-    by (rule PiE_cong)
-  with A have "A = prod_emb I F J (\<Pi>\<^sub>E i\<in>J. G i)"
-    by (simp add: prod_emb_def)
-  moreover
-  from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)"
-    by auto
-  ultimately show "A \<in> prod_algebra I F"
-    apply (simp add: prod_algebra_def image_iff)
-    apply (intro exI[of _ J] exI[of _ G] conjI)
-    apply auto
-    done
-qed
-
-lemma prod_algebra_cong:
-  assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
-  shows "prod_algebra I M = prod_algebra J N"
-proof -
-  have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
-    using sets_eq_imp_space_eq[OF sets] by auto
-  with sets show ?thesis unfolding \<open>I = J\<close>
-    by (intro antisym prod_algebra_mono) auto
-qed
-
-lemma space_in_prod_algebra:
-  "(\<Pi>\<^sub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
-proof cases
-  assume "I = {}" then show ?thesis
-    by (auto simp add: prod_algebra_def image_iff prod_emb_def)
-next
-  assume "I \<noteq> {}"
-  then obtain i where "i \<in> I" by auto
-  then have "(\<Pi>\<^sub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
-    by (auto simp: prod_emb_def)
-  also have "\<dots> \<in> prod_algebra I M"
-    using \<open>i \<in> I\<close> by (intro prod_algebraI) auto
-  finally show ?thesis .
-qed
-
-lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
-
-lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)"
-  by (auto simp: prod_emb_def space_PiM)
-
-lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow>  (\<exists>i\<in>I. space (M i) = {})"
-  by (auto simp: space_PiM PiE_eq_empty_iff)
-
-lemma undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)"
-  by (auto simp: space_PiM)
-
-lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
-  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
-
-lemma sets_PiM_single: "sets (PiM I M) =
-    sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
-    (is "_ = sigma_sets ?\<Omega> ?R")
-  unfolding sets_PiM
-proof (rule sigma_sets_eqI)
-  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
-  fix A assume "A \<in> prod_algebra I M"
-  from prod_algebraE[OF this] guess J X . note X = this
-  show "A \<in> sigma_sets ?\<Omega> ?R"
-  proof cases
-    assume "I = {}"
-    with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def)
-    with \<open>I = {}\<close> show ?thesis by (auto intro!: sigma_sets_top)
-  next
-    assume "I \<noteq> {}"
-    with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^sub>E i\<in>I. space (M i)). f j \<in> X j})"
-      by (auto simp: prod_emb_def)
-    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
-      using X \<open>I \<noteq> {}\<close> by (intro R.finite_INT sigma_sets.Basic) auto
-    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
-  qed
-next
-  fix A assume "A \<in> ?R"
-  then obtain i B where A: "A = {f\<in>\<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)"
-    by auto
-  then have "A = prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. B)"
-     by (auto simp: prod_emb_def)
-  also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)"
-    using A by (intro sigma_sets.Basic prod_algebraI) auto
-  finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" .
-qed
-
-lemma sets_PiM_eq_proj:
-  "I \<noteq> {} \<Longrightarrow> sets (PiM I M) = sets (SUP i:I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) (\<lambda>x. x i) (M i))"
-  apply (simp add: sets_PiM_single)
-  apply (subst sets_Sup_eq[where X="\<Pi>\<^sub>E i\<in>I. space (M i)"])
-  apply auto []
-  apply auto []
-  apply simp
-  apply (subst SUP_cong[OF refl])
-  apply (rule sets_vimage_algebra2)
-  apply auto []
-  apply (auto intro!: arg_cong2[where f=sigma_sets])
-  done
-
-lemma
-  shows space_PiM_empty: "space (Pi\<^sub>M {} M) = {\<lambda>k. undefined}"
-    and sets_PiM_empty: "sets (Pi\<^sub>M {} M) = { {}, {\<lambda>k. undefined} }"
-  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
-
-lemma sets_PiM_sigma:
-  assumes \<Omega>_cover: "\<And>i. i \<in> I \<Longrightarrow> \<exists>S\<subseteq>E i. countable S \<and> \<Omega> i = \<Union>S"
-  assumes E: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (\<Omega> i)"
-  assumes J: "\<And>j. j \<in> J \<Longrightarrow> finite j" "\<Union>J = I"
-  defines "P \<equiv> {{f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i} | A j. j \<in> J \<and> A \<in> Pi j E}"
-  shows "sets (\<Pi>\<^sub>M i\<in>I. sigma (\<Omega> i) (E i)) = sets (sigma (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P)"
-proof cases
-  assume "I = {}"
-  with \<open>\<Union>J = I\<close> have "P = {{\<lambda>_. undefined}} \<or> P = {}"
-    by (auto simp: P_def)
-  with \<open>I = {}\<close> show ?thesis
-    by (auto simp add: sets_PiM_empty sigma_sets_empty_eq)
-next
-  let ?F = "\<lambda>i. {(\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega> |A. A \<in> E i}"
-  assume "I \<noteq> {}"
-  then have "sets (Pi\<^sub>M I (\<lambda>i. sigma (\<Omega> i) (E i))) =
-      sets (SUP i:I. vimage_algebra (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<lambda>x. x i) (sigma (\<Omega> i) (E i)))"
-    by (subst sets_PiM_eq_proj) (auto simp: space_measure_of_conv)
-  also have "\<dots> = sets (SUP i:I. sigma (Pi\<^sub>E I \<Omega>) (?F i))"
-    using E by (intro sets_SUP_cong arg_cong[where f=sets] vimage_algebra_sigma) auto
-  also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i))"
-    using \<open>I \<noteq> {}\<close> by (intro arg_cong[where f=sets] SUP_sigma_sigma) auto
-  also have "\<dots> = sets (sigma (Pi\<^sub>E I \<Omega>) P)"
-  proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)
-    show "(\<Union>i\<in>I. ?F i) \<subseteq> Pow (Pi\<^sub>E I \<Omega>)" "P \<subseteq> Pow (Pi\<^sub>E I \<Omega>)"
-      by (auto simp: P_def)
-  next
-    interpret P: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
-      by (auto intro!: sigma_algebra_sigma_sets simp: P_def)
-
-    fix Z assume "Z \<in> (\<Union>i\<in>I. ?F i)"
-    then obtain i A where i: "i \<in> I" "A \<in> E i" and Z_def: "Z = (\<lambda>x. x i) -` A \<inter> Pi\<^sub>E I \<Omega>"
-      by auto
-    from \<open>i \<in> I\<close> J obtain j where j: "i \<in> j" "j \<in> J" "j \<subseteq> I" "finite j"
-      by auto
-    obtain S where S: "\<And>i. i \<in> j \<Longrightarrow> S i \<subseteq> E i" "\<And>i. i \<in> j \<Longrightarrow> countable (S i)"
-      "\<And>i. i \<in> j \<Longrightarrow> \<Omega> i = \<Union>(S i)"
-      by (metis subset_eq \<Omega>_cover \<open>j \<subseteq> I\<close>)
-    define A' where "A' n = n(i := A)" for n
-    then have A'_i: "\<And>n. A' n i = A"
-      by simp
-    { fix n assume "n \<in> Pi\<^sub>E (j - {i}) S"
-      then have "A' n \<in> Pi j E"
-        unfolding PiE_def Pi_def using S(1) by (auto simp: A'_def \<open>A \<in> E i\<close> )
-      with \<open>j \<in> J\<close> have "{f \<in> Pi\<^sub>E I \<Omega>. \<forall>i\<in>j. f i \<in> A' n i} \<in> P"
-        by (auto simp: P_def) }
-    note A'_in_P = this
-
-    { fix x assume "x i \<in> A" "x \<in> Pi\<^sub>E I \<Omega>"
-      with S(3) \<open>j \<subseteq> I\<close> have "\<forall>i\<in>j. \<exists>s\<in>S i. x i \<in> s"
-        by (auto simp: PiE_def Pi_def)
-      then obtain s where s: "\<And>i. i \<in> j \<Longrightarrow> s i \<in> S i" "\<And>i. i \<in> j \<Longrightarrow> x i \<in> s i"
-        by metis
-      with \<open>x i \<in> A\<close> have "\<exists>n\<in>PiE (j-{i}) S. \<forall>i\<in>j. x i \<in> A' n i"
-        by (intro bexI[of _ "restrict (s(i := A)) (j-{i})"]) (auto simp: A'_def split: if_splits) }
-    then have "Z = (\<Union>n\<in>PiE (j-{i}) S. {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A' n i})"
-      unfolding Z_def
-      by (auto simp add: set_eq_iff ball_conj_distrib \<open>i\<in>j\<close> A'_i dest: bspec[OF _ \<open>i\<in>j\<close>]
-               cong: conj_cong)
-    also have "\<dots> \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P"
-      using \<open>finite j\<close> S(2)
-      by (intro P.countable_UN' countable_PiE) (simp_all add: image_subset_iff A'_in_P)
-    finally show "Z \<in> sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) P" .
-  next
-    interpret F: sigma_algebra "\<Pi>\<^sub>E i\<in>I. \<Omega> i" "sigma_sets (\<Pi>\<^sub>E i\<in>I. \<Omega> i) (\<Union>i\<in>I. ?F i)"
-      by (auto intro!: sigma_algebra_sigma_sets)
-
-    fix b assume "b \<in> P"
-    then obtain A j where b: "b = {f\<in>(\<Pi>\<^sub>E i\<in>I. \<Omega> i). \<forall>i\<in>j. f i \<in> A i}" "j \<in> J" "A \<in> Pi j E"
-      by (auto simp: P_def)
-    show "b \<in> sigma_sets (Pi\<^sub>E I \<Omega>) (\<Union>i\<in>I. ?F i)"
-    proof cases
-      assume "j = {}"
-      with b have "b = (\<Pi>\<^sub>E i\<in>I. \<Omega> i)"
-        by auto
-      then show ?thesis
-        by blast
-    next
-      assume "j \<noteq> {}"
-      with J b(2,3) have eq: "b = (\<Inter>i\<in>j. ((\<lambda>x. x i) -` A i \<inter> Pi\<^sub>E I \<Omega>))"
-        unfolding b(1)
-        by (auto simp: PiE_def Pi_def)
-      show ?thesis
-        unfolding eq using \<open>A \<in> Pi j E\<close> \<open>j \<in> J\<close> J(2)
-        by (intro F.finite_INT J \<open>j \<in> J\<close> \<open>j \<noteq> {}\<close> sigma_sets.Basic) blast
-    qed
-  qed
-  finally show "?thesis" .
-qed
-
-lemma sets_PiM_in_sets:
-  assumes space: "space N = (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  assumes sets: "\<And>i A. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {x\<in>space N. x i \<in> A} \<in> sets N"
-  shows "sets (\<Pi>\<^sub>M i \<in> I. M i) \<subseteq> sets N"
-  unfolding sets_PiM_single space[symmetric]
-  by (intro sets.sigma_sets_subset subsetI) (auto intro: sets)
-
-lemma sets_PiM_cong[measurable_cong]:
-  assumes "I = J" "\<And>i. i \<in> J \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (PiM I M) = sets (PiM J N)"
-  using assms sets_eq_imp_space_eq[OF assms(2)] by (simp add: sets_PiM_single cong: PiE_cong conj_cong)
-
-lemma sets_PiM_I:
-  assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
-  shows "prod_emb I M J (PIE j:J. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
-proof cases
-  assume "J = {}"
-  then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))"
-    by (auto simp: prod_emb_def)
-  then show ?thesis
-    by (auto simp add: sets_PiM intro!: sigma_sets_top)
-next
-  assume "J \<noteq> {}" with assms show ?thesis
-    by (force simp add: sets_PiM prod_algebra_def)
-qed
-
-lemma measurable_PiM:
-  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
-    f -` prod_emb I M J (Pi\<^sub>E J X) \<inter> space N \<in> sets N"
-  shows "f \<in> measurable N (PiM I M)"
-  using sets_PiM prod_algebra_sets_into_space space
-proof (rule measurable_sigma_sets)
-  fix A assume "A \<in> prod_algebra I M"
-  from prod_algebraE[OF this] guess J X .
-  with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto
-qed
-
-lemma measurable_PiM_Collect:
-  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
-    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N"
-  shows "f \<in> measurable N (PiM I M)"
-  using sets_PiM prod_algebra_sets_into_space space
-proof (rule measurable_sigma_sets)
-  fix A assume "A \<in> prod_algebra I M"
-  from prod_algebraE[OF this] guess J X . note X = this
-  then have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}"
-    using space by (auto simp: prod_emb_def del: PiE_I)
-  also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets)
-  finally show "f -` A \<inter> space N \<in> sets N" .
-qed
-
-lemma measurable_PiM_single:
-  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N"
-  shows "f \<in> measurable N (PiM I M)"
-  using sets_PiM_single
-proof (rule measurable_sigma_sets)
-  fix A assume "A \<in> {{f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
-  then obtain B i where "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
-    by auto
-  with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto
-  also have "\<dots> \<in> sets N" using B by (rule sets)
-  finally show "f -` A \<inter> space N \<in> sets N" .
-qed (auto simp: space)
-
-lemma measurable_PiM_single':
-  assumes f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> measurable N (M i)"
-    and "(\<lambda>\<omega> i. f i \<omega>) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-  shows "(\<lambda>\<omega> i. f i \<omega>) \<in> measurable N (Pi\<^sub>M I M)"
-proof (rule measurable_PiM_single)
-  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
-  then have "{\<omega> \<in> space N. f i \<omega> \<in> A} = f i -` A \<inter> space N"
-    by auto
-  then show "{\<omega> \<in> space N. f i \<omega> \<in> A} \<in> sets N"
-    using A f by (auto intro!: measurable_sets)
-qed fact
-
-lemma sets_PiM_I_finite[measurable]:
-  assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
-  shows "(PIE j:I. E j) \<in> sets (\<Pi>\<^sub>M i\<in>I. M i)"
-  using sets_PiM_I[of I I E M] sets.sets_into_space[OF sets] \<open>finite I\<close> sets by auto
-
-lemma measurable_component_singleton[measurable (raw)]:
-  assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^sub>M I M) (M i)"
-proof (unfold measurable_def, intro CollectI conjI ballI)
-  fix A assume "A \<in> sets (M i)"
-  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) = prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. A)"
-    using sets.sets_into_space \<open>i \<in> I\<close>
-    by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: if_split_asm)
-  then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M I M) \<in> sets (Pi\<^sub>M I M)"
-    using \<open>A \<in> sets (M i)\<close> \<open>i \<in> I\<close> by (auto intro!: sets_PiM_I)
-qed (insert \<open>i \<in> I\<close>, auto simp: space_PiM)
-
-lemma measurable_component_singleton'[measurable_dest]:
-  assumes f: "f \<in> measurable N (Pi\<^sub>M I M)"
-  assumes g: "g \<in> measurable L N"
-  assumes i: "i \<in> I"
-  shows "(\<lambda>x. (f (g x)) i) \<in> measurable L (M i)"
-  using measurable_compose[OF measurable_compose[OF g f] measurable_component_singleton, OF i] .
-
-lemma measurable_PiM_component_rev:
-  "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<lambda>x. f (x i)) \<in> measurable (PiM I M) N"
-  by simp
-
-lemma measurable_case_nat[measurable (raw)]:
-  assumes [measurable (raw)]: "i = 0 \<Longrightarrow> f \<in> measurable M N"
-    "\<And>j. i = Suc j \<Longrightarrow> (\<lambda>x. g x j) \<in> measurable M N"
-  shows "(\<lambda>x. case_nat (f x) (g x) i) \<in> measurable M N"
-  by (cases i) simp_all
-
-lemma measurable_case_nat'[measurable (raw)]:
-  assumes fg[measurable]: "f \<in> measurable N M" "g \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
-  shows "(\<lambda>x. case_nat (f x) (g x)) \<in> measurable N (\<Pi>\<^sub>M i\<in>UNIV. M)"
-  using fg[THEN measurable_space]
-  by (auto intro!: measurable_PiM_single' simp add: space_PiM PiE_iff split: nat.split)
-
-lemma measurable_add_dim[measurable]:
-  "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M)"
-    (is "?f \<in> measurable ?P ?I")
-proof (rule measurable_PiM_single)
-  fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)"
-  have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} =
-    (if j = i then space (Pi\<^sub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
-    using sets.sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM)
-  also have "\<dots> \<in> sets ?P"
-    using A j
-    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
-  finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
-qed (auto simp: space_pair_measure space_PiM PiE_def)
-
-lemma measurable_fun_upd:
-  assumes I: "I = J \<union> {i}"
-  assumes f[measurable]: "f \<in> measurable N (PiM J M)"
-  assumes h[measurable]: "h \<in> measurable N (M i)"
-  shows "(\<lambda>x. (f x) (i := h x)) \<in> measurable N (PiM I M)"
-proof (intro measurable_PiM_single')
-  fix j assume "j \<in> I" then show "(\<lambda>\<omega>. ((f \<omega>)(i := h \<omega>)) j) \<in> measurable N (M j)"
-    unfolding I by (cases "j = i") auto
-next
-  show "(\<lambda>x. (f x)(i := h x)) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
-    using I f[THEN measurable_space] h[THEN measurable_space]
-    by (auto simp: space_PiM PiE_iff extensional_def)
-qed
-
-lemma measurable_component_update:
-  "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)"
-  by simp
-
-lemma measurable_merge[measurable]:
-  "merge I J \<in> measurable (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M)"
-    (is "?f \<in> measurable ?P ?U")
-proof (rule measurable_PiM_single)
-  fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J"
-  then have "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} =
-    (if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)"
-    by (auto simp: merge_def)
-  also have "\<dots> \<in> sets ?P"
-    using A
-    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
-  finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" .
-qed (auto simp: space_pair_measure space_PiM PiE_iff merge_def extensional_def)
-
-lemma measurable_restrict[measurable (raw)]:
-  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
-  shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^sub>M I M)"
-proof (rule measurable_PiM_single)
-  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
-  then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N"
-    by auto
-  then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N"
-    using A X by (auto intro!: measurable_sets)
-qed (insert X, auto simp add: PiE_def dest: measurable_space)
-
-lemma measurable_abs_UNIV:
-  "(\<And>n. (\<lambda>\<omega>. f n \<omega>) \<in> measurable M (N n)) \<Longrightarrow> (\<lambda>\<omega> n. f n \<omega>) \<in> measurable M (PiM UNIV N)"
-  by (intro measurable_PiM_single) (auto dest: measurable_space)
-
-lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
-  by (intro measurable_restrict measurable_component_singleton) auto
-
-lemma measurable_restrict_subset':
-  assumes "J \<subseteq> L" "\<And>x. x \<in> J \<Longrightarrow> sets (M x) = sets (N x)"
-  shows "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
-proof-
-  from assms(1) have "(\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M)"
-    by (rule measurable_restrict_subset)
-  also from assms(2) have "measurable (Pi\<^sub>M L M) (Pi\<^sub>M J M) = measurable (Pi\<^sub>M L M) (Pi\<^sub>M J N)"
-    by (intro sets_PiM_cong measurable_cong_sets) simp_all
-  finally show ?thesis .
-qed
-
-lemma measurable_prod_emb[intro, simp]:
-  "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^sub>M L M)"
-  unfolding prod_emb_def space_PiM[symmetric]
-  by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
-
-lemma merge_in_prod_emb:
-  assumes "y \<in> space (PiM I M)" "x \<in> X" and X: "X \<in> sets (Pi\<^sub>M J M)" and "J \<subseteq> I"
-  shows "merge J I (x, y) \<in> prod_emb I M J X"
-  using assms sets.sets_into_space[OF X]
-  by (simp add: merge_def prod_emb_def subset_eq space_PiM PiE_def extensional_restrict Pi_iff
-           cong: if_cong restrict_cong)
-     (simp add: extensional_def)
-
-lemma prod_emb_eq_emptyD:
-  assumes J: "J \<subseteq> I" and ne: "space (PiM I M) \<noteq> {}" and X: "X \<in> sets (Pi\<^sub>M J M)"
-    and *: "prod_emb I M J X = {}"
-  shows "X = {}"
-proof safe
-  fix x assume "x \<in> X"
-  obtain \<omega> where "\<omega> \<in> space (PiM I M)"
-    using ne by blast
-  from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto
-qed
-
-lemma sets_in_Pi_aux:
-  "finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
-  {x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)"
-  by (simp add: subset_eq Pi_iff)
-
-lemma sets_in_Pi[measurable (raw)]:
-  "finite I \<Longrightarrow> f \<in> measurable N (PiM I M) \<Longrightarrow>
-  (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
-  Measurable.pred N (\<lambda>x. f x \<in> Pi I F)"
-  unfolding pred_def
-  by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_Pi_aux]) auto
-
-lemma sets_in_extensional_aux:
-  "{x\<in>space (PiM I M). x \<in> extensional I} \<in> sets (PiM I M)"
-proof -
-  have "{x\<in>space (PiM I M). x \<in> extensional I} = space (PiM I M)"
-    by (auto simp add: extensional_def space_PiM)
-  then show ?thesis by simp
-qed
-
-lemma sets_in_extensional[measurable (raw)]:
-  "f \<in> measurable N (PiM I M) \<Longrightarrow> Measurable.pred N (\<lambda>x. f x \<in> extensional I)"
-  unfolding pred_def
-  by (rule measurable_sets_Collect[of f N "PiM I M", OF _ sets_in_extensional_aux]) auto
-
-lemma sets_PiM_I_countable:
-  assumes I: "countable I" and E: "\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i)" shows "Pi\<^sub>E I E \<in> sets (Pi\<^sub>M I M)"
-proof cases
-  assume "I \<noteq> {}"
-  then have "PiE I E = (\<Inter>i\<in>I. prod_emb I M {i} (PiE {i} E))"
-    using E[THEN sets.sets_into_space] by (auto simp: PiE_iff prod_emb_def fun_eq_iff)
-  also have "\<dots> \<in> sets (PiM I M)"
-    using I \<open>I \<noteq> {}\<close> by (safe intro!: sets.countable_INT' measurable_prod_emb sets_PiM_I_finite E)
-  finally show ?thesis .
-qed (simp add: sets_PiM_empty)
-
-lemma sets_PiM_D_countable:
-  assumes A: "A \<in> PiM I M"
-  shows "\<exists>J\<subseteq>I. \<exists>X\<in>PiM J M. countable J \<and> A = prod_emb I M J X"
-  using A[unfolded sets_PiM_single]
-proof induction
-  case (Basic A)
-  then obtain i X where *: "i \<in> I" "X \<in> sets (M i)" and "A = {f \<in> \<Pi>\<^sub>E i\<in>I. space (M i). f i \<in> X}"
-    by auto
-  then have A: "A = prod_emb I M {i} (\<Pi>\<^sub>E _\<in>{i}. X)"
-    by (auto simp: prod_emb_def)
-  then show ?case
-    by (intro exI[of _ "{i}"] conjI bexI[of _ "\<Pi>\<^sub>E _\<in>{i}. X"])
-       (auto intro: countable_finite * sets_PiM_I_finite)
-next
-  case Empty then show ?case
-    by (intro exI[of _ "{}"] conjI bexI[of _ "{}"]) auto
-next
-  case (Compl A)
-  then obtain J X where "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)" "countable J" "A = prod_emb I M J X"
-    by auto
-  then show ?case
-    by (intro exI[of _ J] bexI[of _ "space (PiM J M) - X"] conjI)
-       (auto simp add: space_PiM prod_emb_PiE intro!: sets_PiM_I_countable)
-next
-  case (Union K)
-  obtain J X where J: "\<And>i. J i \<subseteq> I" "\<And>i. countable (J i)" and X: "\<And>i. X i \<in> sets (Pi\<^sub>M (J i) M)"
-    and K: "\<And>i. K i = prod_emb I M (J i) (X i)"
-    by (metis Union.IH)
-  show ?case
-  proof (intro exI[of _ "\<Union>i. J i"] bexI[of _ "\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i)"] conjI)
-    show "(\<Union>i. J i) \<subseteq> I" "countable (\<Union>i. J i)" using J by auto
-    with J show "UNION UNIV K = prod_emb I M (\<Union>i. J i) (\<Union>i. prod_emb (\<Union>i. J i) M (J i) (X i))"
-      by (simp add: K[abs_def] SUP_upper)
-  qed(auto intro: X)
-qed
-
-lemma measure_eqI_PiM_finite:
-  assumes [simp]: "finite I" "sets P = PiM I M" "sets Q = PiM I M"
-  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = Q (Pi\<^sub>E I A)"
-  assumes A: "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = space (PiM I M)" "\<And>i::nat. P (A i) \<noteq> \<infinity>"
-  shows "P = Q"
-proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
-  show "range A \<subseteq> prod_algebra I M" "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. P (A i) \<noteq> \<infinity>"
-    unfolding space_PiM[symmetric] by fact+
-  fix X assume "X \<in> prod_algebra I M"
-  then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"
-    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
-    by (force elim!: prod_algebraE)
-  then show "emeasure P X = emeasure Q X"
-    unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq)
-qed (simp_all add: sets_PiM)
-
-lemma measure_eqI_PiM_infinite:
-  assumes [simp]: "sets P = PiM I M" "sets Q = PiM I M"
-  assumes eq: "\<And>A J. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow>
-    P (prod_emb I M J (Pi\<^sub>E J A)) = Q (prod_emb I M J (Pi\<^sub>E J A))"
-  assumes A: "finite_measure P"
-  shows "P = Q"
-proof (rule measure_eqI_generator_eq[OF Int_stable_prod_algebra prod_algebra_sets_into_space])
-  interpret finite_measure P by fact
-  define i where "i = (SOME i. i \<in> I)"
-  have i: "I \<noteq> {} \<Longrightarrow> i \<in> I"
-    unfolding i_def by (rule someI_ex) auto
-  define A where "A n =
-    (if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)))"
-    for n :: nat
-  then show "range A \<subseteq> prod_algebra I M"
-    using prod_algebraI[of "{}" I "\<lambda>i. space (M i)" M] by (auto intro!: prod_algebraI i)
-  have "\<And>i. A i = space (PiM I M)"
-    by (auto simp: prod_emb_def space_PiM PiE_iff A_def i ex_in_conv[symmetric] exI)
-  then show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))" "\<And>i. emeasure P (A i) \<noteq> \<infinity>"
-    by (auto simp: space_PiM)
-next
-  fix X assume X: "X \<in> prod_algebra I M"
-  then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"
-    and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
-    by (force elim!: prod_algebraE)
-  then show "emeasure P X = emeasure Q X"
-    by (auto intro!: eq)
-qed (auto simp: sets_PiM)
-
-locale product_sigma_finite =
-  fixes M :: "'i \<Rightarrow> 'a measure"
-  assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
-
-sublocale product_sigma_finite \<subseteq> M?: sigma_finite_measure "M i" for i
-  by (rule sigma_finite_measures)
-
-locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
-  fixes I :: "'i set"
-  assumes finite_index: "finite I"
-
-lemma (in finite_product_sigma_finite) sigma_finite_pairs:
-  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
-    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
-    (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k) \<and>
-    (\<Union>k. \<Pi>\<^sub>E i\<in>I. F i k) = space (PiM I M)"
-proof -
-  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)"
-    using M.sigma_finite_incseq by metis
-  from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>"
-    by auto
-  let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I. F i k"
-  note space_PiM[simp]
-  show ?thesis
-  proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
-    fix i show "range (F i) \<subseteq> sets (M i)" by fact
-  next
-    fix i k show "emeasure (M i) (F i k) \<noteq> \<infinity>" by fact
-  next
-    fix x assume "x \<in> (\<Union>i. ?F i)" with F(1) show "x \<in> space (PiM I M)"
-      by (auto simp: PiE_def dest!: sets.sets_into_space)
-  next
-    fix f assume "f \<in> space (PiM I M)"
-    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
-    show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def PiE_def)
-  next
-    fix i show "?F i \<subseteq> ?F (Suc i)"
-      using \<open>\<And>i. incseq (F i)\<close>[THEN incseq_SucD] by auto
-  qed
-qed
-
-lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
-proof -
-  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ennreal)"
-  have "emeasure (Pi\<^sub>M {} M) (prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = 1"
-  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
-    show "positive (PiM {} M) ?\<mu>"
-      by (auto simp: positive_def)
-    show "countably_additive (PiM {} M) ?\<mu>"
-      by (rule sets.countably_additiveI_finite)
-         (auto simp: additive_def positive_def sets_PiM_empty space_PiM_empty intro!: )
-  qed (auto simp: prod_emb_def)
-  also have "(prod_emb {} M {} (\<Pi>\<^sub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
-    by (auto simp: prod_emb_def)
-  finally show ?thesis
-    by simp
-qed
-
-lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
-  by (rule measure_eqI) (auto simp add: sets_PiM_empty)
-
-lemma (in product_sigma_finite) emeasure_PiM:
-  "finite I \<Longrightarrow> (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
-proof (induct I arbitrary: A rule: finite_induct)
-  case (insert i I)
-  interpret finite_product_sigma_finite M I by standard fact
-  have "finite (insert i I)" using \<open>finite I\<close> by auto
-  interpret I': finite_product_sigma_finite M "insert i I" by standard fact
-  let ?h = "(\<lambda>(f, y). f(i := y))"
-
-  let ?P = "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) (Pi\<^sub>M (insert i I) M) ?h"
-  let ?\<mu> = "emeasure ?P"
-  let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}"
-  let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))"
-
-  have "emeasure (Pi\<^sub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^sub>E (insert i I) A)) =
-    (\<Prod>i\<in>insert i I. emeasure (M i) (A i))"
-  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
-    fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))"
-    then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto
-    let ?p = "prod_emb (insert i I) M J (Pi\<^sub>E J E)"
-    let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^sub>E j\<in>J-{i}. E j)"
-    have "?\<mu> ?p =
-      emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i))"
-      by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
-    also have "?h -` ?p \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
-      using J E[rule_format, THEN sets.sets_into_space]
-      by (force simp: space_pair_measure space_PiM prod_emb_iff PiE_def Pi_iff split: if_split_asm)
-    also have "emeasure (Pi\<^sub>M I M \<Otimes>\<^sub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
-      emeasure (Pi\<^sub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
-      using J E by (intro M.emeasure_pair_measure_Times sets_PiM_I) auto
-    also have "?p' = (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
-      using J E[rule_format, THEN sets.sets_into_space]
-      by (auto simp: prod_emb_iff PiE_def Pi_iff split: if_split_asm) blast+
-    also have "emeasure (Pi\<^sub>M I M) (\<Pi>\<^sub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
-      (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
-      using E by (subst insert) (auto intro!: setprod.cong)
-    also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) *
-       emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)"
-      using insert by (auto simp: mult.commute intro!: arg_cong2[where f="op *"] setprod.cong)
-    also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)"
-      using insert(1,2) J E by (intro setprod.mono_neutral_right) auto
-    finally show "?\<mu> ?p = \<dots>" .
-
-    show "prod_emb (insert i I) M J (Pi\<^sub>E J E) \<in> Pow (\<Pi>\<^sub>E i\<in>insert i I. space (M i))"
-      using J E[rule_format, THEN sets.sets_into_space] by (auto simp: prod_emb_iff PiE_def)
-  next
-    show "positive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^sub>M (insert i I) M)) ?\<mu>"
-      using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all
-  next
-    show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
-      insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))"
-      using insert by auto
-  qed (auto intro!: setprod.cong)
-  with insert show ?case
-    by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space)
-qed simp
-
-lemma (in product_sigma_finite) PiM_eqI:
-  assumes I[simp]: "finite I" and P: "sets P = PiM I M"
-  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
-  shows "P = PiM I M"
-proof -
-  interpret finite_product_sigma_finite M I
-    proof qed fact
-  from sigma_finite_pairs guess C .. note C = this
-  show ?thesis
-  proof (rule measure_eqI_PiM_finite[OF I refl P, symmetric])
-    show "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> (Pi\<^sub>M I M) (Pi\<^sub>E I A) = P (Pi\<^sub>E I A)" for A
-      by (simp add: eq emeasure_PiM)
-    define A where "A n = (\<Pi>\<^sub>E i\<in>I. C i n)" for n
-    with C show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>" "(\<Union>i. A i) = space (PiM I M)"
-      by (auto intro!: prod_algebraI_finite simp: emeasure_PiM subset_eq ennreal_setprod_eq_top)
-  qed
-qed
-
-lemma (in product_sigma_finite) sigma_finite:
-  assumes "finite I"
-  shows "sigma_finite_measure (PiM I M)"
-proof
-  interpret finite_product_sigma_finite M I by standard fact
-
-  obtain F where F: "\<And>j. countable (F j)" "\<And>j f. f \<in> F j \<Longrightarrow> f \<in> sets (M j)"
-    "\<And>j f. f \<in> F j \<Longrightarrow> emeasure (M j) f \<noteq> \<infinity>" and
-    in_space: "\<And>j. space (M j) = (\<Union>F j)"
-    using sigma_finite_countable by (metis subset_eq)
-  moreover have "(\<Union>(PiE I ` PiE I F)) = space (Pi\<^sub>M I M)"
-    using in_space by (auto simp: space_PiM PiE_iff intro!: PiE_choice[THEN iffD2])
-  ultimately show "\<exists>A. countable A \<and> A \<subseteq> sets (Pi\<^sub>M I M) \<and> \<Union>A = space (Pi\<^sub>M I M) \<and> (\<forall>a\<in>A. emeasure (Pi\<^sub>M I M) a \<noteq> \<infinity>)"
-    by (intro exI[of _ "PiE I ` PiE I F"])
-       (auto intro!: countable_PiE sets_PiM_I_finite
-             simp: PiE_iff emeasure_PiM finite_index ennreal_setprod_eq_top)
-qed
-
-sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^sub>M I M"
-  using sigma_finite[OF finite_index] .
-
-lemma (in finite_product_sigma_finite) measure_times:
-  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
-  using emeasure_PiM[OF finite_index] by auto
-
-lemma (in product_sigma_finite) nn_integral_empty:
-  "0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
-  by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2)
-
-lemma (in product_sigma_finite) distr_merge:
-  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
-  shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
-   (is "?D = ?P")
-proof (rule PiM_eqI)
-  interpret I: finite_product_sigma_finite M I by standard fact
-  interpret J: finite_product_sigma_finite M J by standard fact
-  fix A assume A: "\<And>i. i \<in> I \<union> J \<Longrightarrow> A i \<in> sets (M i)"
-  have *: "(merge I J -` Pi\<^sub>E (I \<union> J) A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)) = PiE I A \<times> PiE J A"
-    using A[THEN sets.sets_into_space] by (auto simp: space_PiM space_pair_measure)
-  from A fin show "emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) (Pi\<^sub>E (I \<union> J) A) =
-      (\<Prod>i\<in>I \<union> J. emeasure (M i) (A i))"
-    by (subst emeasure_distr)
-       (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times setprod.union_disjoint)
-qed (insert fin, simp_all)
-
-lemma (in product_sigma_finite) product_nn_integral_fold:
-  assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
-  and f[measurable]: "f \<in> borel_measurable (Pi\<^sub>M (I \<union> J) M)"
-  shows "integral\<^sup>N (Pi\<^sub>M (I \<union> J) M) f =
-    (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ 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
-  interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard
-  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] by (simp add: comp_def)
-  show ?thesis
-    apply (subst distr_merge[OF IJ, symmetric])
-    apply (subst nn_integral_distr[OF measurable_merge])
-    apply measurable []
-    apply (subst J.nn_integral_fst[symmetric, OF P_borel])
-    apply simp
-    done
-qed
-
-lemma (in product_sigma_finite) distr_singleton:
-  "distr (Pi\<^sub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
-proof (intro measure_eqI[symmetric])
-  interpret I: finite_product_sigma_finite M "{i}" by standard simp
-  fix A assume A: "A \<in> sets (M i)"
-  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^sub>M {i} M) = (\<Pi>\<^sub>E i\<in>{i}. A)"
-    using sets.sets_into_space by (auto simp: space_PiM)
-  then show "emeasure (M i) A = emeasure ?D A"
-    using A I.measure_times[of "\<lambda>_. A"]
-    by (simp add: emeasure_distr measurable_component_singleton)
-qed simp
-
-lemma (in product_sigma_finite) product_nn_integral_singleton:
-  assumes f: "f \<in> borel_measurable (M i)"
-  shows "integral\<^sup>N (Pi\<^sub>M {i} M) (\<lambda>x. f (x i)) = integral\<^sup>N (M i) f"
-proof -
-  interpret I: finite_product_sigma_finite M "{i}" by standard simp
-  from f show ?thesis
-    apply (subst distr_singleton[symmetric])
-    apply (subst nn_integral_distr[OF measurable_component_singleton])
-    apply simp_all
-    done
-qed
-
-lemma (in product_sigma_finite) product_nn_integral_insert:
-  assumes I[simp]: "finite I" "i \<notin> I"
-    and f: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
-  shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^sub>M I M))"
-proof -
-  interpret I: finite_product_sigma_finite M I by standard auto
-  interpret i: finite_product_sigma_finite M "{i}" by standard auto
-  have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
-    using f by auto
-  show ?thesis
-    unfolding product_nn_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f]
-  proof (rule nn_integral_cong, subst product_nn_integral_singleton[symmetric])
-    fix x assume x: "x \<in> space (Pi\<^sub>M I M)"
-    let ?f = "\<lambda>y. f (x(i := y))"
-    show "?f \<in> borel_measurable (M i)"
-      using measurable_comp[OF measurable_component_update f, OF x \<open>i \<notin> I\<close>]
-      unfolding comp_def .
-    show "(\<integral>\<^sup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^sub>M {i} M) = (\<integral>\<^sup>+ y. f (x(i := y i)) \<partial>Pi\<^sub>M {i} M)"
-      using x
-      by (auto intro!: nn_integral_cong arg_cong[where f=f]
-               simp add: space_PiM extensional_def PiE_def)
-  qed
-qed
-
-lemma (in product_sigma_finite) product_nn_integral_insert_rev:
-  assumes I[simp]: "finite I" "i \<notin> I"
-    and [measurable]: "f \<in> borel_measurable (Pi\<^sub>M (insert i I) M)"
-  shows "integral\<^sup>N (Pi\<^sub>M (insert i I) M) f = (\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x(i := y)) \<partial>(Pi\<^sub>M I M)) \<partial>(M i))"
-  apply (subst product_nn_integral_insert[OF assms])
-  apply (rule pair_sigma_finite.Fubini')
-  apply intro_locales []
-  apply (rule sigma_finite[OF I(1)])
-  apply measurable
-  done
-
-lemma (in product_sigma_finite) product_nn_integral_setprod:
-  assumes "finite I" "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
-  shows "(\<integral>\<^sup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^sub>M I M) = (\<Prod>i\<in>I. integral\<^sup>N (M i) (f i))"
-using assms proof (induction I)
-  case (insert i I)
-  note insert.prems[measurable]
-  note \<open>finite I\<close>[intro, simp]
-  interpret I: finite_product_sigma_finite M I by standard auto
-  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 insert by (auto intro!: setprod.cong)
-  have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^sub>M J M)"
-    using sets.sets_into_space insert
-    by (intro borel_measurable_setprod_ennreal
-              measurable_comp[OF measurable_component_singleton, unfolded comp_def])
-       auto
-  then show ?case
-    apply (simp add: product_nn_integral_insert[OF insert(1,2)])
-    apply (simp add: insert(2-) * nn_integral_multc)
-    apply (subst nn_integral_cmult)
-    apply (auto simp add: insert(2-))
-    done
-qed (simp add: space_PiM)
-
-lemma (in product_sigma_finite) product_nn_integral_pair:
-  assumes [measurable]: "case_prod f \<in> borel_measurable (M x \<Otimes>\<^sub>M M y)"
-  assumes xy: "x \<noteq> y"
-  shows "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {x, y} M) = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
-proof-
-  interpret psm: pair_sigma_finite "M x" "M y"
-    unfolding pair_sigma_finite_def using sigma_finite_measures by simp_all
-  have "{x, y} = {y, x}" by auto
-  also have "(\<integral>\<^sup>+\<sigma>. f (\<sigma> x) (\<sigma> y) \<partial>PiM {y, x} M) = (\<integral>\<^sup>+y. \<integral>\<^sup>+\<sigma>. f (\<sigma> x) y \<partial>PiM {x} M \<partial>M y)"
-    using xy by (subst product_nn_integral_insert_rev) simp_all
-  also have "... = (\<integral>\<^sup>+y. \<integral>\<^sup>+x. f x y \<partial>M x \<partial>M y)"
-    by (intro nn_integral_cong, subst product_nn_integral_singleton) simp_all
-  also have "... = (\<integral>\<^sup>+z. f (fst z) (snd z) \<partial>(M x \<Otimes>\<^sub>M M y))"
-    by (subst psm.nn_integral_snd[symmetric]) simp_all
-  finally show ?thesis .
-qed
-
-lemma (in product_sigma_finite) distr_component:
-  "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
-proof (intro PiM_eqI)
-  fix A assume A: "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
-  then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i"
-    by (auto dest: sets.sets_into_space)
-  with A show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))"
-    by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict)
-qed simp_all
-
-lemma (in product_sigma_finite)
-  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
-  shows emeasure_fold_integral:
-    "emeasure (Pi\<^sub>M (I \<union> J) M) A = (\<integral>\<^sup>+x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M)) \<partial>Pi\<^sub>M I M)" (is ?I)
-    and emeasure_fold_measurable:
-    "(\<lambda>x. emeasure (Pi\<^sub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^sub>M J M))) \<in> borel_measurable (Pi\<^sub>M I M)" (is ?B)
-proof -
-  interpret I: finite_product_sigma_finite M I by standard fact
-  interpret J: finite_product_sigma_finite M J by standard fact
-  interpret IJ: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" ..
-  have merge: "merge I J -` A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) \<in> sets (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)"
-    by (intro measurable_sets[OF _ A] measurable_merge assms)
-
-  show ?I
-    apply (subst distr_merge[symmetric, OF IJ])
-    apply (subst emeasure_distr[OF measurable_merge A])
-    apply (subst J.emeasure_pair_measure_alt[OF merge])
-    apply (auto intro!: nn_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
-    done
-
-  show ?B
-    using IJ.measurable_emeasure_Pair1[OF merge]
-    by (simp add: vimage_comp comp_def space_pair_measure cong: measurable_cong)
-qed
-
-lemma sets_Collect_single:
-  "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^sub>M I M). x i \<in> A } \<in> sets (Pi\<^sub>M I M)"
-  by simp
-
-lemma pair_measure_eq_distr_PiM:
-  fixes M1 :: "'a measure" and M2 :: "'a measure"
-  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
-  shows "(M1 \<Otimes>\<^sub>M M2) = distr (Pi\<^sub>M UNIV (case_bool M1 M2)) (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. (x True, x False))"
-    (is "?P = ?D")
-proof (rule pair_measure_eqI[OF assms])
-  interpret B: product_sigma_finite "case_bool M1 M2"
-    unfolding product_sigma_finite_def using assms by (auto split: bool.split)
-  let ?B = "Pi\<^sub>M UNIV (case_bool M1 M2)"
-
-  have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)"
-    by auto
-  fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
-  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (case_bool M1 M2 i) (case_bool A B i))"
-    by (simp add: UNIV_bool ac_simps)
-  also have "\<dots> = emeasure ?B (Pi\<^sub>E UNIV (case_bool A B))"
-    using A B by (subst B.emeasure_PiM) (auto split: bool.split)
-  also have "Pi\<^sub>E UNIV (case_bool A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
-    using A[THEN sets.sets_into_space] B[THEN sets.sets_into_space]
-    by (auto simp: PiE_iff all_bool_eq space_PiM split: bool.split)
-  finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)"
-    using A B
-      measurable_component_singleton[of True UNIV "case_bool M1 M2"]
-      measurable_component_singleton[of False UNIV "case_bool M1 M2"]
-    by (subst emeasure_distr) (auto simp: measurable_pair_iff)
-qed simp
-
-end
--- a/src/HOL/Probability/Giry_Monad.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Giry_Monad.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -7,7 +7,7 @@
 *)
 
 theory Giry_Monad
-  imports Probability_Measure Lebesgue_Integral_Substitution "~~/src/HOL/Library/Monad_Syntax"
+  imports Probability_Measure "~~/src/HOL/Library/Monad_Syntax"
 begin
 
 section \<open>Sub-probability spaces\<close>
--- a/src/HOL/Probability/Independent_Family.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Independent_Family.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -6,7 +6,7 @@
 section \<open>Independent families of events, event sets, and random variables\<close>
 
 theory Independent_Family
-  imports Probability_Measure Infinite_Product_Measure
+  imports Infinite_Product_Measure
 begin
 
 definition (in prob_space)
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -5,7 +5,7 @@
 section \<open>Infinite Product Measure\<close>
 
 theory Infinite_Product_Measure
-  imports Probability_Measure Caratheodory Projective_Family
+  imports Probability_Measure Projective_Family
 begin
 
 lemma (in product_prob_space) distr_PiM_restrict_finite:
--- a/src/HOL/Probability/Information.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Information.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -8,7 +8,6 @@
 theory Information
 imports
   Independent_Family
-  "~~/src/HOL/Library/Convex"
 begin
 
 lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
--- a/src/HOL/Probability/Interval_Integral.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1123 +0,0 @@
-(*  Title:      HOL/Probability/Interval_Integral.thy
-    Author:     Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU)
-
-Lebesgue integral over an interval (with endpoints possibly +-\<infinity>)
-*)
-
-theory Interval_Integral
-  imports Set_Integral
-begin
-
-lemma continuous_on_vector_derivative:
-  "(\<And>x. x \<in> S \<Longrightarrow> (f has_vector_derivative f' x) (at x within S)) \<Longrightarrow> continuous_on S f"
-  by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous)
-
-lemma has_vector_derivative_weaken:
-  fixes x D and f g s t
-  assumes f: "(f has_vector_derivative D) (at x within t)"
-    and "x \<in> s" "s \<subseteq> t"
-    and "\<And>x. x \<in> s \<Longrightarrow> f x = g x"
-  shows "(g has_vector_derivative D) (at x within s)"
-proof -
-  have "(f has_vector_derivative D) (at x within s) \<longleftrightarrow> (g has_vector_derivative D) (at x within s)"
-    unfolding has_vector_derivative_def has_derivative_iff_norm
-    using assms by (intro conj_cong Lim_cong_within refl) auto
-  then show ?thesis
-    using has_vector_derivative_within_subset[OF f \<open>s \<subseteq> t\<close>] by simp
-qed
-
-definition "einterval a b = {x. a < ereal x \<and> ereal x < b}"
-
-lemma einterval_eq[simp]:
-  shows einterval_eq_Icc: "einterval (ereal a) (ereal b) = {a <..< b}"
-    and einterval_eq_Ici: "einterval (ereal a) \<infinity> = {a <..}"
-    and einterval_eq_Iic: "einterval (- \<infinity>) (ereal b) = {..< b}"
-    and einterval_eq_UNIV: "einterval (- \<infinity>) \<infinity> = UNIV"
-  by (auto simp: einterval_def)
-
-lemma einterval_same: "einterval a a = {}"
-  by (auto simp add: einterval_def)
-
-lemma einterval_iff: "x \<in> einterval a b \<longleftrightarrow> a < ereal x \<and> ereal x < b"
-  by (simp add: einterval_def)
-
-lemma einterval_nonempty: "a < b \<Longrightarrow> \<exists>c. c \<in> einterval a b"
-  by (cases a b rule: ereal2_cases, auto simp: einterval_def intro!: dense gt_ex lt_ex)
-
-lemma open_einterval[simp]: "open (einterval a b)"
-  by (cases a b rule: ereal2_cases)
-     (auto simp: einterval_def intro!: open_Collect_conj open_Collect_less continuous_intros)
-
-lemma borel_einterval[measurable]: "einterval a b \<in> sets borel"
-  unfolding einterval_def by measurable
-
-(*
-    Approximating a (possibly infinite) interval
-*)
-
-lemma filterlim_sup1: "(LIM x F. f x :> G1) \<Longrightarrow> (LIM x F. f x :> (sup G1 G2))"
- unfolding filterlim_def by (auto intro: le_supI1)
-
-lemma ereal_incseq_approx:
-  fixes a b :: ereal
-  assumes "a < b"
-  obtains X :: "nat \<Rightarrow> real" where
-    "incseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> b"
-proof (cases b)
-  case PInf
-  with \<open>a < b\<close> have "a = -\<infinity> \<or> (\<exists>r. a = ereal r)"
-    by (cases a) auto
-  moreover have "(\<lambda>x. ereal (real (Suc x))) \<longlonglongrightarrow> \<infinity>"
-      apply (subst LIMSEQ_Suc_iff)
-      apply (simp add: Lim_PInfty)
-      using nat_ceiling_le_eq by blast
-  moreover have "\<And>r. (\<lambda>x. ereal (r + real (Suc x))) \<longlonglongrightarrow> \<infinity>"
-    apply (subst LIMSEQ_Suc_iff)
-    apply (subst Lim_PInfty)
-    apply (metis add.commute diff_le_eq nat_ceiling_le_eq ereal_less_eq(3))
-    done
-  ultimately show thesis
-    by (intro that[of "\<lambda>i. real_of_ereal a + Suc i"])
-       (auto simp: incseq_def PInf)
-next
-  case (real b')
-  define d where "d = b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)"
-  with \<open>a < b\<close> have a': "0 < d"
-    by (cases a) (auto simp: real)
-  moreover
-  have "\<And>i r. r < b' \<Longrightarrow> (b' - r) * 1 < (b' - r) * real (Suc (Suc i))"
-    by (intro mult_strict_left_mono) auto
-  with \<open>a < b\<close> a' have "\<And>i. a < ereal (b' - d / real (Suc (Suc i)))"
-    by (cases a) (auto simp: real d_def field_simps)
-  moreover have "(\<lambda>i. b' - d / Suc (Suc i)) \<longlonglongrightarrow> b'"
-    apply (subst filterlim_sequentially_Suc)
-    apply (subst filterlim_sequentially_Suc)
-    apply (rule tendsto_eq_intros)
-    apply (auto intro!: tendsto_divide_0[OF tendsto_const] filterlim_sup1
-                simp: at_infinity_eq_at_top_bot filterlim_real_sequentially)
-    done
-  ultimately show thesis
-    by (intro that[of "\<lambda>i. b' - d / Suc (Suc i)"])
-       (auto simp add: real incseq_def intro!: divide_left_mono)
-qed (insert \<open>a < b\<close>, auto)
-
-lemma ereal_decseq_approx:
-  fixes a b :: ereal
-  assumes "a < b"
-  obtains X :: "nat \<Rightarrow> real" where
-    "decseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> a"
-proof -
-  have "-b < -a" using \<open>a < b\<close> by simp
-  from ereal_incseq_approx[OF this] guess X .
-  then show thesis
-    apply (intro that[of "\<lambda>i. - X i"])
-    apply (auto simp add: uminus_ereal.simps[symmetric] decseq_def incseq_def
-                simp del: uminus_ereal.simps)
-    apply (metis ereal_minus_less_minus ereal_uminus_uminus ereal_Lim_uminus)+
-    done
-qed
-
-lemma einterval_Icc_approximation:
-  fixes a b :: ereal
-  assumes "a < b"
-  obtains u l :: "nat \<Rightarrow> real" where
-    "einterval a b = (\<Union>i. {l i .. u i})"
-    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
-    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
-proof -
-  from dense[OF \<open>a < b\<close>] obtain c where "a < c" "c < b" by safe
-  from ereal_incseq_approx[OF \<open>c < b\<close>] guess u . note u = this
-  from ereal_decseq_approx[OF \<open>a < c\<close>] guess l . note l = this
-  { fix i from less_trans[OF \<open>l i < c\<close> \<open>c < u i\<close>] have "l i < u i" by simp }
-  have "einterval a b = (\<Union>i. {l i .. u i})"
-  proof (auto simp: einterval_iff)
-    fix x assume "a < ereal x" "ereal x < b"
-    have "eventually (\<lambda>i. ereal (l i) < ereal x) sequentially"
-      using l(4) \<open>a < ereal x\<close> by (rule order_tendstoD)
-    moreover
-    have "eventually (\<lambda>i. ereal x < ereal (u i)) sequentially"
-      using u(4) \<open>ereal x< b\<close> by (rule order_tendstoD)
-    ultimately have "eventually (\<lambda>i. l i < x \<and> x < u i) sequentially"
-      by eventually_elim auto
-    then show "\<exists>i. l i \<le> x \<and> x \<le> u i"
-      by (auto intro: less_imp_le simp: eventually_sequentially)
-  next
-    fix x i assume "l i \<le> x" "x \<le> u i"
-    with \<open>a < ereal (l i)\<close> \<open>ereal (u i) < b\<close>
-    show "a < ereal x" "ereal x < b"
-      by (auto simp del: ereal_less_eq(3) simp add: ereal_less_eq(3)[symmetric])
-  qed
-  show thesis
-    by (intro that) fact+
-qed
-
-(* TODO: in this definition, it would be more natural if einterval a b included a and b when
-   they are real. *)
-definition interval_lebesgue_integral :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> 'a::{banach, second_countable_topology}" where
-  "interval_lebesgue_integral M a b f =
-    (if a \<le> b then (LINT x:einterval a b|M. f x) else - (LINT x:einterval b a|M. f x))"
-
-syntax
-  "_ascii_interval_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real measure \<Rightarrow> real \<Rightarrow> real"
-  ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60)
-
-translations
-  "LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\<lambda>x. f)"
-
-definition interval_lebesgue_integrable :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a::{banach, second_countable_topology}) \<Rightarrow> bool" where
-  "interval_lebesgue_integrable M a b f =
-    (if a \<le> b then set_integrable M (einterval a b) f else set_integrable M (einterval b a) f)"
-
-syntax
-  "_ascii_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real"
-  ("(4LBINT _=_.._. _)" [0,60,60,61] 60)
-
-translations
-  "LBINT x=a..b. f" == "CONST interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
-
-(*
-    Basic properties of integration over an interval.
-*)
-
-lemma interval_lebesgue_integral_cong:
-  "a \<le> b \<Longrightarrow> (\<And>x. x \<in> einterval a b \<Longrightarrow> f x = g x) \<Longrightarrow> einterval a b \<in> sets M \<Longrightarrow>
-    interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
-  by (auto intro: set_lebesgue_integral_cong simp: interval_lebesgue_integral_def)
-
-lemma interval_lebesgue_integral_cong_AE:
-  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
-    a \<le> b \<Longrightarrow> AE x \<in> einterval a b in M. f x = g x \<Longrightarrow> einterval a b \<in> sets M \<Longrightarrow>
-    interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
-  by (auto intro: set_lebesgue_integral_cong_AE simp: interval_lebesgue_integral_def)
-
-lemma interval_integrable_mirror:
-  shows "interval_lebesgue_integrable lborel a b (\<lambda>x. f (-x)) \<longleftrightarrow>
-    interval_lebesgue_integrable lborel (-b) (-a) f"
-proof -
-  have *: "indicator (einterval a b) (- x) = (indicator (einterval (-b) (-a)) x :: real)"
-    for a b :: ereal and x :: real
-    by (cases a b rule: ereal2_cases) (auto simp: einterval_def split: split_indicator)
-  show ?thesis
-    unfolding interval_lebesgue_integrable_def
-    using lborel_integrable_real_affine_iff[symmetric, of "-1" "\<lambda>x. indicator (einterval _ _) x *\<^sub>R f x" 0]
-    by (simp add: *)
-qed
-
-lemma interval_lebesgue_integral_add [intro, simp]:
-  fixes M a b f
-  assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g"
-  shows "interval_lebesgue_integrable M a b (\<lambda>x. f x + g x)" and
-    "interval_lebesgue_integral M a b (\<lambda>x. f x + g x) =
-   interval_lebesgue_integral M a b f + interval_lebesgue_integral M a b g"
-using assms by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def
-    field_simps)
-
-lemma interval_lebesgue_integral_diff [intro, simp]:
-  fixes M a b f
-  assumes "interval_lebesgue_integrable M a b f"
-    "interval_lebesgue_integrable M a b g"
-  shows "interval_lebesgue_integrable M a b (\<lambda>x. f x - g x)" and
-    "interval_lebesgue_integral M a b (\<lambda>x. f x - g x) =
-   interval_lebesgue_integral M a b f - interval_lebesgue_integral M a b g"
-using assms by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def
-    field_simps)
-
-lemma interval_lebesgue_integrable_mult_right [intro, simp]:
-  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
-  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
-    interval_lebesgue_integrable M a b (\<lambda>x. c * f x)"
-  by (simp add: interval_lebesgue_integrable_def)
-
-lemma interval_lebesgue_integrable_mult_left [intro, simp]:
-  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
-  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
-    interval_lebesgue_integrable M a b (\<lambda>x. f x * c)"
-  by (simp add: interval_lebesgue_integrable_def)
-
-lemma interval_lebesgue_integrable_divide [intro, simp]:
-  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
-  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
-    interval_lebesgue_integrable M a b (\<lambda>x. f x / c)"
-  by (simp add: interval_lebesgue_integrable_def)
-
-lemma interval_lebesgue_integral_mult_right [simp]:
-  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
-  shows "interval_lebesgue_integral M a b (\<lambda>x. c * f x) =
-    c * interval_lebesgue_integral M a b f"
-  by (simp add: interval_lebesgue_integral_def)
-
-lemma interval_lebesgue_integral_mult_left [simp]:
-  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
-  shows "interval_lebesgue_integral M a b (\<lambda>x. f x * c) =
-    interval_lebesgue_integral M a b f * c"
-  by (simp add: interval_lebesgue_integral_def)
-
-lemma interval_lebesgue_integral_divide [simp]:
-  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
-  shows "interval_lebesgue_integral M a b (\<lambda>x. f x / c) =
-    interval_lebesgue_integral M a b f / c"
-  by (simp add: interval_lebesgue_integral_def)
-
-lemma interval_lebesgue_integral_uminus:
-  "interval_lebesgue_integral M a b (\<lambda>x. - f x) = - interval_lebesgue_integral M a b f"
-  by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def)
-
-lemma interval_lebesgue_integral_of_real:
-  "interval_lebesgue_integral M a b (\<lambda>x. complex_of_real (f x)) =
-    of_real (interval_lebesgue_integral M a b f)"
-  unfolding interval_lebesgue_integral_def
-  by (auto simp add: interval_lebesgue_integral_def set_integral_complex_of_real)
-
-lemma interval_lebesgue_integral_le_eq:
-  fixes a b f
-  assumes "a \<le> b"
-  shows "interval_lebesgue_integral M a b f = (LINT x : einterval a b | M. f x)"
-using assms by (auto simp add: interval_lebesgue_integral_def)
-
-lemma interval_lebesgue_integral_gt_eq:
-  fixes a b f
-  assumes "a > b"
-  shows "interval_lebesgue_integral M a b f = -(LINT x : einterval b a | M. f x)"
-using assms by (auto simp add: interval_lebesgue_integral_def less_imp_le einterval_def)
-
-lemma interval_lebesgue_integral_gt_eq':
-  fixes a b f
-  assumes "a > b"
-  shows "interval_lebesgue_integral M a b f = - interval_lebesgue_integral M b a f"
-using assms by (auto simp add: interval_lebesgue_integral_def less_imp_le einterval_def)
-
-lemma interval_integral_endpoints_same [simp]: "(LBINT x=a..a. f x) = 0"
-  by (simp add: interval_lebesgue_integral_def einterval_same)
-
-lemma interval_integral_endpoints_reverse: "(LBINT x=a..b. f x) = -(LBINT x=b..a. f x)"
-  by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integral_def einterval_same)
-
-lemma interval_integrable_endpoints_reverse:
-  "interval_lebesgue_integrable lborel a b f \<longleftrightarrow>
-    interval_lebesgue_integrable lborel b a f"
-  by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integrable_def einterval_same)
-
-lemma interval_integral_reflect:
-  "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))"
-proof (induct a b rule: linorder_wlog)
-  case (sym a b) then show ?case
-    by (auto simp add: interval_lebesgue_integral_def interval_integrable_endpoints_reverse
-             split: if_split_asm)
-next
-  case (le a b) then show ?case
-    unfolding interval_lebesgue_integral_def
-    by (subst set_integral_reflect)
-       (auto simp: interval_lebesgue_integrable_def einterval_iff
-                   ereal_uminus_le_reorder ereal_uminus_less_reorder not_less
-                   uminus_ereal.simps[symmetric]
-             simp del: uminus_ereal.simps
-             intro!: integral_cong
-             split: split_indicator)
-qed
-
-lemma interval_lebesgue_integral_0_infty:
-  "interval_lebesgue_integrable M 0 \<infinity> f \<longleftrightarrow> set_integrable M {0<..} f"
-  "interval_lebesgue_integral M 0 \<infinity> f = (LINT x:{0<..}|M. f x)"
-  unfolding zero_ereal_def
-  by (auto simp: interval_lebesgue_integral_le_eq interval_lebesgue_integrable_def)
-
-lemma interval_integral_to_infinity_eq: "(LINT x=ereal a..\<infinity> | M. f x) = (LINT x : {a<..} | M. f x)"
-  unfolding interval_lebesgue_integral_def by auto
-
-lemma interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) =
-  (set_integrable M {a<..} f)"
-  unfolding interval_lebesgue_integrable_def by auto
-
-(*
-    Basic properties of integration over an interval wrt lebesgue measure.
-*)
-
-lemma interval_integral_zero [simp]:
-  fixes a b :: ereal
-  shows"LBINT x=a..b. 0 = 0"
-unfolding interval_lebesgue_integral_def einterval_eq
-by simp
-
-lemma interval_integral_const [intro, simp]:
-  fixes a b c :: real
-  shows "interval_lebesgue_integrable lborel a b (\<lambda>x. c)" and "LBINT x=a..b. c = c * (b - a)"
-unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq
-by (auto simp add:  less_imp_le field_simps measure_def)
-
-lemma interval_integral_cong_AE:
-  assumes [measurable]: "f \<in> borel_measurable borel" "g \<in> borel_measurable borel"
-  assumes "AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x"
-  shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
-  using assms
-proof (induct a b rule: linorder_wlog)
-  case (sym a b) then show ?case
-    by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
-next
-  case (le a b) then show ?case
-    by (auto simp: interval_lebesgue_integral_def max_def min_def
-             intro!: set_lebesgue_integral_cong_AE)
-qed
-
-lemma interval_integral_cong:
-  assumes "\<And>x. x \<in> einterval (min a b) (max a b) \<Longrightarrow> f x = g x"
-  shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
-  using assms
-proof (induct a b rule: linorder_wlog)
-  case (sym a b) then show ?case
-    by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
-next
-  case (le a b) then show ?case
-    by (auto simp: interval_lebesgue_integral_def max_def min_def
-             intro!: set_lebesgue_integral_cong)
-qed
-
-lemma interval_lebesgue_integrable_cong_AE:
-    "f \<in> borel_measurable lborel \<Longrightarrow> g \<in> borel_measurable lborel \<Longrightarrow>
-    AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x \<Longrightarrow>
-    interval_lebesgue_integrable lborel a b f = interval_lebesgue_integrable lborel a b g"
-  apply (simp add: interval_lebesgue_integrable_def )
-  apply (intro conjI impI set_integrable_cong_AE)
-  apply (auto simp: min_def max_def)
-  done
-
-lemma interval_integrable_abs_iff:
-  fixes f :: "real \<Rightarrow> real"
-  shows  "f \<in> borel_measurable lborel \<Longrightarrow>
-    interval_lebesgue_integrable lborel a b (\<lambda>x. \<bar>f x\<bar>) = interval_lebesgue_integrable lborel a b f"
-  unfolding interval_lebesgue_integrable_def
-  by (subst (1 2) set_integrable_abs_iff') simp_all
-
-lemma interval_integral_Icc:
-  fixes a b :: real
-  shows "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a..b}. f x)"
-  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
-           simp add: interval_lebesgue_integral_def)
-
-lemma interval_integral_Icc':
-  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {x. a \<le> ereal x \<and> ereal x \<le> b}. f x)"
-  by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"]
-           simp add: interval_lebesgue_integral_def einterval_iff)
-
-lemma interval_integral_Ioc:
-  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a<..b}. f x)"
-  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
-           simp add: interval_lebesgue_integral_def einterval_iff)
-
-(* TODO: other versions as well? *) (* Yes: I need the Icc' version. *)
-lemma interval_integral_Ioc':
-  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {x. a < ereal x \<and> ereal x \<le> b}. f x)"
-  by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"]
-           simp add: interval_lebesgue_integral_def einterval_iff)
-
-lemma interval_integral_Ico:
-  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a..<b}. f x)"
-  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
-           simp add: interval_lebesgue_integral_def einterval_iff)
-
-lemma interval_integral_Ioi:
-  "\<bar>a\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..\<infinity>. f x) = (LBINT x : {real_of_ereal a <..}. f x)"
-  by (auto simp add: interval_lebesgue_integral_def einterval_iff)
-
-lemma interval_integral_Ioo:
-  "a \<le> b \<Longrightarrow> \<bar>a\<bar> < \<infinity> ==> \<bar>b\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {real_of_ereal a <..< real_of_ereal b}. f x)"
-  by (auto simp add: interval_lebesgue_integral_def einterval_iff)
-
-lemma interval_integral_discrete_difference:
-  fixes f :: "real \<Rightarrow> 'b::{banach, second_countable_topology}" and a b :: ereal
-  assumes "countable X"
-  and eq: "\<And>x. a \<le> b \<Longrightarrow> a < x \<Longrightarrow> x < b \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
-  and anti_eq: "\<And>x. b \<le> a \<Longrightarrow> b < x \<Longrightarrow> x < a \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
-  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
-  shows "interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
-  unfolding interval_lebesgue_integral_def
-  apply (intro if_cong refl arg_cong[where f="\<lambda>x. - x"] integral_discrete_difference[of X] assms)
-  apply (auto simp: eq anti_eq einterval_iff split: split_indicator)
-  done
-
-lemma interval_integral_sum:
-  fixes a b c :: ereal
-  assumes integrable: "interval_lebesgue_integrable lborel (min a (min b c)) (max a (max b c)) f"
-  shows "(LBINT x=a..b. f x) + (LBINT x=b..c. f x) = (LBINT x=a..c. f x)"
-proof -
-  let ?I = "\<lambda>a b. LBINT x=a..b. f x"
-  { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
-    then have ord: "a \<le> b" "b \<le> c" "a \<le> c" and f': "set_integrable lborel (einterval a c) f"
-      by (auto simp: interval_lebesgue_integrable_def)
-    then have f: "set_borel_measurable borel (einterval a c) f"
-      by (drule_tac borel_measurable_integrable) simp
-    have "(LBINT x:einterval a c. f x) = (LBINT x:einterval a b \<union> einterval b c. f x)"
-    proof (rule set_integral_cong_set)
-      show "AE x in lborel. (x \<in> einterval a b \<union> einterval b c) = (x \<in> einterval a c)"
-        using AE_lborel_singleton[of "real_of_ereal b"] ord
-        by (cases a b c rule: ereal3_cases) (auto simp: einterval_iff)
-    qed (insert ord, auto intro!: set_borel_measurable_subset[OF f] simp: einterval_iff)
-    also have "\<dots> = (LBINT x:einterval a b. f x) + (LBINT x:einterval b c. f x)"
-      using ord
-      by (intro set_integral_Un_AE) (auto intro!: set_integrable_subset[OF f'] simp: einterval_iff not_less)
-    finally have "?I a b + ?I b c = ?I a c"
-      using ord by (simp add: interval_lebesgue_integral_def)
-  } note 1 = this
-  { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
-    from 1[OF this] have "?I b c + ?I a b = ?I a c"
-      by (metis add.commute)
-  } note 2 = this
-  have 3: "\<And>a b. b \<le> a \<Longrightarrow> (LBINT x=a..b. f x) = - (LBINT x=b..a. f x)"
-    by (rule interval_integral_endpoints_reverse)
-  show ?thesis
-    using integrable
-    by (cases a b b c a c rule: linorder_le_cases[case_product linorder_le_cases linorder_cases])
-       (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3)
-qed
-
-lemma interval_integrable_isCont:
-  fixes a b and f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
-  shows "(\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> isCont f x) \<Longrightarrow>
-    interval_lebesgue_integrable lborel a b f"
-proof (induct a b rule: linorder_wlog)
-  case (le a b) then show ?case
-    by (auto simp: interval_lebesgue_integrable_def
-             intro!: set_integrable_subset[OF borel_integrable_compact[of "{a .. b}"]]
-                     continuous_at_imp_continuous_on)
-qed (auto intro: interval_integrable_endpoints_reverse[THEN iffD1])
-
-lemma interval_integrable_continuous_on:
-  fixes a b :: real and f
-  assumes "a \<le> b" and "continuous_on {a..b} f"
-  shows "interval_lebesgue_integrable lborel a b f"
-using assms unfolding interval_lebesgue_integrable_def apply simp
-  by (rule set_integrable_subset, rule borel_integrable_atLeastAtMost' [of a b], auto)
-
-lemma interval_integral_eq_integral:
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
-  shows "a \<le> b \<Longrightarrow> set_integrable lborel {a..b} f \<Longrightarrow> LBINT x=a..b. f x = integral {a..b} f"
-  by (subst interval_integral_Icc, simp) (rule set_borel_integral_eq_integral)
-
-lemma interval_integral_eq_integral':
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
-  shows "a \<le> b \<Longrightarrow> set_integrable lborel (einterval a b) f \<Longrightarrow> LBINT x=a..b. f x = integral (einterval a b) f"
-  by (subst interval_lebesgue_integral_le_eq, simp) (rule set_borel_integral_eq_integral)
-
-(*
-    General limit approximation arguments
-*)
-
-lemma interval_integral_Icc_approx_nonneg:
-  fixes a b :: ereal
-  assumes "a < b"
-  fixes u l :: "nat \<Rightarrow> real"
-  assumes  approx: "einterval a b = (\<Union>i. {l i .. u i})"
-    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
-    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
-  fixes f :: "real \<Rightarrow> real"
-  assumes f_integrable: "\<And>i. set_integrable lborel {l i..u i} f"
-  assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x"
-  assumes f_measurable: "set_borel_measurable lborel (einterval a b) f"
-  assumes lbint_lim: "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> C"
-  shows
-    "set_integrable lborel (einterval a b) f"
-    "(LBINT x=a..b. f x) = C"
-proof -
-  have 1: "\<And>i. set_integrable lborel {l i..u i} f" by (rule f_integrable)
-  have 2: "AE x in lborel. mono (\<lambda>n. indicator {l n..u n} x *\<^sub>R f x)"
-  proof -
-     from f_nonneg have "AE x in lborel. \<forall>i. l i \<le> x \<longrightarrow> x \<le> u i \<longrightarrow> 0 \<le> f x"
-      by eventually_elim
-         (metis approx(5) approx(6) dual_order.strict_trans1 ereal_less_eq(3) le_less_trans)
-    then show ?thesis
-      apply eventually_elim
-      apply (auto simp: mono_def split: split_indicator)
-      apply (metis approx(3) decseqD order_trans)
-      apply (metis approx(2) incseqD order_trans)
-      done
-  qed
-  have 3: "AE x in lborel. (\<lambda>i. indicator {l i..u i} x *\<^sub>R f x) \<longlonglongrightarrow> indicator (einterval a b) x *\<^sub>R f x"
-  proof -
-    { fix x i assume "l i \<le> x" "x \<le> u i"
-      then have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
-        apply (auto simp: eventually_sequentially intro!: exI[of _ i])
-        apply (metis approx(3) decseqD order_trans)
-        apply (metis approx(2) incseqD order_trans)
-        done
-      then have "eventually (\<lambda>i. f x * indicator {l i..u i} x = f x) sequentially"
-        by eventually_elim auto }
-    then show ?thesis
-      unfolding approx(1) by (auto intro!: AE_I2 Lim_eventually split: split_indicator)
-  qed
-  have 4: "(\<lambda>i. \<integral> x. indicator {l i..u i} x *\<^sub>R f x \<partial>lborel) \<longlonglongrightarrow> C"
-    using lbint_lim by (simp add: interval_integral_Icc approx less_imp_le)
-  have 5: "set_borel_measurable lborel (einterval a b) f" by (rule assms)
-  have "(LBINT x=a..b. f x) = lebesgue_integral lborel (\<lambda>x. indicator (einterval a b) x *\<^sub>R f x)"
-    using assms by (simp add: interval_lebesgue_integral_def less_imp_le)
-  also have "... = C" by (rule integral_monotone_convergence [OF 1 2 3 4 5])
-  finally show "(LBINT x=a..b. f x) = C" .
-
-  show "set_integrable lborel (einterval a b) f"
-    by (rule integrable_monotone_convergence[OF 1 2 3 4 5])
-qed
-
-lemma interval_integral_Icc_approx_integrable:
-  fixes u l :: "nat \<Rightarrow> real" and a b :: ereal
-  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
-  assumes "a < b"
-  assumes  approx: "einterval a b = (\<Union>i. {l i .. u i})"
-    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
-    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
-  assumes f_integrable: "set_integrable lborel (einterval a b) f"
-  shows "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
-proof -
-  have "(\<lambda>i. LBINT x:{l i.. u i}. f x) \<longlonglongrightarrow> (LBINT x:einterval a b. f x)"
-  proof (rule integral_dominated_convergence)
-    show "integrable lborel (\<lambda>x. norm (indicator (einterval a b) x *\<^sub>R f x))"
-      by (rule integrable_norm) fact
-    show "set_borel_measurable lborel (einterval a b) f"
-      using f_integrable by (rule borel_measurable_integrable)
-    then show "\<And>i. set_borel_measurable lborel {l i..u i} f"
-      by (rule set_borel_measurable_subset) (auto simp: approx)
-    show "\<And>i. AE x in lborel. norm (indicator {l i..u i} x *\<^sub>R f x) \<le> norm (indicator (einterval a b) x *\<^sub>R f x)"
-      by (intro AE_I2) (auto simp: approx split: split_indicator)
-    show "AE x in lborel. (\<lambda>i. indicator {l i..u i} x *\<^sub>R f x) \<longlonglongrightarrow> indicator (einterval a b) x *\<^sub>R f x"
-    proof (intro AE_I2 tendsto_intros Lim_eventually)
-      fix x
-      { fix i assume "l i \<le> x" "x \<le> u i"
-        with \<open>incseq u\<close>[THEN incseqD, of i] \<open>decseq l\<close>[THEN decseqD, of i]
-        have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
-          by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) }
-      then show "eventually (\<lambda>xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially"
-        using approx order_tendstoD(2)[OF \<open>l \<longlonglongrightarrow> a\<close>, of x] order_tendstoD(1)[OF \<open>u \<longlonglongrightarrow> b\<close>, of x]
-        by (auto split: split_indicator)
-    qed
-  qed
-  with \<open>a < b\<close> \<open>\<And>i. l i < u i\<close> show ?thesis
-    by (simp add: interval_lebesgue_integral_le_eq[symmetric] interval_integral_Icc less_imp_le)
-qed
-
-(*
-  A slightly stronger version of integral_FTC_atLeastAtMost and related facts,
-  with continuous_on instead of isCont
-
-  TODO: make the older versions corollaries of these (using continuous_at_imp_continuous_on, etc.)
-*)
-
-(*
-TODO: many proofs below require inferences like
-
-  a < ereal x \<Longrightarrow> x < y \<Longrightarrow> a < ereal y
-
-where x and y are real. These should be better automated.
-*)
-
-(*
-    The first Fundamental Theorem of Calculus
-
-    First, for finite intervals, and then two versions for arbitrary intervals.
-*)
-
-lemma interval_integral_FTC_finite:
-  fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: real
-  assumes f: "continuous_on {min a b..max a b} f"
-  assumes F: "\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> (F has_vector_derivative (f x)) (at x within
-    {min a b..max a b})"
-  shows "(LBINT x=a..b. f x) = F b - F a"
-  apply (case_tac "a \<le> b")
-  apply (subst interval_integral_Icc, simp)
-  apply (rule integral_FTC_atLeastAtMost, assumption)
-  apply (metis F max_def min_def)
-  using f apply (simp add: min_absorb1 max_absorb2)
-  apply (subst interval_integral_endpoints_reverse)
-  apply (subst interval_integral_Icc, simp)
-  apply (subst integral_FTC_atLeastAtMost, auto)
-  apply (metis F max_def min_def)
-using f by (simp add: min_absorb2 max_absorb1)
-
-lemma interval_integral_FTC_nonneg:
-  fixes f F :: "real \<Rightarrow> real" and a b :: ereal
-  assumes "a < b"
-  assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV F x :> f x"
-  assumes f: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f x"
-  assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x"
-  assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
-  assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
-  shows
-    "set_integrable lborel (einterval a b) f"
-    "(LBINT x=a..b. f x) = B - A"
-proof -
-  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx = this
-  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
-    by (rule order_less_le_trans, rule approx, force)
-  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
-    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
-  have FTCi: "\<And>i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)"
-    using assms approx apply (intro interval_integral_FTC_finite)
-    apply (auto simp add: less_imp_le min_def max_def
-      has_field_derivative_iff_has_vector_derivative[symmetric])
-    apply (rule continuous_at_imp_continuous_on, auto intro!: f)
-    by (rule DERIV_subset [OF F], auto)
-  have 1: "\<And>i. set_integrable lborel {l i..u i} f"
-  proof -
-    fix i show "set_integrable lborel {l i .. u i} f"
-      using \<open>a < l i\<close> \<open>u i < b\<close>
-      by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI)
-         (auto simp del: ereal_less_eq simp add: ereal_less_eq(3)[symmetric])
-  qed
-  have 2: "set_borel_measurable lborel (einterval a b) f"
-    by (auto simp del: real_scaleR_def intro!: set_borel_measurable_continuous
-             simp: continuous_on_eq_continuous_at einterval_iff f)
-  have 3: "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A"
-    apply (subst FTCi)
-    apply (intro tendsto_intros)
-    using B approx unfolding tendsto_at_iff_sequentially comp_def
-    using tendsto_at_iff_sequentially[where 'a=real]
-    apply (elim allE[of _ "\<lambda>i. ereal (u i)"], auto)
-    using A approx unfolding tendsto_at_iff_sequentially comp_def
-    by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
-  show "(LBINT x=a..b. f x) = B - A"
-    by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
-  show "set_integrable lborel (einterval a b) f"
-    by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
-qed
-
-lemma interval_integral_FTC_integrable:
-  fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: ereal
-  assumes "a < b"
-  assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> (F has_vector_derivative f x) (at x)"
-  assumes f: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f x"
-  assumes f_integrable: "set_integrable lborel (einterval a b) f"
-  assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
-  assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
-  shows "(LBINT x=a..b. f x) = B - A"
-proof -
-  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx = this
-  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
-    by (rule order_less_le_trans, rule approx, force)
-  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
-    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
-  have FTCi: "\<And>i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)"
-    using assms approx
-    by (auto simp add: less_imp_le min_def max_def
-             intro!: f continuous_at_imp_continuous_on interval_integral_FTC_finite
-             intro: has_vector_derivative_at_within)
-  have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A"
-    apply (subst FTCi)
-    apply (intro tendsto_intros)
-    using B approx unfolding tendsto_at_iff_sequentially comp_def
-    apply (elim allE[of _ "\<lambda>i. ereal (u i)"], auto)
-    using A approx unfolding tendsto_at_iff_sequentially comp_def
-    by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
-  moreover have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
-    by (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx f_integrable])
-  ultimately show ?thesis
-    by (elim LIMSEQ_unique)
-qed
-
-(*
-  The second Fundamental Theorem of Calculus and existence of antiderivatives on an
-  einterval.
-*)
-
-lemma interval_integral_FTC2:
-  fixes a b c :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "a \<le> c" "c \<le> b"
-  and contf: "continuous_on {a..b} f"
-  fixes x :: real
-  assumes "a \<le> x" and "x \<le> b"
-  shows "((\<lambda>u. LBINT y=c..u. f y) has_vector_derivative (f x)) (at x within {a..b})"
-proof -
-  let ?F = "(\<lambda>u. LBINT y=a..u. f y)"
-  have intf: "set_integrable lborel {a..b} f"
-    by (rule borel_integrable_atLeastAtMost', rule contf)
-  have "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
-    apply (intro integral_has_vector_derivative)
-    using \<open>a \<le> x\<close> \<open>x \<le> b\<close> by (intro continuous_on_subset [OF contf], auto)
-  then have "((\<lambda>u. integral {a..u} f) has_vector_derivative (f x)) (at x within {a..b})"
-    by simp
-  then have "(?F has_vector_derivative (f x)) (at x within {a..b})"
-    by (rule has_vector_derivative_weaken)
-       (auto intro!: assms interval_integral_eq_integral[symmetric] set_integrable_subset [OF intf])
-  then have "((\<lambda>x. (LBINT y=c..a. f y) + ?F x) has_vector_derivative (f x)) (at x within {a..b})"
-    by (auto intro!: derivative_eq_intros)
-  then show ?thesis
-  proof (rule has_vector_derivative_weaken)
-    fix u assume "u \<in> {a .. b}"
-    then show "(LBINT y=c..a. f y) + (LBINT y=a..u. f y) = (LBINT y=c..u. f y)"
-      using assms
-      apply (intro interval_integral_sum)
-      apply (auto simp add: interval_lebesgue_integrable_def simp del: real_scaleR_def)
-      by (rule set_integrable_subset [OF intf], auto simp add: min_def max_def)
-  qed (insert assms, auto)
-qed
-
-lemma einterval_antiderivative:
-  fixes a b :: ereal and f :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "a < b" and contf: "\<And>x :: real. a < x \<Longrightarrow> x < b \<Longrightarrow> isCont f x"
-  shows "\<exists>F. \<forall>x :: real. a < x \<longrightarrow> x < b \<longrightarrow> (F has_vector_derivative f x) (at x)"
-proof -
-  from einterval_nonempty [OF \<open>a < b\<close>] obtain c :: real where [simp]: "a < c" "c < b"
-    by (auto simp add: einterval_def)
-  let ?F = "(\<lambda>u. LBINT y=c..u. f y)"
-  show ?thesis
-  proof (rule exI, clarsimp)
-    fix x :: real
-    assume [simp]: "a < x" "x < b"
-    have 1: "a < min c x" by simp
-    from einterval_nonempty [OF 1] obtain d :: real where [simp]: "a < d" "d < c" "d < x"
-      by (auto simp add: einterval_def)
-    have 2: "max c x < b" by simp
-    from einterval_nonempty [OF 2] obtain e :: real where [simp]: "c < e" "x < e" "e < b"
-      by (auto simp add: einterval_def)
-    show "(?F has_vector_derivative f x) (at x)"
-      (* TODO: factor out the next three lines to has_field_derivative_within_open *)
-      unfolding has_vector_derivative_def
-      apply (subst has_derivative_within_open [of _ "{d<..<e}", symmetric], auto)
-      apply (subst has_vector_derivative_def [symmetric])
-      apply (rule has_vector_derivative_within_subset [of _ _ _ "{d..e}"])
-      apply (rule interval_integral_FTC2, auto simp add: less_imp_le)
-      apply (rule continuous_at_imp_continuous_on)
-      apply (auto intro!: contf)
-      apply (rule order_less_le_trans, rule \<open>a < d\<close>, auto)
-      apply (rule order_le_less_trans) prefer 2
-      by (rule \<open>e < b\<close>, auto)
-  qed
-qed
-
-(*
-    The substitution theorem
-
-    Once again, three versions: first, for finite intervals, and then two versions for
-    arbitrary intervals.
-*)
-
-lemma interval_integral_substitution_finite:
-  fixes a b :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "a \<le> b"
-  and derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_real_derivative (g' x)) (at x within {a..b})"
-  and contf : "continuous_on (g ` {a..b}) f"
-  and contg': "continuous_on {a..b} g'"
-  shows "LBINT x=a..b. g' x *\<^sub>R f (g x) = LBINT y=g a..g b. f y"
-proof-
-  have v_derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_vector_derivative (g' x)) (at x within {a..b})"
-    using derivg unfolding has_field_derivative_iff_has_vector_derivative .
-  then have contg [simp]: "continuous_on {a..b} g"
-    by (rule continuous_on_vector_derivative) auto
-  have 1: "\<And>u. min (g a) (g b) \<le> u \<Longrightarrow> u \<le> max (g a) (g b) \<Longrightarrow>
-      \<exists>x\<in>{a..b}. u = g x"
-    apply (case_tac "g a \<le> g b")
-    apply (auto simp add: min_def max_def less_imp_le)
-    apply (frule (1) IVT' [of g], auto simp add: assms)
-    by (frule (1) IVT2' [of g], auto simp add: assms)
-  from contg \<open>a \<le> b\<close> have "\<exists>c d. g ` {a..b} = {c..d} \<and> c \<le> d"
-    by (elim continuous_image_closed_interval)
-  then obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \<le> d" by auto
-  have "\<exists>F. \<forall>x\<in>{a..b}. (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))"
-    apply (rule exI, auto, subst g_im)
-    apply (rule interval_integral_FTC2 [of c c d])
-    using \<open>c \<le> d\<close> apply auto
-    apply (rule continuous_on_subset [OF contf])
-    using g_im by auto
-  then guess F ..
-  then have derivF: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow>
-    (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))" by auto
-  have contf2: "continuous_on {min (g a) (g b)..max (g a) (g b)} f"
-    apply (rule continuous_on_subset [OF contf])
-    apply (auto simp add: image_def)
-    by (erule 1)
-  have contfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
-    by (blast intro: continuous_on_compose2 contf contg)
-  have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)"
-    apply (subst interval_integral_Icc, simp add: assms)
-    apply (rule integral_FTC_atLeastAtMost[of a b "\<lambda>x. F (g x)", OF \<open>a \<le> b\<close>])
-    apply (rule vector_diff_chain_within[OF v_derivg derivF, unfolded comp_def])
-    apply (auto intro!: continuous_on_scaleR contg' contfg)
-    done
-  moreover have "LBINT y=(g a)..(g b). f y = F (g b) - F (g a)"
-    apply (rule interval_integral_FTC_finite)
-    apply (rule contf2)
-    apply (frule (1) 1, auto)
-    apply (rule has_vector_derivative_within_subset [OF derivF])
-    apply (auto simp add: image_def)
-    by (rule 1, auto)
-  ultimately show ?thesis by simp
-qed
-
-(* TODO: is it possible to lift the assumption here that g' is nonnegative? *)
-
-lemma interval_integral_substitution_integrable:
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and a b u v :: ereal
-  assumes "a < b"
-  and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
-  and contf: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f (g x)"
-  and contg': "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont g' x"
-  and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x"
-  and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
-  and B: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
-  and integrable: "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
-  and integrable2: "set_integrable lborel (einterval A B) (\<lambda>x. f x)"
-  shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
-proof -
-  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx [simp] = this
-  note less_imp_le [simp]
-  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
-    by (rule order_less_le_trans, rule approx, force)
-  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
-    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
-  have [simp]: "\<And>i. l i < b"
-    apply (rule order_less_trans) prefer 2
-    by (rule approx, auto, rule approx)
-  have [simp]: "\<And>i. a < u i"
-    by (rule order_less_trans, rule approx, auto, rule approx)
-  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> l j \<le> l i" by (rule decseqD, rule approx)
-  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> u i \<le> u j" by (rule incseqD, rule approx)
-  have g_nondec [simp]: "\<And>x y. a < x \<Longrightarrow> x \<le> y \<Longrightarrow> y < b \<Longrightarrow> g x \<le> g y"
-    apply (erule DERIV_nonneg_imp_nondecreasing, auto)
-    apply (rule exI, rule conjI, rule deriv_g)
-    apply (erule order_less_le_trans, auto)
-    apply (rule order_le_less_trans, subst ereal_less_eq(3), assumption, auto)
-    apply (rule g'_nonneg)
-    apply (rule less_imp_le, erule order_less_le_trans, auto)
-    by (rule less_imp_le, rule le_less_trans, subst ereal_less_eq(3), assumption, auto)
-  have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
-  proof -
-    have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A"
-      using A apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
-      by (drule_tac x = "\<lambda>i. ereal (l i)" in spec, auto)
-    hence A3: "\<And>i. g (l i) \<ge> A"
-      by (intro decseq_le, auto simp add: decseq_def)
-    have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B"
-      using B apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
-      by (drule_tac x = "\<lambda>i. ereal (u i)" in spec, auto)
-    hence B3: "\<And>i. g (u i) \<le> B"
-      by (intro incseq_le, auto simp add: incseq_def)
-    show "A \<le> B"
-      apply (rule order_trans [OF A3 [of 0]])
-      apply (rule order_trans [OF _ B3 [of 0]])
-      by auto
-    { fix x :: real
-      assume "A < x" and "x < B"
-      then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially"
-        apply (intro eventually_conj order_tendstoD)
-        by (rule A2, assumption, rule B2, assumption)
-      hence "\<exists>i. g (l i) < x \<and> x < g (u i)"
-        by (simp add: eventually_sequentially, auto)
-    } note AB = this
-    show "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
-      apply (auto simp add: einterval_def)
-      apply (erule (1) AB)
-      apply (rule order_le_less_trans, rule A3, simp)
-      apply (rule order_less_le_trans) prefer 2
-      by (rule B3, simp)
-  qed
-  (* finally, the main argument *)
-  {
-     fix i
-     have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)"
-        apply (rule interval_integral_substitution_finite, auto)
-        apply (rule DERIV_subset)
-        unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
-        apply (rule deriv_g)
-        apply (auto intro!: continuous_at_imp_continuous_on contf contg')
-        done
-  } note eq1 = this
-  have "(\<lambda>i. LBINT x=l i..u i. g' x *\<^sub>R f (g x)) \<longlonglongrightarrow> (LBINT x=a..b. g' x *\<^sub>R f (g x))"
-    apply (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx])
-    by (rule assms)
-  hence 2: "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x=a..b. g' x *\<^sub>R f (g x))"
-    by (simp add: eq1)
-  have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
-    apply (auto simp add: incseq_def)
-    apply (rule order_le_less_trans)
-    prefer 2 apply (assumption, auto)
-    by (erule order_less_le_trans, rule g_nondec, auto)
-  have "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x = A..B. f x)"
-    apply (subst interval_lebesgue_integral_le_eq, auto simp del: real_scaleR_def)
-    apply (subst interval_lebesgue_integral_le_eq, rule \<open>A \<le> B\<close>)
-    apply (subst un, rule set_integral_cont_up, auto simp del: real_scaleR_def)
-    apply (rule incseq)
-    apply (subst un [symmetric])
-    by (rule integrable2)
-  thus ?thesis by (intro LIMSEQ_unique [OF _ 2])
-qed
-
-(* TODO: the last two proofs are only slightly different. Factor out common part?
-   An alternative: make the second one the main one, and then have another lemma
-   that says that if f is nonnegative and all the other hypotheses hold, then it is integrable. *)
-
-lemma interval_integral_substitution_nonneg:
-  fixes f g g':: "real \<Rightarrow> real" and a b u v :: ereal
-  assumes "a < b"
-  and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
-  and contf: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f (g x)"
-  and contg': "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont g' x"
-  and f_nonneg: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> 0 \<le> f (g x)" (* TODO: make this AE? *)
-  and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x"
-  and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
-  and B: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
-  and integrable_fg: "set_integrable lborel (einterval a b) (\<lambda>x. f (g x) * g' x)"
-  shows
-    "set_integrable lborel (einterval A B) f"
-    "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
-proof -
-  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx [simp] = this
-  note less_imp_le [simp]
-  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
-    by (rule order_less_le_trans, rule approx, force)
-  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
-    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
-  have [simp]: "\<And>i. l i < b"
-    apply (rule order_less_trans) prefer 2
-    by (rule approx, auto, rule approx)
-  have [simp]: "\<And>i. a < u i"
-    by (rule order_less_trans, rule approx, auto, rule approx)
-  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> l j \<le> l i" by (rule decseqD, rule approx)
-  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> u i \<le> u j" by (rule incseqD, rule approx)
-  have g_nondec [simp]: "\<And>x y. a < x \<Longrightarrow> x \<le> y \<Longrightarrow> y < b \<Longrightarrow> g x \<le> g y"
-    apply (erule DERIV_nonneg_imp_nondecreasing, auto)
-    apply (rule exI, rule conjI, rule deriv_g)
-    apply (erule order_less_le_trans, auto)
-    apply (rule order_le_less_trans, subst ereal_less_eq(3), assumption, auto)
-    apply (rule g'_nonneg)
-    apply (rule less_imp_le, erule order_less_le_trans, auto)
-    by (rule less_imp_le, rule le_less_trans, subst ereal_less_eq(3), assumption, auto)
-  have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
-  proof -
-    have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A"
-      using A apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
-      by (drule_tac x = "\<lambda>i. ereal (l i)" in spec, auto)
-    hence A3: "\<And>i. g (l i) \<ge> A"
-      by (intro decseq_le, auto simp add: decseq_def)
-    have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B"
-      using B apply (auto simp add: einterval_def tendsto_at_iff_sequentially comp_def)
-      by (drule_tac x = "\<lambda>i. ereal (u i)" in spec, auto)
-    hence B3: "\<And>i. g (u i) \<le> B"
-      by (intro incseq_le, auto simp add: incseq_def)
-    show "A \<le> B"
-      apply (rule order_trans [OF A3 [of 0]])
-      apply (rule order_trans [OF _ B3 [of 0]])
-      by auto
-    { fix x :: real
-      assume "A < x" and "x < B"
-      then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially"
-        apply (intro eventually_conj order_tendstoD)
-        by (rule A2, assumption, rule B2, assumption)
-      hence "\<exists>i. g (l i) < x \<and> x < g (u i)"
-        by (simp add: eventually_sequentially, auto)
-    } note AB = this
-    show "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
-      apply (auto simp add: einterval_def)
-      apply (erule (1) AB)
-      apply (rule order_le_less_trans, rule A3, simp)
-      apply (rule order_less_le_trans) prefer 2
-      by (rule B3, simp)
-  qed
-  (* finally, the main argument *)
-  {
-     fix i
-     have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)"
-        apply (rule interval_integral_substitution_finite, auto)
-        apply (rule DERIV_subset, rule deriv_g, auto)
-        apply (rule continuous_at_imp_continuous_on, auto, rule contf, auto)
-        by (rule continuous_at_imp_continuous_on, auto, rule contg', auto)
-     then have "(LBINT x=l i.. u i. (f (g x) * g' x)) = (LBINT y=g (l i)..g (u i). f y)"
-       by (simp add: ac_simps)
-  } note eq1 = this
-  have "(\<lambda>i. LBINT x=l i..u i. f (g x) * g' x)
-      \<longlonglongrightarrow> (LBINT x=a..b. f (g x) * g' x)"
-    apply (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx])
-    by (rule assms)
-  hence 2: "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x=a..b. f (g x) * g' x)"
-    by (simp add: eq1)
-  have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
-    apply (auto simp add: incseq_def)
-    apply (rule order_le_less_trans)
-    prefer 2 apply assumption
-    apply (rule g_nondec, auto)
-    by (erule order_less_le_trans, rule g_nondec, auto)
-  have img: "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> \<exists>c \<ge> l i. c \<le> u i \<and> x = g c"
-    apply (frule (1) IVT' [of g], auto)
-    apply (rule continuous_at_imp_continuous_on, auto)
-    by (rule DERIV_isCont, rule deriv_g, auto)
-  have nonneg_f2: "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> 0 \<le> f x"
-    by (frule (1) img, auto, rule f_nonneg, auto)
-  have contf_2: "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> isCont f x"
-    by (frule (1) img, auto, rule contf, auto)
-  have integrable: "set_integrable lborel (\<Union>i. {g (l i)<..<g (u i)}) f"
-    apply (rule pos_integrable_to_top, auto simp del: real_scaleR_def)
-    apply (rule incseq)
-    apply (rule nonneg_f2, erule less_imp_le, erule less_imp_le)
-    apply (rule set_integrable_subset)
-    apply (rule borel_integrable_atLeastAtMost')
-    apply (rule continuous_at_imp_continuous_on)
-    apply (clarsimp, erule (1) contf_2, auto)
-    apply (erule less_imp_le)+
-    using 2 unfolding interval_lebesgue_integral_def
-    by auto
-  thus "set_integrable lborel (einterval A B) f"
-    by (simp add: un)
-
-  have "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
-  proof (rule interval_integral_substitution_integrable)
-    show "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
-      using integrable_fg by (simp add: ac_simps)
-  qed fact+
-  then show "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
-    by (simp add: ac_simps)
-qed
-
-
-syntax
-"_complex_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> complex"
-("(2CLBINT _. _)" [0,60] 60)
-
-translations
-"CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\<lambda>x. f)"
-
-syntax
-"_complex_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> complex"
-("(3CLBINT _:_. _)" [0,60,61] 60)
-
-translations
-"CLBINT x:A. f" == "CONST complex_set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
-
-abbreviation complex_interval_lebesgue_integral ::
-    "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> complex" where
-  "complex_interval_lebesgue_integral M a b f \<equiv> interval_lebesgue_integral M a b f"
-
-abbreviation complex_interval_lebesgue_integrable ::
-  "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool" where
-  "complex_interval_lebesgue_integrable M a b f \<equiv> interval_lebesgue_integrable M a b f"
-
-syntax
-  "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> real \<Rightarrow> complex"
-  ("(4CLBINT _=_.._. _)" [0,60,60,61] 60)
-
-translations
-  "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
-
-lemma interval_integral_norm:
-  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  shows "interval_lebesgue_integrable lborel a b f \<Longrightarrow> a \<le> b \<Longrightarrow>
-    norm (LBINT t=a..b. f t) \<le> LBINT t=a..b. norm (f t)"
-  using integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
-  by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def)
-
-lemma interval_integral_norm2:
-  "interval_lebesgue_integrable lborel a b f \<Longrightarrow>
-    norm (LBINT t=a..b. f t) \<le> \<bar>LBINT t=a..b. norm (f t)\<bar>"
-proof (induct a b rule: linorder_wlog)
-  case (sym a b) then show ?case
-    by (simp add: interval_integral_endpoints_reverse[of a b] interval_integrable_endpoints_reverse[of a b])
-next
-  case (le a b)
-  then have "\<bar>LBINT t=a..b. norm (f t)\<bar> = LBINT t=a..b. norm (f t)"
-    using integrable_norm[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
-    by (auto simp add: interval_lebesgue_integral_def interval_lebesgue_integrable_def
-             intro!: integral_nonneg_AE abs_of_nonneg)
-  then show ?case
-    using le by (simp add: interval_integral_norm)
-qed
-
-(* TODO: should we have a library of facts like these? *)
-lemma integral_cos: "t \<noteq> 0 \<Longrightarrow> LBINT x=a..b. cos (t * x) = sin (t * b) / t - sin (t * a) / t"
-  apply (intro interval_integral_FTC_finite continuous_intros)
-  by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative[symmetric])
-
-
-end
--- a/src/HOL/Probability/Lebesgue_Integral_Substitution.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,400 +0,0 @@
-(*  Title:      HOL/Probability/Lebesgue_Integral_Substitution.thy
-    Author:     Manuel Eberl
-
-    Provides lemmas for integration by substitution for the basic integral types.
-    Note that the substitution function must have a nonnegative derivative.
-    This could probably be weakened somehow.
-*)
-
-section \<open>Integration by Substition\<close>
-
-theory Lebesgue_Integral_Substitution
-imports Interval_Integral
-begin
-
-lemma nn_integral_substitution_aux:
-  fixes f :: "real \<Rightarrow> ennreal"
-  assumes Mf: "f \<in> borel_measurable borel"
-  assumes nonnegf: "\<And>x. f x \<ge> 0"
-  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'"
-  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
-  assumes "a < b"
-  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
-             (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
-proof-
-  from \<open>a < b\<close> have [simp]: "a \<le> b" by simp
-  from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
-  from this and contg' have Mg: "set_borel_measurable borel {a..b} g" and
-                             Mg': "set_borel_measurable borel {a..b} g'"
-      by (simp_all only: set_measurable_continuous_on_ivl)
-  from derivg have derivg': "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
-    by (simp only: has_field_derivative_iff_has_vector_derivative)
-
-  have real_ind[simp]: "\<And>A x. enn2real (indicator A x) = indicator A x"
-      by (auto split: split_indicator)
-  have ennreal_ind[simp]: "\<And>A x. ennreal (indicator A x) = indicator A x"
-      by (auto split: split_indicator)
-  have [simp]: "\<And>x A. indicator A (g x) = indicator (g -` A) x"
-      by (auto split: split_indicator)
-
-  from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
-    by (rule deriv_nonneg_imp_mono) simp_all
-  with monog have [simp]: "g a \<le> g b" by (auto intro: mono_onD)
-
-  show ?thesis
-  proof (induction rule: borel_measurable_induct[OF Mf, case_names cong set mult add sup])
-    case (cong f1 f2)
-    from cong.hyps(3) have "f1 = f2" by auto
-    with cong show ?case by simp
-  next
-    case (set A)
-    from set.hyps show ?case
-    proof (induction rule: borel_set_induct)
-      case empty
-      thus ?case by simp
-    next
-      case (interval c d)
-      {
-        fix u v :: real assume asm: "{u..v} \<subseteq> {g a..g b}" "u \<le> v"
-
-        obtain u' v' where u'v': "{a..b} \<inter> g-`{u..v} = {u'..v'}" "u' \<le> v'" "g u' = u" "g v' = v"
-             using asm by (rule_tac continuous_interval_vimage_Int[OF contg monog, of u v]) simp_all
-        hence "{u'..v'} \<subseteq> {a..b}" "{u'..v'} \<subseteq> g -` {u..v}" by blast+
-        with u'v'(2) have "u' \<in> g -` {u..v}" "v' \<in> g -` {u..v}" by auto
-        from u'v'(1) have [simp]: "{a..b} \<inter> g -` {u..v} \<in> sets borel" by simp
-
-        have A: "continuous_on {min u' v'..max u' v'} g'"
-            by (simp only: u'v' max_absorb2 min_absorb1)
-               (intro continuous_on_subset[OF contg'], insert u'v', auto)
-        have "\<And>x. x \<in> {u'..v'} \<Longrightarrow> (g has_real_derivative g' x) (at x within {u'..v'})"
-           using asm by (intro has_field_derivative_subset[OF derivg] set_mp[OF \<open>{u'..v'} \<subseteq> {a..b}\<close>]) auto
-        hence B: "\<And>x. min u' v' \<le> x \<Longrightarrow> x \<le> max u' v' \<Longrightarrow>
-                      (g has_vector_derivative g' x) (at x within {min u' v'..max u' v'})"
-            by (simp only: u'v' max_absorb2 min_absorb1)
-               (auto simp: has_field_derivative_iff_has_vector_derivative)
-        have "integrable lborel (\<lambda>x. indicator ({a..b} \<inter> g -` {u..v}) x *\<^sub>R g' x)"
-          by (rule set_integrable_subset[OF borel_integrable_atLeastAtMost'[OF contg']]) simp_all
-        hence "(\<integral>\<^sup>+x. ennreal (g' x) * indicator ({a..b} \<inter> g-` {u..v}) x \<partial>lborel) =
-                   LBINT x:{a..b} \<inter> g-`{u..v}. g' x"
-          by (subst nn_integral_eq_integral[symmetric])
-             (auto intro!: derivg_nonneg nn_integral_cong split: split_indicator)
-        also from interval_integral_FTC_finite[OF A B]
-            have "LBINT x:{a..b} \<inter> g-`{u..v}. g' x = v - u"
-                by (simp add: u'v' interval_integral_Icc \<open>u \<le> v\<close>)
-        finally have "(\<integral>\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \<inter> g -` {u..v}) x \<partial>lborel) =
-                           ennreal (v - u)" .
-      } note A = this
-
-      have "(\<integral>\<^sup>+x. indicator {c..d} (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) =
-               (\<integral>\<^sup>+ x. ennreal (g' x) * indicator ({a..b} \<inter> g -` {c..d}) x \<partial>lborel)"
-        by (intro nn_integral_cong) (simp split: split_indicator)
-      also have "{a..b} \<inter> g-`{c..d} = {a..b} \<inter> g-`{max (g a) c..min (g b) d}"
-        using \<open>a \<le> b\<close> \<open>c \<le> d\<close>
-        by (auto intro!: monog intro: order.trans)
-      also have "(\<integral>\<^sup>+ x. ennreal (g' x) * indicator ... x \<partial>lborel) =
-        (if max (g a) c \<le> min (g b) d then min (g b) d - max (g a) c else 0)"
-         using \<open>c \<le> d\<close> by (simp add: A)
-      also have "... = (\<integral>\<^sup>+ x. indicator ({g a..g b} \<inter> {c..d}) x \<partial>lborel)"
-        by (subst nn_integral_indicator) (auto intro!: measurable_sets Mg simp:)
-      also have "... = (\<integral>\<^sup>+ x. indicator {c..d} x * indicator {g a..g b} x \<partial>lborel)"
-        by (intro nn_integral_cong) (auto split: split_indicator)
-      finally show ?case ..
-
-      next
-
-      case (compl A)
-      note \<open>A \<in> sets borel\<close>[measurable]
-      from emeasure_mono[of "A \<inter> {g a..g b}" "{g a..g b}" lborel]
-      have [simp]: "emeasure lborel (A \<inter> {g a..g b}) \<noteq> top" by (auto simp: top_unique)
-      have [simp]: "g -` A \<inter> {a..b} \<in> sets borel"
-        by (rule set_borel_measurable_sets[OF Mg]) auto
-      have [simp]: "g -` (-A) \<inter> {a..b} \<in> sets borel"
-        by (rule set_borel_measurable_sets[OF Mg]) auto
-
-      have "(\<integral>\<^sup>+x. indicator (-A) x * indicator {g a..g b} x \<partial>lborel) =
-                (\<integral>\<^sup>+x. indicator (-A \<inter> {g a..g b}) x \<partial>lborel)"
-        by (rule nn_integral_cong) (simp split: split_indicator)
-      also from compl have "... = emeasure lborel ({g a..g b} - A)" using derivg_nonneg
-        by (simp add: vimage_Compl diff_eq Int_commute[of "-A"])
-      also have "{g a..g b} - A = {g a..g b} - A \<inter> {g a..g b}" by blast
-      also have "emeasure lborel ... = g b - g a - emeasure lborel (A \<inter> {g a..g b})"
-             using \<open>A \<in> sets borel\<close> by (subst emeasure_Diff) (auto simp: )
-     also have "emeasure lborel (A \<inter> {g a..g b}) =
-                    \<integral>\<^sup>+x. indicator A x * indicator {g a..g b} x \<partial>lborel"
-       using \<open>A \<in> sets borel\<close>
-       by (subst nn_integral_indicator[symmetric], simp, intro nn_integral_cong)
-          (simp split: split_indicator)
-      also have "... = \<integral>\<^sup>+ x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel" (is "_ = ?I")
-        by (subst compl.IH, intro nn_integral_cong) (simp split: split_indicator)
-      also have "g b - g a = LBINT x:{a..b}. g' x" using derivg'
-        by (intro integral_FTC_atLeastAtMost[symmetric])
-           (auto intro: continuous_on_subset[OF contg'] has_field_derivative_subset[OF derivg]
-                 has_vector_derivative_at_within)
-      also have "ennreal ... = \<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel"
-        using borel_integrable_atLeastAtMost'[OF contg']
-        by (subst nn_integral_eq_integral)
-           (simp_all add: mult.commute derivg_nonneg split: split_indicator)
-      also have Mg'': "(\<lambda>x. indicator (g -` A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x))
-                            \<in> borel_measurable borel" using Mg'
-        by (intro borel_measurable_times_ennreal borel_measurable_indicator)
-           (simp_all add: mult.commute)
-      have le: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel) \<le>
-                        (\<integral>\<^sup>+x. ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
-         by (intro nn_integral_mono) (simp split: split_indicator add: derivg_nonneg)
-      note integrable = borel_integrable_atLeastAtMost'[OF contg']
-      with le have notinf: "(\<integral>\<^sup>+x. indicator (g-`A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel) \<noteq> top"
-          by (auto simp: real_integrable_def nn_integral_set_ennreal mult.commute top_unique)
-      have "(\<integral>\<^sup>+ x. g' x * indicator {a..b} x \<partial>lborel) - ?I =
-                  \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) -
-                        indicator (g -` A \<inter> {a..b}) x * ennreal (g' x * indicator {a..b} x) \<partial>lborel"
-        apply (intro nn_integral_diff[symmetric])
-        apply (insert Mg', simp add: mult.commute) []
-        apply (insert Mg'', simp) []
-        apply (simp split: split_indicator add: derivg_nonneg)
-        apply (rule notinf)
-        apply (simp split: split_indicator add: derivg_nonneg)
-        done
-      also have "... = \<integral>\<^sup>+ x. indicator (-A) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
-        by (intro nn_integral_cong) (simp split: split_indicator)
-      finally show ?case .
-
-    next
-      case (union f)
-      then have [simp]: "\<And>i. {a..b} \<inter> g -` f i \<in> sets borel"
-        by (subst Int_commute, intro set_borel_measurable_sets[OF Mg]) auto
-      have "g -` (\<Union>i. f i) \<inter> {a..b} = (\<Union>i. {a..b} \<inter> g -` f i)" by auto
-      hence "g -` (\<Union>i. f i) \<inter> {a..b} \<in> sets borel" by (auto simp del: UN_simps)
-
-      have "(\<integral>\<^sup>+x. indicator (\<Union>i. f i) x * indicator {g a..g b} x \<partial>lborel) =
-                \<integral>\<^sup>+x. indicator (\<Union>i. {g a..g b} \<inter> f i) x \<partial>lborel"
-          by (intro nn_integral_cong) (simp split: split_indicator)
-      also from union have "... = emeasure lborel (\<Union>i. {g a..g b} \<inter> f i)" by simp
-      also from union have "... = (\<Sum>i. emeasure lborel ({g a..g b} \<inter> f i))"
-        by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def)
-      also from union have "... = (\<Sum>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel)" by simp
-      also have "(\<lambda>i. \<integral>\<^sup>+x. indicator ({g a..g b} \<inter> f i) x \<partial>lborel) =
-                           (\<lambda>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel)"
-        by (intro ext nn_integral_cong) (simp split: split_indicator)
-      also from union.IH have "(\<Sum>i. \<integral>\<^sup>+x. indicator (f i) x * indicator {g a..g b} x \<partial>lborel) =
-          (\<Sum>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)" by simp
-      also have "(\<lambda>i. \<integral>\<^sup>+ x. indicator (f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) =
-                         (\<lambda>i. \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x \<partial>lborel)"
-        by (intro ext nn_integral_cong) (simp split: split_indicator)
-      also have "(\<Sum>i. ... i) = \<integral>\<^sup>+ x. (\<Sum>i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) \<partial>lborel"
-        using Mg'
-        apply (intro nn_integral_suminf[symmetric])
-        apply (rule borel_measurable_times_ennreal, simp add: mult.commute)
-        apply (rule borel_measurable_indicator, subst sets_lborel)
-        apply (simp_all split: split_indicator add: derivg_nonneg)
-        done
-      also have "(\<lambda>x i. ennreal (g' x * indicator {a..b} x) * indicator ({a..b} \<inter> g -` f i) x) =
-                      (\<lambda>x i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x)"
-        by (intro ext) (simp split: split_indicator)
-      also have "(\<integral>\<^sup>+ x. (\<Sum>i. ennreal (g' x * indicator {a..b} x) * indicator (g -` f i) x) \<partial>lborel) =
-                     \<integral>\<^sup>+ x. ennreal (g' x * indicator {a..b} x) * (\<Sum>i. indicator (g -` f i) x) \<partial>lborel"
-        by (intro nn_integral_cong) (auto split: split_indicator simp: derivg_nonneg)
-      also from union have "(\<lambda>x. \<Sum>i. indicator (g -` f i) x :: ennreal) = (\<lambda>x. indicator (\<Union>i. g -` f i) x)"
-        by (intro ext suminf_indicator) (auto simp: disjoint_family_on_def)
-      also have "(\<integral>\<^sup>+x. ennreal (g' x * indicator {a..b} x) * ... x \<partial>lborel) =
-                    (\<integral>\<^sup>+x. indicator (\<Union>i. f i) (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
-       by (intro nn_integral_cong) (simp split: split_indicator)
-      finally show ?case .
-  qed
-
-next
-  case (mult f c)
-    note Mf[measurable] = \<open>f \<in> borel_measurable borel\<close>
-    let ?I = "indicator {a..b}"
-    have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
-      by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf])
-         (simp_all add: mult.commute)
-    also have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. f (g x) * ennreal (g' x) * ?I x)"
-      by (intro ext) (simp split: split_indicator)
-    finally have Mf': "(\<lambda>x. f (g x) * ennreal (g' x) * ?I x) \<in> borel_measurable borel" .
-    with mult show ?case
-      by (subst (1 2 3) mult_ac, subst (1 2) nn_integral_cmult) (simp_all add: mult_ac)
-
-next
-  case (add f2 f1)
-    let ?I = "indicator {a..b}"
-    {
-      fix f :: "real \<Rightarrow> ennreal" assume Mf: "f \<in> borel_measurable borel"
-      have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
-        by (intro borel_measurable_times_ennreal measurable_compose[OF _ Mf])
-           (simp_all add:  mult.commute)
-      also have "(\<lambda>x. f (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. f (g x) * ennreal (g' x) * ?I x)"
-        by (intro ext) (simp split: split_indicator)
-      finally have "(\<lambda>x. f (g x) * ennreal (g' x) * ?I x) \<in> borel_measurable borel" .
-    } note Mf' = this[OF \<open>f1 \<in> borel_measurable borel\<close>] this[OF \<open>f2 \<in> borel_measurable borel\<close>]
-
-    have "(\<integral>\<^sup>+ x. (f1 x + f2 x) * indicator {g a..g b} x \<partial>lborel) =
-             (\<integral>\<^sup>+ x. f1 x * indicator {g a..g b} x + f2 x * indicator {g a..g b} x \<partial>lborel)"
-      by (intro nn_integral_cong) (simp split: split_indicator)
-    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel) +
-                                (\<integral>\<^sup>+ x. f2 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
-      by (simp_all add: nn_integral_add)
-    also from add have "... = (\<integral>\<^sup>+ x. f1 (g x) * ennreal (g' x) * indicator {a..b} x +
-                                      f2 (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
-      by (intro nn_integral_add[symmetric])
-         (auto simp add: Mf' derivg_nonneg split: split_indicator)
-    also have "... = \<integral>\<^sup>+ x. (f1 (g x) + f2 (g x)) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
-      by (intro nn_integral_cong) (simp split: split_indicator add: distrib_right)
-    finally show ?case .
-
-next
-  case (sup F)
-  {
-    fix i
-    let ?I = "indicator {a..b}"
-    have "(\<lambda>x. F i (g x * ?I x) * ennreal (g' x * ?I x)) \<in> borel_measurable borel" using Mg Mg'
-      by (rule_tac borel_measurable_times_ennreal, rule_tac measurable_compose[OF _ sup.hyps(1)])
-         (simp_all add: mult.commute)
-    also have "(\<lambda>x. F i (g x * ?I x) * ennreal (g' x * ?I x)) = (\<lambda>x. F i (g x) * ennreal (g' x) * ?I x)"
-      by (intro ext) (simp split: split_indicator)
-     finally have "... \<in> borel_measurable borel" .
-  } note Mf' = this
-
-    have "(\<integral>\<^sup>+x. (SUP i. F i x) * indicator {g a..g b} x \<partial>lborel) =
-               \<integral>\<^sup>+x. (SUP i. F i x* indicator {g a..g b} x) \<partial>lborel"
-      by (intro nn_integral_cong) (simp split: split_indicator)
-    also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i x* indicator {g a..g b} x \<partial>lborel)"
-      by (intro nn_integral_monotone_convergence_SUP)
-         (auto simp: incseq_def le_fun_def split: split_indicator)
-    also from sup have "... = (SUP i. \<integral>\<^sup>+x. F i (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel)"
-      by simp
-    also from sup have "... =  \<integral>\<^sup>+x. (SUP i. F i (g x) * ennreal (g' x) * indicator {a..b} x) \<partial>lborel"
-      by (intro nn_integral_monotone_convergence_SUP[symmetric])
-         (auto simp: incseq_def le_fun_def derivg_nonneg Mf' split: split_indicator
-               intro!: mult_right_mono)
-    also from sup have "... = \<integral>\<^sup>+x. (SUP i. F i (g x)) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
-      by (subst mult.assoc, subst mult.commute, subst SUP_mult_left_ennreal)
-         (auto split: split_indicator simp: derivg_nonneg mult_ac)
-    finally show ?case by simp
-  qed
-qed
-
-lemma nn_integral_substitution:
-  fixes f :: "real \<Rightarrow> real"
-  assumes Mf[measurable]: "set_borel_measurable borel {g a..g b} f"
-  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'"
-  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
-  assumes "a \<le> b"
-  shows "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
-             (\<integral>\<^sup>+x. f (g x) * g' x * indicator {a..b} x \<partial>lborel)"
-proof (cases "a = b")
-  assume "a \<noteq> b"
-  with \<open>a \<le> b\<close> have "a < b" by auto
-  let ?f' = "\<lambda>x. f x * indicator {g a..g b} x"
-
-  from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
-    by (rule deriv_nonneg_imp_mono) simp_all
-  have bounds: "\<And>x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g x \<ge> g a" "\<And>x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g x \<le> g b"
-    by (auto intro: monog)
-
-  from derivg_nonneg have nonneg:
-    "\<And>f x. x \<ge> a \<Longrightarrow> x \<le> b \<Longrightarrow> g' x \<noteq> 0 \<Longrightarrow> f x * ennreal (g' x) \<ge> 0 \<Longrightarrow> f x \<ge> 0"
-    by (force simp: field_simps)
-  have nonneg': "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<not> 0 \<le> f (g x) \<Longrightarrow> 0 \<le> f (g x) * g' x \<Longrightarrow> g' x = 0"
-    by (metis atLeastAtMost_iff derivg_nonneg eq_iff mult_eq_0_iff mult_le_0_iff)
-
-  have "(\<integral>\<^sup>+x. f x * indicator {g a..g b} x \<partial>lborel) =
-            (\<integral>\<^sup>+x. ennreal (?f' x) * indicator {g a..g b} x \<partial>lborel)"
-    by (intro nn_integral_cong)
-       (auto split: split_indicator split_max simp: zero_ennreal.rep_eq ennreal_neg)
-  also have "... = \<integral>\<^sup>+ x. ?f' (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel" using Mf
-    by (subst nn_integral_substitution_aux[OF _ _ derivg contg' derivg_nonneg \<open>a < b\<close>])
-       (auto simp add: mult.commute)
-  also have "... = \<integral>\<^sup>+ x. f (g x) * ennreal (g' x) * indicator {a..b} x \<partial>lborel"
-    by (intro nn_integral_cong) (auto split: split_indicator simp: max_def dest: bounds)
-  also have "... = \<integral>\<^sup>+x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel"
-    by (intro nn_integral_cong) (auto simp: mult.commute derivg_nonneg ennreal_mult' split: split_indicator)
-  finally show ?thesis .
-qed auto
-
-lemma integral_substitution:
-  assumes integrable: "set_integrable lborel {g a..g b} f"
-  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'"
-  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
-  assumes "a \<le> b"
-  shows "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
-    and "(LBINT x. f x * indicator {g a..g b} x) = (LBINT x. f (g x) * g' x * indicator {a..b} x)"
-proof-
-  from derivg have contg: "continuous_on {a..b} g" by (rule has_real_derivative_imp_continuous_on)
-  with contg' have Mg: "set_borel_measurable borel {a..b} g"
-    and Mg': "set_borel_measurable borel {a..b} g'"
-    by (simp_all only: set_measurable_continuous_on_ivl)
-  from derivg derivg_nonneg have monog: "\<And>x y. a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b \<Longrightarrow> g x \<le> g y"
-    by (rule deriv_nonneg_imp_mono) simp_all
-
-  have "(\<lambda>x. ennreal (f x) * indicator {g a..g b} x) =
-           (\<lambda>x. ennreal (f x * indicator {g a..g b} x))"
-    by (intro ext) (simp split: split_indicator)
-  with integrable have M1: "(\<lambda>x. f x * indicator {g a..g b} x) \<in> borel_measurable borel"
-    unfolding real_integrable_def by (force simp: mult.commute)
-  from integrable have M2: "(\<lambda>x. -f x * indicator {g a..g b} x) \<in> borel_measurable borel"
-    unfolding real_integrable_def by (force simp: mult.commute)
-
-  have "LBINT x. (f x :: real) * indicator {g a..g b} x =
-          enn2real (\<integral>\<^sup>+ x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) -
-          enn2real (\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel)" using integrable
-    by (subst real_lebesgue_integral_def) (simp_all add: nn_integral_set_ennreal mult.commute)
-  also have *: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {g a..g b} x \<partial>lborel) =
-      (\<integral>\<^sup>+x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel)"
-    by (intro nn_integral_cong) (simp split: split_indicator)
-  also from M1 * have A: "(\<integral>\<^sup>+ x. ennreal (f x * indicator {g a..g b} x) \<partial>lborel) =
-                            (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel)"
-    by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
-       (auto simp: nn_integral_set_ennreal mult.commute)
-  also have **: "(\<integral>\<^sup>+ x. ennreal (- (f x)) * indicator {g a..g b} x \<partial>lborel) =
-      (\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel)"
-    by (intro nn_integral_cong) (simp split: split_indicator)
-  also from M2 ** have B: "(\<integral>\<^sup>+ x. ennreal (- (f x) * indicator {g a..g b} x) \<partial>lborel) =
-        (\<integral>\<^sup>+ x. ennreal (- (f (g x)) * g' x * indicator {a..b} x) \<partial>lborel)"
-    by (subst nn_integral_substitution[OF _ derivg contg' derivg_nonneg \<open>a \<le> b\<close>])
-       (auto simp: nn_integral_set_ennreal mult.commute)
-
-  also {
-    from integrable have Mf: "set_borel_measurable borel {g a..g b} f"
-      unfolding real_integrable_def by simp
-    from borel_measurable_times[OF measurable_compose[OF Mg Mf] Mg']
-      have "(\<lambda>x. f (g x * indicator {a..b} x) * indicator {g a..g b} (g x * indicator {a..b} x) *
-                     (g' x * indicator {a..b} x)) \<in> borel_measurable borel"  (is "?f \<in> _")
-      by (simp add: mult.commute)
-    also have "?f = (\<lambda>x. f (g x) * g' x * indicator {a..b} x)"
-      using monog by (intro ext) (auto split: split_indicator)
-    finally show "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
-      using A B integrable unfolding real_integrable_def
-      by (simp_all add: nn_integral_set_ennreal mult.commute)
-  } note integrable' = this
-
-  have "enn2real (\<integral>\<^sup>+ x. ennreal (f (g x) * g' x * indicator {a..b} x) \<partial>lborel) -
-                  enn2real (\<integral>\<^sup>+ x. ennreal (-f (g x) * g' x * indicator {a..b} x) \<partial>lborel) =
-                (LBINT x. f (g x) * g' x * indicator {a..b} x)" using integrable'
-    by (subst real_lebesgue_integral_def) (simp_all add: field_simps)
-  finally show "(LBINT x. f x * indicator {g a..g b} x) =
-                     (LBINT x. f (g x) * g' x * indicator {a..b} x)" .
-qed
-
-lemma interval_integral_substitution:
-  assumes integrable: "set_integrable lborel {g a..g b} f"
-  assumes derivg: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_real_derivative g' x) (at x)"
-  assumes contg': "continuous_on {a..b} g'"
-  assumes derivg_nonneg: "\<And>x. x \<in> {a..b} \<Longrightarrow> g' x \<ge> 0"
-  assumes "a \<le> b"
-  shows "set_integrable lborel {a..b} (\<lambda>x. f (g x) * g' x)"
-    and "(LBINT x=g a..g b. f x) = (LBINT x=a..b. f (g x) * g' x)"
-  apply (rule integral_substitution[OF assms], simp, simp)
-  apply (subst (1 2) interval_integral_Icc, fact)
-  apply (rule deriv_nonneg_imp_mono[OF derivg derivg_nonneg], simp, simp, fact)
-  using integral_substitution(2)[OF assms]
-  apply (simp add: mult.commute)
-  done
-
-lemma set_borel_integrable_singleton[simp]:
-  "set_integrable lborel {x} (f :: real \<Rightarrow> real)"
-  by (subst integrable_discrete_difference[where X="{x}" and g="\<lambda>_. 0"]) auto
-
-end
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1382 +0,0 @@
-(*  Title:      HOL/Probability/Lebesgue_Measure.thy
-    Author:     Johannes Hölzl, TU München
-    Author:     Robert Himmelmann, TU München
-    Author:     Jeremy Avigad
-    Author:     Luke Serafin
-*)
-
-section \<open>Lebesgue measure\<close>
-
-theory Lebesgue_Measure
-  imports Finite_Product_Measure Bochner_Integration Caratheodory
-begin
-
-subsection \<open>Every right continuous and nondecreasing function gives rise to a measure\<close>
-
-definition interval_measure :: "(real \<Rightarrow> real) \<Rightarrow> real measure" where
-  "interval_measure F = extend_measure UNIV {(a, b). a \<le> b} (\<lambda>(a, b). {a <.. b}) (\<lambda>(a, b). ennreal (F b - F a))"
-
-lemma emeasure_interval_measure_Ioc:
-  assumes "a \<le> b"
-  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
-  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
-  shows "emeasure (interval_measure F) {a <.. b} = F b - F a"
-proof (rule extend_measure_caratheodory_pair[OF interval_measure_def \<open>a \<le> b\<close>])
-  show "semiring_of_sets UNIV {{a<..b} |a b :: real. a \<le> b}"
-  proof (unfold_locales, safe)
-    fix a b c d :: real assume *: "a \<le> b" "c \<le> d"
-    then show "\<exists>C\<subseteq>{{a<..b} |a b. a \<le> b}. finite C \<and> disjoint C \<and> {a<..b} - {c<..d} = \<Union>C"
-    proof cases
-      let ?C = "{{a<..b}}"
-      assume "b < c \<or> d \<le> a \<or> d \<le> c"
-      with * have "?C \<subseteq> {{a<..b} |a b. a \<le> b} \<and> finite ?C \<and> disjoint ?C \<and> {a<..b} - {c<..d} = \<Union>?C"
-        by (auto simp add: disjoint_def)
-      thus ?thesis ..
-    next
-      let ?C = "{{a<..c}, {d<..b}}"
-      assume "\<not> (b < c \<or> d \<le> a \<or> d \<le> c)"
-      with * have "?C \<subseteq> {{a<..b} |a b. a \<le> b} \<and> finite ?C \<and> disjoint ?C \<and> {a<..b} - {c<..d} = \<Union>?C"
-        by (auto simp add: disjoint_def Ioc_inj) (metis linear)+
-      thus ?thesis ..
-    qed
-  qed (auto simp: Ioc_inj, metis linear)
-next
-  fix l r :: "nat \<Rightarrow> real" and a b :: real
-  assume l_r[simp]: "\<And>n. l n \<le> r n" and "a \<le> b" and disj: "disjoint_family (\<lambda>n. {l n<..r n})"
-  assume lr_eq_ab: "(\<Union>i. {l i<..r i}) = {a<..b}"
-
-  have [intro, simp]: "\<And>a b. a \<le> b \<Longrightarrow> F a \<le> F b"
-    by (auto intro!: l_r mono_F)
-
-  { fix S :: "nat set" assume "finite S"
-    moreover note \<open>a \<le> b\<close>
-    moreover have "\<And>i. i \<in> S \<Longrightarrow> {l i <.. r i} \<subseteq> {a <.. b}"
-      unfolding lr_eq_ab[symmetric] by auto
-    ultimately have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<le> F b - F a"
-    proof (induction S arbitrary: a rule: finite_psubset_induct)
-      case (psubset S)
-      show ?case
-      proof cases
-        assume "\<exists>i\<in>S. l i < r i"
-        with \<open>finite S\<close> have "Min (l ` {i\<in>S. l i < r i}) \<in> l ` {i\<in>S. l i < r i}"
-          by (intro Min_in) auto
-        then obtain m where m: "m \<in> S" "l m < r m" "l m = Min (l ` {i\<in>S. l i < r i})"
-          by fastforce
-
-        have "(\<Sum>i\<in>S. F (r i) - F (l i)) = (F (r m) - F (l m)) + (\<Sum>i\<in>S - {m}. F (r i) - F (l i))"
-          using m psubset by (intro setsum.remove) auto
-        also have "(\<Sum>i\<in>S - {m}. F (r i) - F (l i)) \<le> F b - F (r m)"
-        proof (intro psubset.IH)
-          show "S - {m} \<subset> S"
-            using \<open>m\<in>S\<close> by auto
-          show "r m \<le> b"
-            using psubset.prems(2)[OF \<open>m\<in>S\<close>] \<open>l m < r m\<close> by auto
-        next
-          fix i assume "i \<in> S - {m}"
-          then have i: "i \<in> S" "i \<noteq> m" by auto
-          { assume i': "l i < r i" "l i < r m"
-            with \<open>finite S\<close> i m have "l m \<le> l i"
-              by auto
-            with i' have "{l i <.. r i} \<inter> {l m <.. r m} \<noteq> {}"
-              by auto
-            then have False
-              using disjoint_family_onD[OF disj, of i m] i by auto }
-          then have "l i \<noteq> r i \<Longrightarrow> r m \<le> l i"
-            unfolding not_less[symmetric] using l_r[of i] by auto
-          then show "{l i <.. r i} \<subseteq> {r m <.. b}"
-            using psubset.prems(2)[OF \<open>i\<in>S\<close>] by auto
-        qed
-        also have "F (r m) - F (l m) \<le> F (r m) - F a"
-          using psubset.prems(2)[OF \<open>m \<in> S\<close>] \<open>l m < r m\<close>
-          by (auto simp add: Ioc_subset_iff intro!: mono_F)
-        finally show ?case
-          by (auto intro: add_mono)
-      qed (auto simp add: \<open>a \<le> b\<close> less_le)
-    qed }
-  note claim1 = this
-
-  (* second key induction: a lower bound on the measures of any finite collection of Ai's
-     that cover an interval {u..v} *)
-
-  { fix S u v and l r :: "nat \<Rightarrow> real"
-    assume "finite S" "\<And>i. i\<in>S \<Longrightarrow> l i < r i" "{u..v} \<subseteq> (\<Union>i\<in>S. {l i<..< r i})"
-    then have "F v - F u \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
-    proof (induction arbitrary: v u rule: finite_psubset_induct)
-      case (psubset S)
-      show ?case
-      proof cases
-        assume "S = {}" then show ?case
-          using psubset by (simp add: mono_F)
-      next
-        assume "S \<noteq> {}"
-        then obtain j where "j \<in> S"
-          by auto
-
-        let ?R = "r j < u \<or> l j > v \<or> (\<exists>i\<in>S-{j}. l i \<le> l j \<and> r j \<le> r i)"
-        show ?case
-        proof cases
-          assume "?R"
-          with \<open>j \<in> S\<close> psubset.prems have "{u..v} \<subseteq> (\<Union>i\<in>S-{j}. {l i<..< r i})"
-            apply (auto simp: subset_eq Ball_def)
-            apply (metis Diff_iff less_le_trans leD linear singletonD)
-            apply (metis Diff_iff less_le_trans leD linear singletonD)
-            apply (metis order_trans less_le_not_le linear)
-            done
-          with \<open>j \<in> S\<close> have "F v - F u \<le> (\<Sum>i\<in>S - {j}. F (r i) - F (l i))"
-            by (intro psubset) auto
-          also have "\<dots> \<le> (\<Sum>i\<in>S. F (r i) - F (l i))"
-            using psubset.prems
-            by (intro setsum_mono2 psubset) (auto intro: less_imp_le)
-          finally show ?thesis .
-        next
-          assume "\<not> ?R"
-          then have j: "u \<le> r j" "l j \<le> v" "\<And>i. i \<in> S - {j} \<Longrightarrow> r i < r j \<or> l i > l j"
-            by (auto simp: not_less)
-          let ?S1 = "{i \<in> S. l i < l j}"
-          let ?S2 = "{i \<in> S. r i > r j}"
-
-          have "(\<Sum>i\<in>S. F (r i) - F (l i)) \<ge> (\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i))"
-            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
-            by (intro setsum_mono2) (auto intro: less_imp_le)
-          also have "(\<Sum>i\<in>?S1 \<union> ?S2 \<union> {j}. F (r i) - F (l i)) =
-            (\<Sum>i\<in>?S1. F (r i) - F (l i)) + (\<Sum>i\<in>?S2 . F (r i) - F (l i)) + (F (r j) - F (l j))"
-            using psubset(1) psubset.prems(1) j
-            apply (subst setsum.union_disjoint)
-            apply simp_all
-            apply (subst setsum.union_disjoint)
-            apply auto
-            apply (metis less_le_not_le)
-            done
-          also (xtrans) have "(\<Sum>i\<in>?S1. F (r i) - F (l i)) \<ge> F (l j) - F u"
-            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
-            apply (intro psubset.IH psubset)
-            apply (auto simp: subset_eq Ball_def)
-            apply (metis less_le_trans not_le)
-            done
-          also (xtrans) have "(\<Sum>i\<in>?S2. F (r i) - F (l i)) \<ge> F v - F (r j)"
-            using \<open>j \<in> S\<close> \<open>finite S\<close> psubset.prems j
-            apply (intro psubset.IH psubset)
-            apply (auto simp: subset_eq Ball_def)
-            apply (metis le_less_trans not_le)
-            done
-          finally (xtrans) show ?case
-            by (auto simp: add_mono)
-        qed
-      qed
-    qed }
-  note claim2 = this
-
-  (* now prove the inequality going the other way *)
-  have "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i)))"
-  proof (rule ennreal_le_epsilon)
-    fix epsilon :: real assume egt0: "epsilon > 0"
-    have "\<forall>i. \<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
-    proof
-      fix i
-      note right_cont_F [of "r i"]
-      thus "\<exists>d>0. F (r i + d) < F (r i) + epsilon / 2^(i+2)"
-        apply -
-        apply (subst (asm) continuous_at_right_real_increasing)
-        apply (rule mono_F, assumption)
-        apply (drule_tac x = "epsilon / 2 ^ (i + 2)" in spec)
-        apply (erule impE)
-        using egt0 by (auto simp add: field_simps)
-    qed
-    then obtain delta where
-        deltai_gt0: "\<And>i. delta i > 0" and
-        deltai_prop: "\<And>i. F (r i + delta i) < F (r i) + epsilon / 2^(i+2)"
-      by metis
-    have "\<exists>a' > a. F a' - F a < epsilon / 2"
-      apply (insert right_cont_F [of a])
-      apply (subst (asm) continuous_at_right_real_increasing)
-      using mono_F apply force
-      apply (drule_tac x = "epsilon / 2" in spec)
-      using egt0 unfolding mult.commute [of 2] by force
-    then obtain a' where a'lea [arith]: "a' > a" and
-      a_prop: "F a' - F a < epsilon / 2"
-      by auto
-    define S' where "S' = {i. l i < r i}"
-    obtain S :: "nat set" where
-      "S \<subseteq> S'" and finS: "finite S" and
-      Sprop: "{a'..b} \<subseteq> (\<Union>i \<in> S. {l i<..<r i + delta i})"
-    proof (rule compactE_image)
-      show "compact {a'..b}"
-        by (rule compact_Icc)
-      show "\<forall>i \<in> S'. open ({l i<..<r i + delta i})" by auto
-      have "{a'..b} \<subseteq> {a <.. b}"
-        by auto
-      also have "{a <.. b} = (\<Union>i\<in>S'. {l i<..r i})"
-        unfolding lr_eq_ab[symmetric] by (fastforce simp add: S'_def intro: less_le_trans)
-      also have "\<dots> \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})"
-        apply (intro UN_mono)
-        apply (auto simp: S'_def)
-        apply (cut_tac i=i in deltai_gt0)
-        apply simp
-        done
-      finally show "{a'..b} \<subseteq> (\<Union>i \<in> S'. {l i<..<r i + delta i})" .
-    qed
-    with S'_def have Sprop2: "\<And>i. i \<in> S \<Longrightarrow> l i < r i" by auto
-    from finS have "\<exists>n. \<forall>i \<in> S. i \<le> n"
-      by (subst finite_nat_set_iff_bounded_le [symmetric])
-    then obtain n where Sbound [rule_format]: "\<forall>i \<in> S. i \<le> n" ..
-    have "F b - F a' \<le> (\<Sum>i\<in>S. F (r i + delta i) - F (l i))"
-      apply (rule claim2 [rule_format])
-      using finS Sprop apply auto
-      apply (frule Sprop2)
-      apply (subgoal_tac "delta i > 0")
-      apply arith
-      by (rule deltai_gt0)
-    also have "... \<le> (\<Sum>i \<in> S. F(r i) - F(l i) + epsilon / 2^(i+2))"
-      apply (rule setsum_mono)
-      apply simp
-      apply (rule order_trans)
-      apply (rule less_imp_le)
-      apply (rule deltai_prop)
-      by auto
-    also have "... = (\<Sum>i \<in> S. F(r i) - F(l i)) +
-        (epsilon / 4) * (\<Sum>i \<in> S. (1 / 2)^i)" (is "_ = ?t + _")
-      by (subst setsum.distrib) (simp add: field_simps setsum_right_distrib)
-    also have "... \<le> ?t + (epsilon / 4) * (\<Sum> i < Suc n. (1 / 2)^i)"
-      apply (rule add_left_mono)
-      apply (rule mult_left_mono)
-      apply (rule setsum_mono2)
-      using egt0 apply auto
-      by (frule Sbound, auto)
-    also have "... \<le> ?t + (epsilon / 2)"
-      apply (rule add_left_mono)
-      apply (subst geometric_sum)
-      apply auto
-      apply (rule mult_left_mono)
-      using egt0 apply auto
-      done
-    finally have aux2: "F b - F a' \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2"
-      by simp
-
-    have "F b - F a = (F b - F a') + (F a' - F a)"
-      by auto
-    also have "... \<le> (F b - F a') + epsilon / 2"
-      using a_prop by (intro add_left_mono) simp
-    also have "... \<le> (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon / 2 + epsilon / 2"
-      apply (intro add_right_mono)
-      apply (rule aux2)
-      done
-    also have "... = (\<Sum>i\<in>S. F (r i) - F (l i)) + epsilon"
-      by auto
-    also have "... \<le> (\<Sum>i\<le>n. F (r i) - F (l i)) + epsilon"
-      using finS Sbound Sprop by (auto intro!: add_right_mono setsum_mono3)
-    finally have "ennreal (F b - F a) \<le> (\<Sum>i\<le>n. ennreal (F (r i) - F (l i))) + epsilon"
-      using egt0 by (simp add: ennreal_plus[symmetric] setsum_nonneg del: ennreal_plus)
-    then show "ennreal (F b - F a) \<le> (\<Sum>i. ennreal (F (r i) - F (l i))) + (epsilon :: real)"
-      by (rule order_trans) (auto intro!: add_mono setsum_le_suminf simp del: setsum_ennreal)
-  qed
-  moreover have "(\<Sum>i. ennreal (F (r i) - F (l i))) \<le> ennreal (F b - F a)"
-    using \<open>a \<le> b\<close> by (auto intro!: suminf_le_const ennreal_le_iff[THEN iffD2] claim1)
-  ultimately show "(\<Sum>n. ennreal (F (r n) - F (l n))) = ennreal (F b - F a)"
-    by (rule antisym[rotated])
-qed (auto simp: Ioc_inj mono_F)
-
-lemma measure_interval_measure_Ioc:
-  assumes "a \<le> b"
-  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
-  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
-  shows "measure (interval_measure F) {a <.. b} = F b - F a"
-  unfolding measure_def
-  apply (subst emeasure_interval_measure_Ioc)
-  apply fact+
-  apply (simp add: assms)
-  done
-
-lemma emeasure_interval_measure_Ioc_eq:
-  "(\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y) \<Longrightarrow> (\<And>a. continuous (at_right a) F) \<Longrightarrow>
-    emeasure (interval_measure F) {a <.. b} = (if a \<le> b then F b - F a else 0)"
-  using emeasure_interval_measure_Ioc[of a b F] by auto
-
-lemma sets_interval_measure [simp, measurable_cong]: "sets (interval_measure F) = sets borel"
-  apply (simp add: sets_extend_measure interval_measure_def borel_sigma_sets_Ioc)
-  apply (rule sigma_sets_eqI)
-  apply auto
-  apply (case_tac "a \<le> ba")
-  apply (auto intro: sigma_sets.Empty)
-  done
-
-lemma space_interval_measure [simp]: "space (interval_measure F) = UNIV"
-  by (simp add: interval_measure_def space_extend_measure)
-
-lemma emeasure_interval_measure_Icc:
-  assumes "a \<le> b"
-  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
-  assumes cont_F : "continuous_on UNIV F"
-  shows "emeasure (interval_measure F) {a .. b} = F b - F a"
-proof (rule tendsto_unique)
-  { fix a b :: real assume "a \<le> b" then have "emeasure (interval_measure F) {a <.. b} = F b - F a"
-      using cont_F
-      by (subst emeasure_interval_measure_Ioc)
-         (auto intro: mono_F continuous_within_subset simp: continuous_on_eq_continuous_within) }
-  note * = this
-
-  let ?F = "interval_measure F"
-  show "((\<lambda>a. F b - F a) \<longlongrightarrow> emeasure ?F {a..b}) (at_left a)"
-  proof (rule tendsto_at_left_sequentially)
-    show "a - 1 < a" by simp
-    fix X assume "\<And>n. X n < a" "incseq X" "X \<longlonglongrightarrow> a"
-    with \<open>a \<le> b\<close> have "(\<lambda>n. emeasure ?F {X n<..b}) \<longlonglongrightarrow> emeasure ?F (\<Inter>n. {X n <..b})"
-      apply (intro Lim_emeasure_decseq)
-      apply (auto simp: decseq_def incseq_def emeasure_interval_measure_Ioc *)
-      apply force
-      apply (subst (asm ) *)
-      apply (auto intro: less_le_trans less_imp_le)
-      done
-    also have "(\<Inter>n. {X n <..b}) = {a..b}"
-      using \<open>\<And>n. X n < a\<close>
-      apply auto
-      apply (rule LIMSEQ_le_const2[OF \<open>X \<longlonglongrightarrow> a\<close>])
-      apply (auto intro: less_imp_le)
-      apply (auto intro: less_le_trans)
-      done
-    also have "(\<lambda>n. emeasure ?F {X n<..b}) = (\<lambda>n. F b - F (X n))"
-      using \<open>\<And>n. X n < a\<close> \<open>a \<le> b\<close> by (subst *) (auto intro: less_imp_le less_le_trans)
-    finally show "(\<lambda>n. F b - F (X n)) \<longlonglongrightarrow> emeasure ?F {a..b}" .
-  qed
-  show "((\<lambda>a. ennreal (F b - F a)) \<longlongrightarrow> F b - F a) (at_left a)"
-    by (rule continuous_on_tendsto_compose[where g="\<lambda>x. x" and s=UNIV])
-       (auto simp: continuous_on_ennreal continuous_on_diff cont_F continuous_on_const)
-qed (rule trivial_limit_at_left_real)
-
-lemma sigma_finite_interval_measure:
-  assumes mono_F: "\<And>x y. x \<le> y \<Longrightarrow> F x \<le> F y"
-  assumes right_cont_F : "\<And>a. continuous (at_right a) F"
-  shows "sigma_finite_measure (interval_measure F)"
-  apply unfold_locales
-  apply (intro exI[of _ "(\<lambda>(a, b). {a <.. b}) ` (\<rat> \<times> \<rat>)"])
-  apply (auto intro!: Rats_no_top_le Rats_no_bot_less countable_rat simp: emeasure_interval_measure_Ioc_eq[OF assms])
-  done
-
-subsection \<open>Lebesgue-Borel measure\<close>
-
-definition lborel :: "('a :: euclidean_space) measure" where
-  "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. interval_measure (\<lambda>x. x)) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
-
-lemma
-  shows sets_lborel[simp, measurable_cong]: "sets lborel = sets borel"
-    and space_lborel[simp]: "space lborel = space borel"
-    and measurable_lborel1[simp]: "measurable M lborel = measurable M borel"
-    and measurable_lborel2[simp]: "measurable lborel M = measurable borel M"
-  by (simp_all add: lborel_def)
-
-context
-begin
-
-interpretation sigma_finite_measure "interval_measure (\<lambda>x. x)"
-  by (rule sigma_finite_interval_measure) auto
-interpretation finite_product_sigma_finite "\<lambda>_. interval_measure (\<lambda>x. x)" Basis
-  proof qed simp
-
-lemma lborel_eq_real: "lborel = interval_measure (\<lambda>x. x)"
-  unfolding lborel_def Basis_real_def
-  using distr_id[of "interval_measure (\<lambda>x. x)"]
-  by (subst distr_component[symmetric])
-     (simp_all add: distr_distr comp_def del: distr_id cong: distr_cong)
-
-lemma lborel_eq: "lborel = distr (\<Pi>\<^sub>M b\<in>Basis. lborel) borel (\<lambda>f. \<Sum>b\<in>Basis. f b *\<^sub>R b)"
-  by (subst lborel_def) (simp add: lborel_eq_real)
-
-lemma nn_integral_lborel_setprod:
-  assumes [measurable]: "\<And>b. b \<in> Basis \<Longrightarrow> f b \<in> borel_measurable borel"
-  assumes nn[simp]: "\<And>b x. b \<in> Basis \<Longrightarrow> 0 \<le> f b x"
-  shows "(\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. f b (x \<bullet> b)) \<partial>lborel) = (\<Prod>b\<in>Basis. (\<integral>\<^sup>+x. f b x \<partial>lborel))"
-  by (simp add: lborel_def nn_integral_distr product_nn_integral_setprod
-                product_nn_integral_singleton)
-
-lemma emeasure_lborel_Icc[simp]:
-  fixes l u :: real
-  assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l .. u} = u - l"
-proof -
-  have "((\<lambda>f. f 1) -` {l..u} \<inter> space (Pi\<^sub>M {1} (\<lambda>b. interval_measure (\<lambda>x. x)))) = {1::real} \<rightarrow>\<^sub>E {l..u}"
-    by (auto simp: space_PiM)
-  then show ?thesis
-    by (simp add: lborel_def emeasure_distr emeasure_PiM emeasure_interval_measure_Icc continuous_on_id)
-qed
-
-lemma emeasure_lborel_Icc_eq: "emeasure lborel {l .. u} = ennreal (if l \<le> u then u - l else 0)"
-  by simp
-
-lemma emeasure_lborel_cbox[simp]:
-  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
-  shows "emeasure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-proof -
-  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (cbox l u)"
-    by (auto simp: fun_eq_iff cbox_def split: split_indicator)
-  then have "emeasure lborel (cbox l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b .. u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
-    by simp
-  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
-  finally show ?thesis .
-qed
-
-lemma AE_lborel_singleton: "AE x in lborel::'a::euclidean_space measure. x \<noteq> c"
-  using SOME_Basis AE_discrete_difference [of "{c}" lborel] emeasure_lborel_cbox [of c c]
-  by (auto simp add: cbox_sing setprod_constant power_0_left)
-
-lemma emeasure_lborel_Ioo[simp]:
-  assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l <..< u} = ennreal (u - l)"
-proof -
-  have "emeasure lborel {l <..< u} = emeasure lborel {l .. u}"
-    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
-  then show ?thesis
-    by simp
-qed
-
-lemma emeasure_lborel_Ioc[simp]:
-  assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l <.. u} = ennreal (u - l)"
-proof -
-  have "emeasure lborel {l <.. u} = emeasure lborel {l .. u}"
-    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
-  then show ?thesis
-    by simp
-qed
-
-lemma emeasure_lborel_Ico[simp]:
-  assumes [simp]: "l \<le> u"
-  shows "emeasure lborel {l ..< u} = ennreal (u - l)"
-proof -
-  have "emeasure lborel {l ..< u} = emeasure lborel {l .. u}"
-    using AE_lborel_singleton[of u] AE_lborel_singleton[of l] by (intro emeasure_eq_AE) auto
-  then show ?thesis
-    by simp
-qed
-
-lemma emeasure_lborel_box[simp]:
-  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
-  shows "emeasure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-proof -
-  have "(\<lambda>x. \<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b) :: ennreal) = indicator (box l u)"
-    by (auto simp: fun_eq_iff box_def split: split_indicator)
-  then have "emeasure lborel (box l u) = (\<integral>\<^sup>+x. (\<Prod>b\<in>Basis. indicator {l\<bullet>b <..< u\<bullet>b} (x \<bullet> b)) \<partial>lborel)"
-    by simp
-  also have "\<dots> = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-    by (subst nn_integral_lborel_setprod) (simp_all add: setprod_ennreal inner_diff_left)
-  finally show ?thesis .
-qed
-
-lemma emeasure_lborel_cbox_eq:
-  "emeasure lborel (cbox l u) = (if \<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b then \<Prod>b\<in>Basis. (u - l) \<bullet> b else 0)"
-  using box_eq_empty(2)[THEN iffD2, of u l] by (auto simp: not_le)
-
-lemma emeasure_lborel_box_eq:
-  "emeasure lborel (box l u) = (if \<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b then \<Prod>b\<in>Basis. (u - l) \<bullet> b else 0)"
-  using box_eq_empty(1)[THEN iffD2, of u l] by (auto simp: not_le dest!: less_imp_le) force
-
-lemma
-  fixes l u :: real
-  assumes [simp]: "l \<le> u"
-  shows measure_lborel_Icc[simp]: "measure lborel {l .. u} = u - l"
-    and measure_lborel_Ico[simp]: "measure lborel {l ..< u} = u - l"
-    and measure_lborel_Ioc[simp]: "measure lborel {l <.. u} = u - l"
-    and measure_lborel_Ioo[simp]: "measure lborel {l <..< u} = u - l"
-  by (simp_all add: measure_def)
-
-lemma
-  assumes [simp]: "\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
-  shows measure_lborel_box[simp]: "measure lborel (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-    and measure_lborel_cbox[simp]: "measure lborel (cbox l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-  by (simp_all add: measure_def inner_diff_left setprod_nonneg)
-
-lemma sigma_finite_lborel: "sigma_finite_measure lborel"
-proof
-  show "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets lborel \<and> \<Union>A = space lborel \<and> (\<forall>a\<in>A. emeasure lborel a \<noteq> \<infinity>)"
-    by (intro exI[of _ "range (\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One))"])
-       (auto simp: emeasure_lborel_cbox_eq UN_box_eq_UNIV)
-qed
-
-end
-
-lemma emeasure_lborel_UNIV: "emeasure lborel (UNIV::'a::euclidean_space set) = \<infinity>"
-proof -
-  { fix n::nat
-    let ?Ba = "Basis :: 'a set"
-    have "real n \<le> (2::real) ^ card ?Ba * real n"
-      by (simp add: mult_le_cancel_right1)
-    also
-    have "... \<le> (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba"
-      apply (rule mult_left_mono)
-      apply (metis DIM_positive One_nat_def less_eq_Suc_le less_imp_le of_nat_le_iff of_nat_power self_le_power zero_less_Suc)
-      apply (simp add: DIM_positive)
-      done
-    finally have "real n \<le> (2::real) ^ card ?Ba * real (Suc n) ^ card ?Ba" .
-  } note [intro!] = this
-  show ?thesis
-    unfolding UN_box_eq_UNIV[symmetric]
-    apply (subst SUP_emeasure_incseq[symmetric])
-    apply (auto simp: incseq_def subset_box inner_add_left setprod_constant
-      simp del: Sup_eq_top_iff SUP_eq_top_iff
-      intro!: ennreal_SUP_eq_top)
-    done
-qed
-
-lemma emeasure_lborel_singleton[simp]: "emeasure lborel {x} = 0"
-  using emeasure_lborel_cbox[of x x] nonempty_Basis
-  by (auto simp del: emeasure_lborel_cbox nonempty_Basis simp add: cbox_sing setprod_constant)
-
-lemma emeasure_lborel_countable:
-  fixes A :: "'a::euclidean_space set"
-  assumes "countable A"
-  shows "emeasure lborel A = 0"
-proof -
-  have "A \<subseteq> (\<Union>i. {from_nat_into A i})" using from_nat_into_surj assms by force
-  then have "emeasure lborel A \<le> emeasure lborel (\<Union>i. {from_nat_into A i})"
-    by (intro emeasure_mono) auto
-  also have "emeasure lborel (\<Union>i. {from_nat_into A i}) = 0"
-    by (rule emeasure_UN_eq_0) auto
-  finally show ?thesis
-    by (auto simp add: )
-qed
-
-lemma countable_imp_null_set_lborel: "countable A \<Longrightarrow> A \<in> null_sets lborel"
-  by (simp add: null_sets_def emeasure_lborel_countable sets.countable)
-
-lemma finite_imp_null_set_lborel: "finite A \<Longrightarrow> A \<in> null_sets lborel"
-  by (intro countable_imp_null_set_lborel countable_finite)
-
-lemma lborel_neq_count_space[simp]: "lborel \<noteq> count_space (A::('a::ordered_euclidean_space) set)"
-proof
-  assume asm: "lborel = count_space A"
-  have "space lborel = UNIV" by simp
-  hence [simp]: "A = UNIV" by (subst (asm) asm) (simp only: space_count_space)
-  have "emeasure lborel {undefined::'a} = 1"
-      by (subst asm, subst emeasure_count_space_finite) auto
-  moreover have "emeasure lborel {undefined} \<noteq> 1" by simp
-  ultimately show False by contradiction
-qed
-
-subsection \<open>Affine transformation on the Lebesgue-Borel\<close>
-
-lemma lborel_eqI:
-  fixes M :: "'a::euclidean_space measure"
-  assumes emeasure_eq: "\<And>l u. (\<And>b. b \<in> Basis \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b) \<Longrightarrow> emeasure M (box l u) = (\<Prod>b\<in>Basis. (u - l) \<bullet> b)"
-  assumes sets_eq: "sets M = sets borel"
-  shows "lborel = M"
-proof (rule measure_eqI_generator_eq)
-  let ?E = "range (\<lambda>(a, b). box a b::'a set)"
-  show "Int_stable ?E"
-    by (auto simp: Int_stable_def box_Int_box)
-
-  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
-    by (simp_all add: borel_eq_box sets_eq)
-
-  let ?A = "\<lambda>n::nat. box (- (real n *\<^sub>R One)) (real n *\<^sub>R One) :: 'a set"
-  show "range ?A \<subseteq> ?E" "(\<Union>i. ?A i) = UNIV"
-    unfolding UN_box_eq_UNIV by auto
-
-  { fix i show "emeasure lborel (?A i) \<noteq> \<infinity>" by auto }
-  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
-      apply (auto simp: emeasure_eq emeasure_lborel_box_eq )
-      apply (subst box_eq_empty(1)[THEN iffD2])
-      apply (auto intro: less_imp_le simp: not_le)
-      done }
-qed
-
-lemma lborel_affine:
-  fixes t :: "'a::euclidean_space" assumes "c \<noteq> 0"
-  shows "lborel = density (distr lborel borel (\<lambda>x. t + c *\<^sub>R x)) (\<lambda>_. \<bar>c\<bar>^DIM('a))" (is "_ = ?D")
-proof (rule lborel_eqI)
-  let ?B = "Basis :: 'a set"
-  fix l u assume le: "\<And>b. b \<in> ?B \<Longrightarrow> l \<bullet> b \<le> u \<bullet> b"
-  show "emeasure ?D (box l u) = (\<Prod>b\<in>?B. (u - l) \<bullet> b)"
-  proof cases
-    assume "0 < c"
-    then have "(\<lambda>x. t + c *\<^sub>R x) -` box l u = box ((l - t) /\<^sub>R c) ((u - t) /\<^sub>R c)"
-      by (auto simp: field_simps box_def inner_simps)
-    with \<open>0 < c\<close> show ?thesis
-      using le
-      by (auto simp: field_simps inner_simps setprod_dividef setprod_constant setprod_nonneg
-                     ennreal_mult[symmetric] emeasure_density nn_integral_distr emeasure_distr
-                     nn_integral_cmult emeasure_lborel_box_eq borel_measurable_indicator')
-  next
-    assume "\<not> 0 < c" with \<open>c \<noteq> 0\<close> have "c < 0" by auto
-    then have "box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c) = (\<lambda>x. t + c *\<^sub>R x) -` box l u"
-      by (auto simp: field_simps box_def inner_simps)
-    then have *: "\<And>x. indicator (box l u) (t + c *\<^sub>R x) = (indicator (box ((u - t) /\<^sub>R c) ((l - t) /\<^sub>R c)) x :: ennreal)"
-      by (auto split: split_indicator)
-    have **: "(\<Prod>x\<in>Basis. (l \<bullet> x - u \<bullet> x) / c) = (\<Prod>x\<in>Basis. u \<bullet> x - l \<bullet> x) / (-c) ^ card (Basis::'a set)"
-      using \<open>c < 0\<close>
-      by (auto simp add: field_simps setprod_dividef[symmetric] setprod_constant[symmetric]
-               intro!: setprod.cong)
-    show ?thesis
-      using \<open>c < 0\<close> le
-      by (auto simp: * ** field_simps emeasure_density nn_integral_distr nn_integral_cmult
-                     emeasure_lborel_box_eq inner_simps setprod_nonneg ennreal_mult[symmetric]
-                     borel_measurable_indicator')
-  qed
-qed simp
-
-lemma lborel_real_affine:
-  "c \<noteq> 0 \<Longrightarrow> lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. ennreal (abs c))"
-  using lborel_affine[of c t] by simp
-
-lemma AE_borel_affine:
-  fixes P :: "real \<Rightarrow> bool"
-  shows "c \<noteq> 0 \<Longrightarrow> Measurable.pred borel P \<Longrightarrow> AE x in lborel. P x \<Longrightarrow> AE x in lborel. P (t + c * x)"
-  by (subst lborel_real_affine[where t="- t / c" and c="1 / c"])
-     (simp_all add: AE_density AE_distr_iff field_simps)
-
-lemma nn_integral_real_affine:
-  fixes c :: real assumes [measurable]: "f \<in> borel_measurable borel" and c: "c \<noteq> 0"
-  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = \<bar>c\<bar> * (\<integral>\<^sup>+x. f (t + c * x) \<partial>lborel)"
-  by (subst lborel_real_affine[OF c, of t])
-     (simp add: nn_integral_density nn_integral_distr nn_integral_cmult)
-
-lemma lborel_integrable_real_affine:
-  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  assumes f: "integrable lborel f"
-  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x))"
-  using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded
-  by (subst (asm) nn_integral_real_affine[where c=c and t=t]) (auto simp: ennreal_mult_less_top)
-
-lemma lborel_integrable_real_affine_iff:
-  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x)) \<longleftrightarrow> integrable lborel f"
-  using
-    lborel_integrable_real_affine[of f c t]
-    lborel_integrable_real_affine[of "\<lambda>x. f (t + c * x)" "1/c" "-t/c"]
-  by (auto simp add: field_simps)
-
-lemma lborel_integral_real_affine:
-  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" and c :: real
-  assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lborel) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f (t + c * x) \<partial>lborel)"
-proof cases
-  assume f[measurable]: "integrable lborel f" then show ?thesis
-    using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t]
-    by (subst lborel_real_affine[OF c, of t])
-       (simp add: integral_density integral_distr)
-next
-  assume "\<not> integrable lborel f" with c show ?thesis
-    by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq)
-qed
-
-lemma divideR_right:
-  fixes x y :: "'a::real_normed_vector"
-  shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x"
-  using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp
-
-lemma lborel_has_bochner_integral_real_affine_iff:
-  fixes x :: "'a :: {banach, second_countable_topology}"
-  shows "c \<noteq> 0 \<Longrightarrow>
-    has_bochner_integral lborel f x \<longleftrightarrow>
-    has_bochner_integral lborel (\<lambda>x. f (t + c * x)) (x /\<^sub>R \<bar>c\<bar>)"
-  unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff
-  by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong)
-
-lemma lborel_distr_uminus: "distr lborel borel uminus = (lborel :: real measure)"
-  by (subst lborel_real_affine[of "-1" 0])
-     (auto simp: density_1 one_ennreal_def[symmetric])
-
-lemma lborel_distr_mult:
-  assumes "(c::real) \<noteq> 0"
-  shows "distr lborel borel (op * c) = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
-proof-
-  have "distr lborel borel (op * c) = distr lborel lborel (op * c)" by (simp cong: distr_cong)
-  also from assms have "... = density lborel (\<lambda>_. inverse \<bar>c\<bar>)"
-    by (subst lborel_real_affine[of "inverse c" 0]) (auto simp: o_def distr_density_distr)
-  finally show ?thesis .
-qed
-
-lemma lborel_distr_mult':
-  assumes "(c::real) \<noteq> 0"
-  shows "lborel = density (distr lborel borel (op * c)) (\<lambda>_. \<bar>c\<bar>)"
-proof-
-  have "lborel = density lborel (\<lambda>_. 1)" by (rule density_1[symmetric])
-  also from assms have "(\<lambda>_. 1 :: ennreal) = (\<lambda>_. inverse \<bar>c\<bar> * \<bar>c\<bar>)" by (intro ext) simp
-  also have "density lborel ... = density (density lborel (\<lambda>_. inverse \<bar>c\<bar>)) (\<lambda>_. \<bar>c\<bar>)"
-    by (subst density_density_eq) (auto simp: ennreal_mult)
-  also from assms have "density lborel (\<lambda>_. inverse \<bar>c\<bar>) = distr lborel borel (op * c)"
-    by (rule lborel_distr_mult[symmetric])
-  finally show ?thesis .
-qed
-
-lemma lborel_distr_plus: "distr lborel borel (op + c) = (lborel :: real measure)"
-  by (subst lborel_real_affine[of 1 c]) (auto simp: density_1 one_ennreal_def[symmetric])
-
-interpretation lborel: sigma_finite_measure lborel
-  by (rule sigma_finite_lborel)
-
-interpretation lborel_pair: pair_sigma_finite lborel lborel ..
-
-lemma lborel_prod:
-  "lborel \<Otimes>\<^sub>M lborel = (lborel :: ('a::euclidean_space \<times> 'b::euclidean_space) measure)"
-proof (rule lborel_eqI[symmetric], clarify)
-  fix la ua :: 'a and lb ub :: 'b
-  assume lu: "\<And>a b. (a, b) \<in> Basis \<Longrightarrow> (la, lb) \<bullet> (a, b) \<le> (ua, ub) \<bullet> (a, b)"
-  have [simp]:
-    "\<And>b. b \<in> Basis \<Longrightarrow> la \<bullet> b \<le> ua \<bullet> b"
-    "\<And>b. b \<in> Basis \<Longrightarrow> lb \<bullet> b \<le> ub \<bullet> b"
-    "inj_on (\<lambda>u. (u, 0)) Basis" "inj_on (\<lambda>u. (0, u)) Basis"
-    "(\<lambda>u. (u, 0)) ` Basis \<inter> (\<lambda>u. (0, u)) ` Basis = {}"
-    "box (la, lb) (ua, ub) = box la ua \<times> box lb ub"
-    using lu[of _ 0] lu[of 0] by (auto intro!: inj_onI simp add: Basis_prod_def ball_Un box_def)
-  show "emeasure (lborel \<Otimes>\<^sub>M lborel) (box (la, lb) (ua, ub)) =
-      ennreal (setprod (op \<bullet> ((ua, ub) - (la, lb))) Basis)"
-    by (simp add: lborel.emeasure_pair_measure_Times Basis_prod_def setprod.union_disjoint
-                  setprod.reindex ennreal_mult inner_diff_left setprod_nonneg)
-qed (simp add: borel_prod[symmetric])
-
-(* FIXME: conversion in measurable prover *)
-lemma lborelD_Collect[measurable (raw)]: "{x\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" by simp
-lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp
-
-subsection \<open>Equivalence Lebesgue integral on @{const lborel} and HK-integral\<close>
-
-lemma has_integral_measure_lborel:
-  fixes A :: "'a::euclidean_space set"
-  assumes A[measurable]: "A \<in> sets borel" and finite: "emeasure lborel A < \<infinity>"
-  shows "((\<lambda>x. 1) has_integral measure lborel A) A"
-proof -
-  { fix l u :: 'a
-    have "((\<lambda>x. 1) has_integral measure lborel (box l u)) (box l u)"
-    proof cases
-      assume "\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b"
-      then show ?thesis
-        apply simp
-        apply (subst has_integral_restrict[symmetric, OF box_subset_cbox])
-        apply (subst has_integral_spike_interior_eq[where g="\<lambda>_. 1"])
-        using has_integral_const[of "1::real" l u]
-        apply (simp_all add: inner_diff_left[symmetric] content_cbox_cases)
-        done
-    next
-      assume "\<not> (\<forall>b\<in>Basis. l \<bullet> b \<le> u \<bullet> b)"
-      then have "box l u = {}"
-        unfolding box_eq_empty by (auto simp: not_le intro: less_imp_le)
-      then show ?thesis
-        by simp
-    qed }
-  note has_integral_box = this
-
-  { fix a b :: 'a let ?M = "\<lambda>A. measure lborel (A \<inter> box a b)"
-    have "Int_stable  (range (\<lambda>(a, b). box a b))"
-      by (auto simp: Int_stable_def box_Int_box)
-    moreover have "(range (\<lambda>(a, b). box a b)) \<subseteq> Pow UNIV"
-      by auto
-    moreover have "A \<in> sigma_sets UNIV (range (\<lambda>(a, b). box a b))"
-       using A unfolding borel_eq_box by simp
-    ultimately have "((\<lambda>x. 1) has_integral ?M A) (A \<inter> box a b)"
-    proof (induction rule: sigma_sets_induct_disjoint)
-      case (basic A) then show ?case
-        by (auto simp: box_Int_box has_integral_box)
-    next
-      case empty then show ?case
-        by simp
-    next
-      case (compl A)
-      then have [measurable]: "A \<in> sets borel"
-        by (simp add: borel_eq_box)
-
-      have "((\<lambda>x. 1) has_integral ?M (box a b)) (box a b)"
-        by (simp add: has_integral_box)
-      moreover have "((\<lambda>x. if x \<in> A \<inter> box a b then 1 else 0) has_integral ?M A) (box a b)"
-        by (subst has_integral_restrict) (auto intro: compl)
-      ultimately have "((\<lambda>x. 1 - (if x \<in> A \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
-        by (rule has_integral_sub)
-      then have "((\<lambda>x. (if x \<in> (UNIV - A) \<inter> box a b then 1 else 0)) has_integral ?M (box a b) - ?M A) (box a b)"
-        by (rule has_integral_cong[THEN iffD1, rotated 1]) auto
-      then have "((\<lambda>x. 1) has_integral ?M (box a b) - ?M A) ((UNIV - A) \<inter> box a b)"
-        by (subst (asm) has_integral_restrict) auto
-      also have "?M (box a b) - ?M A = ?M (UNIV - A)"
-        by (subst measure_Diff[symmetric]) (auto simp: emeasure_lborel_box_eq Diff_Int_distrib2)
-      finally show ?case .
-    next
-      case (union F)
-      then have [measurable]: "\<And>i. F i \<in> sets borel"
-        by (simp add: borel_eq_box subset_eq)
-      have "((\<lambda>x. if x \<in> UNION UNIV F \<inter> box a b then 1 else 0) has_integral ?M (\<Union>i. F i)) (box a b)"
-      proof (rule has_integral_monotone_convergence_increasing)
-        let ?f = "\<lambda>k x. \<Sum>i<k. if x \<in> F i \<inter> box a b then 1 else 0 :: real"
-        show "\<And>k. (?f k has_integral (\<Sum>i<k. ?M (F i))) (box a b)"
-          using union.IH by (auto intro!: has_integral_setsum simp del: Int_iff)
-        show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
-          by (intro setsum_mono2) auto
-        from union(1) have *: "\<And>x i j. x \<in> F i \<Longrightarrow> x \<in> F j \<longleftrightarrow> j = i"
-          by (auto simp add: disjoint_family_on_def)
-        show "\<And>x. (\<lambda>k. ?f k x) \<longlonglongrightarrow> (if x \<in> UNION UNIV F \<inter> box a b then 1 else 0)"
-          apply (auto simp: * setsum.If_cases Iio_Int_singleton)
-          apply (rule_tac k="Suc xa" in LIMSEQ_offset)
-          apply simp
-          done
-        have *: "emeasure lborel ((\<Union>x. F x) \<inter> box a b) \<le> emeasure lborel (box a b)"
-          by (intro emeasure_mono) auto
-
-        with union(1) show "(\<lambda>k. \<Sum>i<k. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
-          unfolding sums_def[symmetric] UN_extend_simps
-          by (intro measure_UNION) (auto simp: disjoint_family_on_def emeasure_lborel_box_eq top_unique)
-      qed
-      then show ?case
-        by (subst (asm) has_integral_restrict) auto
-    qed }
-  note * = this
-
-  show ?thesis
-  proof (rule has_integral_monotone_convergence_increasing)
-    let ?B = "\<lambda>n::nat. box (- real n *\<^sub>R One) (real n *\<^sub>R One) :: 'a set"
-    let ?f = "\<lambda>n::nat. \<lambda>x. if x \<in> A \<inter> ?B n then 1 else 0 :: real"
-    let ?M = "\<lambda>n. measure lborel (A \<inter> ?B n)"
-
-    show "\<And>n::nat. (?f n has_integral ?M n) A"
-      using * by (subst has_integral_restrict) simp_all
-    show "\<And>k x. ?f k x \<le> ?f (Suc k) x"
-      by (auto simp: box_def)
-    { fix x assume "x \<in> A"
-      moreover have "(\<lambda>k. indicator (A \<inter> ?B k) x :: real) \<longlonglongrightarrow> indicator (\<Union>k::nat. A \<inter> ?B k) x"
-        by (intro LIMSEQ_indicator_incseq) (auto simp: incseq_def box_def)
-      ultimately show "(\<lambda>k. if x \<in> A \<inter> ?B k then 1 else 0::real) \<longlonglongrightarrow> 1"
-        by (simp add: indicator_def UN_box_eq_UNIV) }
-
-    have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> emeasure lborel (\<Union>n::nat. A \<inter> ?B n)"
-      by (intro Lim_emeasure_incseq) (auto simp: incseq_def box_def)
-    also have "(\<lambda>n. emeasure lborel (A \<inter> ?B n)) = (\<lambda>n. measure lborel (A \<inter> ?B n))"
-    proof (intro ext emeasure_eq_ennreal_measure)
-      fix n have "emeasure lborel (A \<inter> ?B n) \<le> emeasure lborel (?B n)"
-        by (intro emeasure_mono) auto
-      then show "emeasure lborel (A \<inter> ?B n) \<noteq> top"
-        by (auto simp: top_unique)
-    qed
-    finally show "(\<lambda>n. measure lborel (A \<inter> ?B n)) \<longlonglongrightarrow> measure lborel A"
-      using emeasure_eq_ennreal_measure[of lborel A] finite
-      by (simp add: UN_box_eq_UNIV less_top)
-  qed
-qed
-
-lemma nn_integral_has_integral:
-  fixes f::"'a::euclidean_space \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
-  shows "(f has_integral r) UNIV"
-using f proof (induct f arbitrary: r rule: borel_measurable_induct_real)
-  case (set A)
-  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
-    by (intro has_integral_measure_lborel) (auto simp: ennreal_indicator)
-  with set show ?case
-    by (simp add: ennreal_indicator measure_def) (simp add: indicator_def)
-next
-  case (mult g c)
-  then have "ennreal c * (\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal r"
-    by (subst nn_integral_cmult[symmetric]) (auto simp: ennreal_mult)
-  with \<open>0 \<le> r\<close> \<open>0 \<le> c\<close>
-  obtain r' where "(c = 0 \<and> r = 0) \<or> (0 \<le> r' \<and> (\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel) = ennreal r' \<and> r = c * r')"
-    by (cases "\<integral>\<^sup>+ x. ennreal (g x) \<partial>lborel" rule: ennreal_cases)
-       (auto split: if_split_asm simp: ennreal_mult_top ennreal_mult[symmetric])
-  with mult show ?case
-    by (auto intro!: has_integral_cmult_real)
-next
-  case (add g h)
-  then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lborel) = (\<integral>\<^sup>+ x. h x \<partial>lborel) + (\<integral>\<^sup>+ x. g x \<partial>lborel)"
-    by (simp add: nn_integral_add)
-  with add obtain a b where "0 \<le> a" "0 \<le> b" "(\<integral>\<^sup>+ x. h x \<partial>lborel) = ennreal a" "(\<integral>\<^sup>+ x. g x \<partial>lborel) = ennreal b" "r = a + b"
-    by (cases "\<integral>\<^sup>+ x. h x \<partial>lborel" "\<integral>\<^sup>+ x. g x \<partial>lborel" rule: ennreal2_cases)
-       (auto simp: add_top nn_integral_add top_add ennreal_plus[symmetric] simp del: ennreal_plus)
-  with add show ?case
-    by (auto intro!: has_integral_add)
-next
-  case (seq U)
-  note seq(1)[measurable] and f[measurable]
-
-  { fix i x
-    have "U i x \<le> f x"
-      using seq(5)
-      apply (rule LIMSEQ_le_const)
-      using seq(4)
-      apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def)
-      done }
-  note U_le_f = this
-
-  { fix i
-    have "(\<integral>\<^sup>+x. U i x \<partial>lborel) \<le> (\<integral>\<^sup>+x. f x \<partial>lborel)"
-      using seq(2) f(2) U_le_f by (intro nn_integral_mono) simp
-    then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p" "p \<le> r" "0 \<le> p"
-      using seq(6) \<open>0\<le>r\<close> by (cases "\<integral>\<^sup>+x. U i x \<partial>lborel" rule: ennreal_cases) (auto simp: top_unique)
-    moreover note seq
-    ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lborel) = ennreal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV"
-      by auto }
-  then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ennreal (U i x) \<partial>lborel) = ennreal (p i)"
-    and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i"
-    and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis
-
-  have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique)
-
-  have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) \<longlonglongrightarrow> integral UNIV f"
-  proof (rule monotone_convergence_increasing)
-    show "\<forall>k. U k integrable_on UNIV" using U_int by auto
-    show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using \<open>incseq U\<close> by (auto simp: incseq_def le_fun_def)
-    then show "bounded {integral UNIV (U k) |k. True}"
-      using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r])
-    show "\<forall>x\<in>UNIV. (\<lambda>k. U k x) \<longlonglongrightarrow> f x"
-      using seq by auto
-  qed
-  moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lborel)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. f x \<partial>lborel)"
-    using seq f(2) U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto
-  ultimately have "integral UNIV f = r"
-    by (auto simp add: bnd int_eq p seq intro: LIMSEQ_unique)
-  with * show ?case
-    by (simp add: has_integral_integral)
-qed
-
-lemma nn_integral_lborel_eq_integral:
-  fixes f::"'a::euclidean_space \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
-  shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = integral UNIV f"
-proof -
-  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
-    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
-  then show ?thesis
-    using nn_integral_has_integral[OF f(1,2) r] by (simp add: integral_unique)
-qed
-
-lemma nn_integral_integrable_on:
-  fixes f::"'a::euclidean_space \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lborel) < \<infinity>"
-  shows "f integrable_on UNIV"
-proof -
-  from f(3) obtain r where r: "(\<integral>\<^sup>+x. f x \<partial>lborel) = ennreal r" "0 \<le> r"
-    by (cases "\<integral>\<^sup>+x. f x \<partial>lborel" rule: ennreal_cases) auto
-  then show ?thesis
-    by (intro has_integral_integrable[where i=r] nn_integral_has_integral[where r=r] f)
-qed
-
-lemma nn_integral_has_integral_lborel:
-  fixes f :: "'a::euclidean_space \<Rightarrow> real"
-  assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"
-  assumes I: "(f has_integral I) UNIV"
-  shows "integral\<^sup>N lborel f = I"
-proof -
-  from f_borel have "(\<lambda>x. ennreal (f x)) \<in> borel_measurable lborel" by auto
-  from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
-  let ?B = "\<lambda>i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One) :: 'a set"
-
-  note F(1)[THEN borel_measurable_simple_function, measurable]
-
-  have "0 \<le> I"
-    using I by (rule has_integral_nonneg) (simp add: nonneg)
-
-  have F_le_f: "enn2real (F i x) \<le> f x" for i x
-    using F(3,4)[where x=x] nonneg SUP_upper[of i UNIV "\<lambda>i. F i x"]
-    by (cases "F i x" rule: ennreal_cases) auto
-  let ?F = "\<lambda>i x. F i x * indicator (?B i) x"
-  have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (SUP i. integral\<^sup>N lborel (\<lambda>x. ?F i x))"
-  proof (subst nn_integral_monotone_convergence_SUP[symmetric])
-    { fix x
-      obtain j where j: "x \<in> ?B j"
-        using UN_box_eq_UNIV by auto
-
-      have "ennreal (f x) = (SUP i. F i x)"
-        using F(4)[of x] nonneg[of x] by (simp add: max_def)
-      also have "\<dots> = (SUP i. ?F i x)"
-      proof (rule SUP_eq)
-        fix i show "\<exists>j\<in>UNIV. F i x \<le> ?F j x"
-          using j F(2)
-          by (intro bexI[of _ "max i j"])
-             (auto split: split_max split_indicator simp: incseq_def le_fun_def box_def)
-      qed (auto intro!: F split: split_indicator)
-      finally have "ennreal (f x) =  (SUP i. ?F i x)" . }
-    then show "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) = (\<integral>\<^sup>+ x. (SUP i. ?F i x) \<partial>lborel)"
-      by simp
-  qed (insert F, auto simp: incseq_def le_fun_def box_def split: split_indicator)
-  also have "\<dots> \<le> ennreal I"
-  proof (rule SUP_least)
-    fix i :: nat
-    have finite_F: "(\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel) < \<infinity>"
-    proof (rule nn_integral_bound_simple_function)
-      have "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} \<le>
-        emeasure lborel (?B i)"
-        by (intro emeasure_mono)  (auto split: split_indicator)
-      then show "emeasure lborel {x \<in> space lborel. ennreal (enn2real (F i x) * indicator (?B i) x) \<noteq> 0} < \<infinity>"
-        by (auto simp: less_top[symmetric] top_unique)
-    qed (auto split: split_indicator
-              intro!: F simple_function_compose1[where g="enn2real"] simple_function_ennreal)
-
-    have int_F: "(\<lambda>x. enn2real (F i x) * indicator (?B i) x) integrable_on UNIV"
-      using F(4) finite_F
-      by (intro nn_integral_integrable_on) (auto split: split_indicator simp: enn2real_nonneg)
-
-    have "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) =
-      (\<integral>\<^sup>+ x. ennreal (enn2real (F i x) * indicator (?B i) x) \<partial>lborel)"
-      using F(3,4)
-      by (intro nn_integral_cong) (auto simp: image_iff eq_commute split: split_indicator)
-    also have "\<dots> = ennreal (integral UNIV (\<lambda>x. enn2real (F i x) * indicator (?B i) x))"
-      using F
-      by (intro nn_integral_lborel_eq_integral[OF _ _ finite_F])
-         (auto split: split_indicator intro: enn2real_nonneg)
-    also have "\<dots> \<le> ennreal I"
-      by (auto intro!: has_integral_le[OF integrable_integral[OF int_F] I] nonneg F_le_f
-               simp: \<open>0 \<le> I\<close> split: split_indicator )
-    finally show "(\<integral>\<^sup>+ x. F i x * indicator (?B i) x \<partial>lborel) \<le> ennreal I" .
-  qed
-  finally have "(\<integral>\<^sup>+ x. ennreal (f x) \<partial>lborel) < \<infinity>"
-    by (auto simp: less_top[symmetric] top_unique)
-  from nn_integral_lborel_eq_integral[OF assms(1,2) this] I show ?thesis
-    by (simp add: integral_unique)
-qed
-
-lemma has_integral_iff_emeasure_lborel:
-  fixes A :: "'a::euclidean_space set"
-  assumes A[measurable]: "A \<in> sets borel" and [simp]: "0 \<le> r"
-  shows "((\<lambda>x. 1) has_integral r) A \<longleftrightarrow> emeasure lborel A = ennreal r"
-proof (cases "emeasure lborel A = \<infinity>")
-  case emeasure_A: True
-  have "\<not> (\<lambda>x. 1::real) integrable_on A"
-  proof
-    assume int: "(\<lambda>x. 1::real) integrable_on A"
-    then have "(indicator A::'a \<Rightarrow> real) integrable_on UNIV"
-      unfolding indicator_def[abs_def] integrable_restrict_univ .
-    then obtain r where "((indicator A::'a\<Rightarrow>real) has_integral r) UNIV"
-      by auto
-    from nn_integral_has_integral_lborel[OF _ _ this] emeasure_A show False
-      by (simp add: ennreal_indicator)
-  qed
-  with emeasure_A show ?thesis
-    by auto
-next
-  case False
-  then have "((\<lambda>x. 1) has_integral measure lborel A) A"
-    by (simp add: has_integral_measure_lborel less_top)
-  with False show ?thesis
-    by (auto simp: emeasure_eq_ennreal_measure has_integral_unique)
-qed
-
-lemma has_integral_integral_real:
-  fixes f::"'a::euclidean_space \<Rightarrow> real"
-  assumes f: "integrable lborel f"
-  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
-using f proof induct
-  case (base A c) then show ?case
-    by (auto intro!: has_integral_mult_left simp: )
-       (simp add: emeasure_eq_ennreal_measure indicator_def has_integral_measure_lborel)
-next
-  case (add f g) then show ?case
-    by (auto intro!: has_integral_add)
-next
-  case (lim f s)
-  show ?case
-  proof (rule has_integral_dominated_convergence)
-    show "\<And>i. (s i has_integral integral\<^sup>L lborel (s i)) UNIV" by fact
-    show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV"
-      using \<open>integrable lborel f\<close>
-      by (intro nn_integral_integrable_on)
-         (auto simp: integrable_iff_bounded abs_mult  nn_integral_cmult ennreal_mult ennreal_mult_less_top)
-    show "\<And>k. \<forall>x\<in>UNIV. norm (s k x) \<le> norm (2 * f x)"
-      using lim by (auto simp add: abs_mult)
-    show "\<forall>x\<in>UNIV. (\<lambda>k. s k x) \<longlonglongrightarrow> f x"
-      using lim by auto
-    show "(\<lambda>k. integral\<^sup>L lborel (s k)) \<longlonglongrightarrow> integral\<^sup>L lborel f"
-      using lim lim(1)[THEN borel_measurable_integrable]
-      by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto
-  qed
-qed
-
-context
-  fixes f::"'a::euclidean_space \<Rightarrow> 'b::euclidean_space"
-begin
-
-lemma has_integral_integral_lborel:
-  assumes f: "integrable lborel f"
-  shows "(f has_integral (integral\<^sup>L lborel f)) UNIV"
-proof -
-  have "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b)) UNIV"
-    using f by (intro has_integral_setsum finite_Basis ballI has_integral_scaleR_left has_integral_integral_real) auto
-  also have eq_f: "(\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) = f"
-    by (simp add: fun_eq_iff euclidean_representation)
-  also have "(\<Sum>b\<in>Basis. integral\<^sup>L lborel (\<lambda>x. f x \<bullet> b) *\<^sub>R b) = integral\<^sup>L lborel f"
-    using f by (subst (2) eq_f[symmetric]) simp
-  finally show ?thesis .
-qed
-
-lemma integrable_on_lborel: "integrable lborel f \<Longrightarrow> f integrable_on UNIV"
-  using has_integral_integral_lborel by auto
-
-lemma integral_lborel: "integrable lborel f \<Longrightarrow> integral UNIV f = (\<integral>x. f x \<partial>lborel)"
-  using has_integral_integral_lborel by auto
-
-end
-
-subsection \<open>Fundamental Theorem of Calculus for the Lebesgue integral\<close>
-
-lemma emeasure_bounded_finite:
-  assumes "bounded A" shows "emeasure lborel A < \<infinity>"
-proof -
-  from bounded_subset_cbox[OF \<open>bounded A\<close>] obtain a b where "A \<subseteq> cbox a b"
-    by auto
-  then have "emeasure lborel A \<le> emeasure lborel (cbox a b)"
-    by (intro emeasure_mono) auto
-  then show ?thesis
-    by (auto simp: emeasure_lborel_cbox_eq setprod_nonneg less_top[symmetric] top_unique split: if_split_asm)
-qed
-
-lemma emeasure_compact_finite: "compact A \<Longrightarrow> emeasure lborel A < \<infinity>"
-  using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded)
-
-lemma borel_integrable_compact:
-  fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}"
-  assumes "compact S" "continuous_on S f"
-  shows "integrable lborel (\<lambda>x. indicator S x *\<^sub>R f x)"
-proof cases
-  assume "S \<noteq> {}"
-  have "continuous_on S (\<lambda>x. norm (f x))"
-    using assms by (intro continuous_intros)
-  from continuous_attains_sup[OF \<open>compact S\<close> \<open>S \<noteq> {}\<close> this]
-  obtain M where M: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> M"
-    by auto
-
-  show ?thesis
-  proof (rule integrable_bound)
-    show "integrable lborel (\<lambda>x. indicator S x * M)"
-      using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left)
-    show "(\<lambda>x. indicator S x *\<^sub>R f x) \<in> borel_measurable lborel"
-      using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact)
-    show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \<le> norm (indicator S x * M)"
-      by (auto split: split_indicator simp: abs_real_def dest!: M)
-  qed
-qed simp
-
-lemma borel_integrable_atLeastAtMost:
-  fixes f :: "real \<Rightarrow> real"
-  assumes f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
-  shows "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is "integrable _ ?f")
-proof -
-  have "integrable lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x)"
-  proof (rule borel_integrable_compact)
-    from f show "continuous_on {a..b} f"
-      by (auto intro: continuous_at_imp_continuous_on)
-  qed simp
-  then show ?thesis
-    by (auto simp: mult.commute)
-qed
-
-text \<open>
-
-For the positive integral we replace continuity with Borel-measurability.
-
-\<close>
-
-lemma
-  fixes f :: "real \<Rightarrow> real"
-  assumes [measurable]: "f \<in> borel_measurable borel"
-  assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b"
-  shows nn_integral_FTC_Icc: "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?nn)
-    and has_bochner_integral_FTC_Icc_nonneg:
-      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
-    and integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
-    and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int)
-proof -
-  have *: "(\<lambda>x. f x * indicator {a..b} x) \<in> borel_measurable borel" "\<And>x. 0 \<le> f x * indicator {a..b} x"
-    using f(2) by (auto split: split_indicator)
-
-  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> y \<le> b\<Longrightarrow> F x \<le> F y" for x y
-    using f by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
-
-  have "(f has_integral F b - F a) {a..b}"
-    by (intro fundamental_theorem_of_calculus)
-       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric]
-             intro: has_field_derivative_subset[OF f(1)] \<open>a \<le> b\<close>)
-  then have i: "((\<lambda>x. f x * indicator {a .. b} x) has_integral F b - F a) UNIV"
-    unfolding indicator_def if_distrib[where f="\<lambda>x. a * x" for a]
-    by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
-  then have nn: "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
-    by (rule nn_integral_has_integral_lborel[OF *])
-  then show ?has
-    by (rule has_bochner_integral_nn_integral[rotated 3]) (simp_all add: * F_mono \<open>a \<le> b\<close>)
-  then show ?eq ?int
-    unfolding has_bochner_integral_iff by auto
-  show ?nn
-    by (subst nn[symmetric])
-       (auto intro!: nn_integral_cong simp add: ennreal_mult f split: split_indicator)
-qed
-
-lemma
-  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
-  assumes "a \<le> b"
-  assumes "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
-  assumes cont: "continuous_on {a .. b} f"
-  shows has_bochner_integral_FTC_Icc:
-      "has_bochner_integral lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) (F b - F a)" (is ?has)
-    and integral_FTC_Icc: "(\<integral>x. indicator {a .. b} x *\<^sub>R f x \<partial>lborel) = F b - F a" (is ?eq)
-proof -
-  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
-  have int: "integrable lborel ?f"
-    using borel_integrable_compact[OF _ cont] by auto
-  have "(f has_integral F b - F a) {a..b}"
-    using assms(1,2) by (intro fundamental_theorem_of_calculus) auto
-  moreover
-  have "(f has_integral integral\<^sup>L lborel ?f) {a..b}"
-    using has_integral_integral_lborel[OF int]
-    unfolding indicator_def if_distrib[where f="\<lambda>x. x *\<^sub>R a" for a]
-    by (simp cong del: if_weak_cong del: atLeastAtMost_iff)
-  ultimately show ?eq
-    by (auto dest: has_integral_unique)
-  then show ?has
-    using int by (auto simp: has_bochner_integral_iff)
-qed
-
-lemma
-  fixes f :: "real \<Rightarrow> real"
-  assumes "a \<le> b"
-  assumes deriv: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x"
-  assumes cont: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x"
-  shows has_bochner_integral_FTC_Icc_real:
-      "has_bochner_integral lborel (\<lambda>x. f x * indicator {a .. b} x) (F b - F a)" (is ?has)
-    and integral_FTC_Icc_real: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq)
-proof -
-  have 1: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
-    unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
-    using deriv by (auto intro: DERIV_subset)
-  have 2: "continuous_on {a .. b} f"
-    using cont by (intro continuous_at_imp_continuous_on) auto
-  show ?has ?eq
-    using has_bochner_integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2] integral_FTC_Icc[OF \<open>a \<le> b\<close> 1 2]
-    by (auto simp: mult.commute)
-qed
-
-lemma nn_integral_FTC_atLeast:
-  fixes f :: "real \<Rightarrow> real"
-  assumes f_borel: "f \<in> borel_measurable borel"
-  assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x"
-  assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x"
-  assumes lim: "(F \<longlongrightarrow> T) at_top"
-  shows "(\<integral>\<^sup>+x. ennreal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
-proof -
-  let ?f = "\<lambda>(i::nat) (x::real). ennreal (f x) * indicator {a..a + real i} x"
-  let ?fR = "\<lambda>x. ennreal (f x) * indicator {a ..} x"
-
-  have F_mono: "a \<le> x \<Longrightarrow> x \<le> y \<Longrightarrow> F x \<le> F y" for x y
-    using f nonneg by (intro DERIV_nonneg_imp_nondecreasing[of x y F]) (auto intro: order_trans)
-  then have F_le_T: "a \<le> x \<Longrightarrow> F x \<le> T" for x
-    by (intro tendsto_le_const[OF _ lim])
-       (auto simp: trivial_limit_at_top_linorder eventually_at_top_linorder)
-
-  have "(SUP i::nat. ?f i x) = ?fR x" for x
-  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
-    from reals_Archimedean2[of "x - a"] guess n ..
-    then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially"
-      by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator)
-    then show "(\<lambda>n. ?f n x) \<longlonglongrightarrow> ?fR x"
-      by (rule Lim_eventually)
-  qed (auto simp: nonneg incseq_def le_fun_def split: split_indicator)
-  then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"
-    by simp
-  also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))"
-  proof (rule nn_integral_monotone_convergence_SUP)
-    show "incseq ?f"
-      using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator)
-    show "\<And>i. (?f i) \<in> borel_measurable lborel"
-      using f_borel by auto
-  qed
-  also have "\<dots> = (SUP i::nat. ennreal (F (a + real i) - F a))"
-    by (subst nn_integral_FTC_Icc[OF f_borel f nonneg]) auto
-  also have "\<dots> = T - F a"
-  proof (rule LIMSEQ_unique[OF LIMSEQ_SUP])
-    have "(\<lambda>x. F (a + real x)) \<longlonglongrightarrow> T"
-      apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top])
-      apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])
-      apply (rule filterlim_real_sequentially)
-      done
-    then show "(\<lambda>n. ennreal (F (a + real n) - F a)) \<longlonglongrightarrow> ennreal (T - F a)"
-      by (simp add: F_mono F_le_T tendsto_diff)
-  qed (auto simp: incseq_def intro!: ennreal_le_iff[THEN iffD2] F_mono)
-  finally show ?thesis .
-qed
-
-lemma integral_power:
-  "a \<le> b \<Longrightarrow> (\<integral>x. x^k * indicator {a..b} x \<partial>lborel) = (b^Suc k - a^Suc k) / Suc k"
-proof (subst integral_FTC_Icc_real)
-  fix x show "DERIV (\<lambda>x. x^Suc k / Suc k) x :> x^k"
-    by (intro derivative_eq_intros) auto
-qed (auto simp: field_simps simp del: of_nat_Suc)
-
-subsection \<open>Integration by parts\<close>
-
-lemma integral_by_parts_integrable:
-  fixes f g F G::"real \<Rightarrow> real"
-  assumes "a \<le> b"
-  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
-  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
-  assumes [intro]: "!!x. DERIV F x :> f x"
-  assumes [intro]: "!!x. DERIV G x :> g x"
-  shows  "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)"
-  by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont)
-
-lemma integral_by_parts:
-  fixes f g F G::"real \<Rightarrow> real"
-  assumes [arith]: "a \<le> b"
-  assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
-  assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
-  assumes [intro]: "!!x. DERIV F x :> f x"
-  assumes [intro]: "!!x. DERIV G x :> g x"
-  shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel)
-            =  F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
-proof-
-  have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a"
-    by (rule integral_FTC_Icc_real, auto intro!: derivative_eq_intros continuous_intros)
-      (auto intro!: DERIV_isCont)
-
-  have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) =
-    (\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel"
-    apply (subst integral_add[symmetric])
-    apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros)
-    by (auto intro!: DERIV_isCont integral_cong split:split_indicator)
-
-  thus ?thesis using 0 by auto
-qed
-
-lemma integral_by_parts':
-  fixes f g F G::"real \<Rightarrow> real"
-  assumes "a \<le> b"
-  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x"
-  assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x"
-  assumes "!!x. DERIV F x :> f x"
-  assumes "!!x. DERIV G x :> g x"
-  shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel)
-            =  F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel"
-  using integral_by_parts[OF assms] by (simp add: ac_simps)
-
-lemma has_bochner_integral_even_function:
-  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
-  assumes even: "\<And>x. f (- x) = f x"
-  shows "has_bochner_integral lborel f (2 *\<^sub>R x)"
-proof -
-  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
-    by (auto split: split_indicator)
-  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x"
-    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
-       (auto simp: indicator even f)
-  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)"
-    by (rule has_bochner_integral_add)
-  then have "has_bochner_integral lborel f (x + x)"
-    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
-       (auto split: split_indicator)
-  then show ?thesis
-    by (simp add: scaleR_2)
-qed
-
-lemma has_bochner_integral_odd_function:
-  fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x"
-  assumes odd: "\<And>x. f (- x) = - f x"
-  shows "has_bochner_integral lborel f 0"
-proof -
-  have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x"
-    by (auto split: split_indicator)
-  have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x"
-    by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0])
-       (auto simp: indicator odd f)
-  from has_bochner_integral_minus[OF this]
-  have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)"
-    by simp
-  with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)"
-    by (rule has_bochner_integral_add)
-  then have "has_bochner_integral lborel f (x + - x)"
-    by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4])
-       (auto split: split_indicator)
-  then show ?thesis
-    by simp
-qed
-
-end
-
--- a/src/HOL/Probability/Measurable.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,624 +0,0 @@
-(*  Title:      HOL/Probability/Measurable.thy
-    Author:     Johannes Hölzl <hoelzl@in.tum.de>
-*)
-theory Measurable
-  imports
-    Sigma_Algebra
-    "~~/src/HOL/Library/Order_Continuity"
-begin
-
-subsection \<open>Measurability prover\<close>
-
-lemma (in algebra) sets_Collect_finite_All:
-  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
-  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (if S = {} then \<Omega> else \<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
-    by auto
-  with assms show ?thesis by (auto intro!: sets_Collect_finite_All')
-qed
-
-abbreviation "pred M P \<equiv> P \<in> measurable M (count_space (UNIV::bool set))"
-
-lemma pred_def: "pred M P \<longleftrightarrow> {x\<in>space M. P x} \<in> sets M"
-proof
-  assume "pred M P"
-  then have "P -` {True} \<inter> space M \<in> sets M"
-    by (auto simp: measurable_count_space_eq2)
-  also have "P -` {True} \<inter> space M = {x\<in>space M. P x}" by auto
-  finally show "{x\<in>space M. P x} \<in> sets M" .
-next
-  assume P: "{x\<in>space M. P x} \<in> sets M"
-  moreover
-  { fix X
-    have "X \<in> Pow (UNIV :: bool set)" by simp
-    then have "P -` X \<inter> space M = {x\<in>space M. ((X = {True} \<longrightarrow> P x) \<and> (X = {False} \<longrightarrow> \<not> P x) \<and> X \<noteq> {})}"
-      unfolding UNIV_bool Pow_insert Pow_empty by auto
-    then have "P -` X \<inter> space M \<in> sets M"
-      by (auto intro!: sets.sets_Collect_neg sets.sets_Collect_imp sets.sets_Collect_conj sets.sets_Collect_const P) }
-  then show "pred M P"
-    by (auto simp: measurable_def)
-qed
-
-lemma pred_sets1: "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> f \<in> measurable N M \<Longrightarrow> pred N (\<lambda>x. P (f x))"
-  by (rule measurable_compose[where f=f and N=M]) (auto simp: pred_def)
-
-lemma pred_sets2: "A \<in> sets N \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A)"
-  by (rule measurable_compose[where f=f and N=N]) (auto simp: pred_def Int_def[symmetric])
-
-ML_file "measurable.ML"
-
-attribute_setup measurable = \<open>
-  Scan.lift (
-    (Args.add >> K true || Args.del >> K false || Scan.succeed true) --
-    Scan.optional (Args.parens (
-      Scan.optional (Args.$$$ "raw" >> K true) false --
-      Scan.optional (Args.$$$ "generic" >> K Measurable.Generic) Measurable.Concrete))
-    (false, Measurable.Concrete) >>
-    Measurable.measurable_thm_attr)
-\<close> "declaration of measurability theorems"
-
-attribute_setup measurable_dest = Measurable.dest_thm_attr
-  "add dest rule to measurability prover"
-
-attribute_setup measurable_cong = Measurable.cong_thm_attr
-  "add congurence rules to measurability prover"
-
-method_setup measurable = \<open> Scan.lift (Scan.succeed (METHOD o Measurable.measurable_tac)) \<close>
-  "measurability prover"
-
-simproc_setup measurable ("A \<in> sets M" | "f \<in> measurable M N") = \<open>K Measurable.simproc\<close>
-
-setup \<open>
-  Global_Theory.add_thms_dynamic (@{binding measurable}, Measurable.get_all)
-\<close>
-
-declare
-  pred_sets1[measurable_dest]
-  pred_sets2[measurable_dest]
-  sets.sets_into_space[measurable_dest]
-
-declare
-  sets.top[measurable]
-  sets.empty_sets[measurable (raw)]
-  sets.Un[measurable (raw)]
-  sets.Diff[measurable (raw)]
-
-declare
-  measurable_count_space[measurable (raw)]
-  measurable_ident[measurable (raw)]
-  measurable_id[measurable (raw)]
-  measurable_const[measurable (raw)]
-  measurable_If[measurable (raw)]
-  measurable_comp[measurable (raw)]
-  measurable_sets[measurable (raw)]
-
-declare measurable_cong_sets[measurable_cong]
-declare sets_restrict_space_cong[measurable_cong]
-declare sets_restrict_UNIV[measurable_cong]
-
-lemma predE[measurable (raw)]:
-  "pred M P \<Longrightarrow> {x\<in>space M. P x} \<in> sets M"
-  unfolding pred_def .
-
-lemma pred_intros_imp'[measurable (raw)]:
-  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<longrightarrow> P x)"
-  by (cases K) auto
-
-lemma pred_intros_conj1'[measurable (raw)]:
-  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<and> P x)"
-  by (cases K) auto
-
-lemma pred_intros_conj2'[measurable (raw)]:
-  "(K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<and> K)"
-  by (cases K) auto
-
-lemma pred_intros_disj1'[measurable (raw)]:
-  "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. K \<or> P x)"
-  by (cases K) auto
-
-lemma pred_intros_disj2'[measurable (raw)]:
-  "(\<not> K \<Longrightarrow> pred M (\<lambda>x. P x)) \<Longrightarrow> pred M (\<lambda>x. P x \<or> K)"
-  by (cases K) auto
-
-lemma pred_intros_logic[measurable (raw)]:
-  "pred M (\<lambda>x. x \<in> space M)"
-  "pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. \<not> P x)"
-  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<and> P x)"
-  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<longrightarrow> P x)"
-  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x \<or> P x)"
-  "pred M (\<lambda>x. Q x) \<Longrightarrow> pred M (\<lambda>x. P x) \<Longrightarrow> pred M (\<lambda>x. Q x = P x)"
-  "pred M (\<lambda>x. f x \<in> UNIV)"
-  "pred M (\<lambda>x. f x \<in> {})"
-  "pred M (\<lambda>x. P' (f x) x) \<Longrightarrow> pred M (\<lambda>x. f x \<in> {y. P' y x})"
-  "pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> - (B x))"
-  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) - (B x))"
-  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<inter> (B x))"
-  "pred M (\<lambda>x. f x \<in> (A x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (B x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (A x) \<union> (B x))"
-  "pred M (\<lambda>x. g x (f x) \<in> (X x)) \<Longrightarrow> pred M (\<lambda>x. f x \<in> (g x) -` (X x))"
-  by (auto simp: iff_conv_conj_imp pred_def)
-
-lemma pred_intros_countable[measurable (raw)]:
-  fixes P :: "'a \<Rightarrow> 'i :: countable \<Rightarrow> bool"
-  shows
-    "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i. P x i)"
-    "(\<And>i. pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i. P x i)"
-  by (auto intro!: sets.sets_Collect_countable_All sets.sets_Collect_countable_Ex simp: pred_def)
-
-lemma pred_intros_countable_bounded[measurable (raw)]:
-  fixes X :: "'i :: countable set"
-  shows
-    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>X. N x i))"
-    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>X. N x i))"
-    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
-    "(\<And>i. i \<in> X \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
-  by simp_all (auto simp: Bex_def Ball_def)
-
-lemma pred_intros_finite[measurable (raw)]:
-  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Inter>i\<in>I. N x i))"
-  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. x \<in> N x i)) \<Longrightarrow> pred M (\<lambda>x. x \<in> (\<Union>i\<in>I. N x i))"
-  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<forall>i\<in>I. P x i)"
-  "finite I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> pred M (\<lambda>x. P x i)) \<Longrightarrow> pred M (\<lambda>x. \<exists>i\<in>I. P x i)"
-  by (auto intro!: sets.sets_Collect_finite_Ex sets.sets_Collect_finite_All simp: iff_conv_conj_imp pred_def)
-
-lemma countable_Un_Int[measurable (raw)]:
-  "(\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Union>i\<in>I. N i) \<in> sets M"
-  "I \<noteq> {} \<Longrightarrow> (\<And>i :: 'i :: countable. i \<in> I \<Longrightarrow> N i \<in> sets M) \<Longrightarrow> (\<Inter>i\<in>I. N i) \<in> sets M"
-  by auto
-
-declare
-  finite_UN[measurable (raw)]
-  finite_INT[measurable (raw)]
-
-lemma sets_Int_pred[measurable (raw)]:
-  assumes space: "A \<inter> B \<subseteq> space M" and [measurable]: "pred M (\<lambda>x. x \<in> A)" "pred M (\<lambda>x. x \<in> B)"
-  shows "A \<inter> B \<in> sets M"
-proof -
-  have "{x\<in>space M. x \<in> A \<inter> B} \<in> sets M" by auto
-  also have "{x\<in>space M. x \<in> A \<inter> B} = A \<inter> B"
-    using space by auto
-  finally show ?thesis .
-qed
-
-lemma [measurable (raw generic)]:
-  assumes f: "f \<in> measurable M N" and c: "c \<in> space N \<Longrightarrow> {c} \<in> sets N"
-  shows pred_eq_const1: "pred M (\<lambda>x. f x = c)"
-    and pred_eq_const2: "pred M (\<lambda>x. c = f x)"
-proof -
-  show "pred M (\<lambda>x. f x = c)"
-  proof cases
-    assume "c \<in> space N"
-    with measurable_sets[OF f c] show ?thesis
-      by (auto simp: Int_def conj_commute pred_def)
-  next
-    assume "c \<notin> space N"
-    with f[THEN measurable_space] have "{x \<in> space M. f x = c} = {}" by auto
-    then show ?thesis by (auto simp: pred_def cong: conj_cong)
-  qed
-  then show "pred M (\<lambda>x. c = f x)"
-    by (simp add: eq_commute)
-qed
-
-lemma pred_count_space_const1[measurable (raw)]:
-  "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. f x = c)"
-  by (intro pred_eq_const1[where N="count_space UNIV"]) (auto )
-
-lemma pred_count_space_const2[measurable (raw)]:
-  "f \<in> measurable M (count_space UNIV) \<Longrightarrow> Measurable.pred M (\<lambda>x. c = f x)"
-  by (intro pred_eq_const2[where N="count_space UNIV"]) (auto )
-
-lemma pred_le_const[measurable (raw generic)]:
-  assumes f: "f \<in> measurable M N" and c: "{.. c} \<in> sets N" shows "pred M (\<lambda>x. f x \<le> c)"
-  using measurable_sets[OF f c]
-  by (auto simp: Int_def conj_commute eq_commute pred_def)
-
-lemma pred_const_le[measurable (raw generic)]:
-  assumes f: "f \<in> measurable M N" and c: "{c ..} \<in> sets N" shows "pred M (\<lambda>x. c \<le> f x)"
-  using measurable_sets[OF f c]
-  by (auto simp: Int_def conj_commute eq_commute pred_def)
-
-lemma pred_less_const[measurable (raw generic)]:
-  assumes f: "f \<in> measurable M N" and c: "{..< c} \<in> sets N" shows "pred M (\<lambda>x. f x < c)"
-  using measurable_sets[OF f c]
-  by (auto simp: Int_def conj_commute eq_commute pred_def)
-
-lemma pred_const_less[measurable (raw generic)]:
-  assumes f: "f \<in> measurable M N" and c: "{c <..} \<in> sets N" shows "pred M (\<lambda>x. c < f x)"
-  using measurable_sets[OF f c]
-  by (auto simp: Int_def conj_commute eq_commute pred_def)
-
-declare
-  sets.Int[measurable (raw)]
-
-lemma pred_in_If[measurable (raw)]:
-  "(P \<Longrightarrow> pred M (\<lambda>x. x \<in> A x)) \<Longrightarrow> (\<not> P \<Longrightarrow> pred M (\<lambda>x. x \<in> B x)) \<Longrightarrow>
-    pred M (\<lambda>x. x \<in> (if P then A x else B x))"
-  by auto
-
-lemma sets_range[measurable_dest]:
-  "A ` I \<subseteq> sets M \<Longrightarrow> i \<in> I \<Longrightarrow> A i \<in> sets M"
-  by auto
-
-lemma pred_sets_range[measurable_dest]:
-  "A ` I \<subseteq> sets N \<Longrightarrow> i \<in> I \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
-  using pred_sets2[OF sets_range] by auto
-
-lemma sets_All[measurable_dest]:
-  "\<forall>i. A i \<in> sets (M i) \<Longrightarrow> A i \<in> sets (M i)"
-  by auto
-
-lemma pred_sets_All[measurable_dest]:
-  "\<forall>i. A i \<in> sets (N i) \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
-  using pred_sets2[OF sets_All, of A N f] by auto
-
-lemma sets_Ball[measurable_dest]:
-  "\<forall>i\<in>I. A i \<in> sets (M i) \<Longrightarrow> i\<in>I \<Longrightarrow> A i \<in> sets (M i)"
-  by auto
-
-lemma pred_sets_Ball[measurable_dest]:
-  "\<forall>i\<in>I. A i \<in> sets (N i) \<Longrightarrow> i\<in>I \<Longrightarrow> f \<in> measurable M (N i) \<Longrightarrow> pred M (\<lambda>x. f x \<in> A i)"
-  using pred_sets2[OF sets_Ball, of _ _ _ f] by auto
-
-lemma measurable_finite[measurable (raw)]:
-  fixes S :: "'a \<Rightarrow> nat set"
-  assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
-  shows "pred M (\<lambda>x. finite (S x))"
-  unfolding finite_nat_set_iff_bounded by (simp add: Ball_def)
-
-lemma measurable_Least[measurable]:
-  assumes [measurable]: "(\<And>i::nat. (\<lambda>x. P i x) \<in> measurable M (count_space UNIV))"
-  shows "(\<lambda>x. LEAST i. P i x) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_def by (safe intro!: sets_Least) simp_all
-
-lemma measurable_Max_nat[measurable (raw)]:
-  fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
-  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
-  shows "(\<lambda>x. Max {i. P i x}) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_count_space_eq2_countable
-proof safe
-  fix n
-
-  { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
-    then have "infinite {i. P i x}"
-      unfolding infinite_nat_iff_unbounded_le by auto
-    then have "Max {i. P i x} = the None"
-      by (rule Max.infinite) }
-  note 1 = this
-
-  { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
-    then have "finite {i. P i x}"
-      by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded)
-    with \<open>P i x\<close> have "P (Max {i. P i x}) x" "i \<le> Max {i. P i x}" "finite {i. P i x}"
-      using Max_in[of "{i. P i x}"] by auto }
-  note 2 = this
-
-  have "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Max {i. P i x} = n}"
-    by auto
-  also have "\<dots> =
-    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
-      if (\<exists>i. P i x) then P n x \<and> (\<forall>i>n. \<not> P i x)
-      else Max {} = n}"
-    by (intro arg_cong[where f=Collect] ext conj_cong)
-       (auto simp add: 1 2 not_le[symmetric] intro!: Max_eqI)
-  also have "\<dots> \<in> sets M"
-    by measurable
-  finally show "(\<lambda>x. Max {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
-qed simp
-
-lemma measurable_Min_nat[measurable (raw)]:
-  fixes P :: "nat \<Rightarrow> 'a \<Rightarrow> bool"
-  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
-  shows "(\<lambda>x. Min {i. P i x}) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_count_space_eq2_countable
-proof safe
-  fix n
-
-  { fix x assume "\<forall>i. \<exists>n\<ge>i. P n x"
-    then have "infinite {i. P i x}"
-      unfolding infinite_nat_iff_unbounded_le by auto
-    then have "Min {i. P i x} = the None"
-      by (rule Min.infinite) }
-  note 1 = this
-
-  { fix x i j assume "P i x" "\<forall>n\<ge>j. \<not> P n x"
-    then have "finite {i. P i x}"
-      by (auto simp: subset_eq not_le[symmetric] finite_nat_iff_bounded)
-    with \<open>P i x\<close> have "P (Min {i. P i x}) x" "Min {i. P i x} \<le> i" "finite {i. P i x}"
-      using Min_in[of "{i. P i x}"] by auto }
-  note 2 = this
-
-  have "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M = {x\<in>space M. Min {i. P i x} = n}"
-    by auto
-  also have "\<dots> =
-    {x\<in>space M. if (\<forall>i. \<exists>n\<ge>i. P n x) then the None = n else
-      if (\<exists>i. P i x) then P n x \<and> (\<forall>i<n. \<not> P i x)
-      else Min {} = n}"
-    by (intro arg_cong[where f=Collect] ext conj_cong)
-       (auto simp add: 1 2 not_le[symmetric] intro!: Min_eqI)
-  also have "\<dots> \<in> sets M"
-    by measurable
-  finally show "(\<lambda>x. Min {i. P i x}) -` {n} \<inter> space M \<in> sets M" .
-qed simp
-
-lemma measurable_count_space_insert[measurable (raw)]:
-  "s \<in> S \<Longrightarrow> A \<in> sets (count_space S) \<Longrightarrow> insert s A \<in> sets (count_space S)"
-  by simp
-
-lemma sets_UNIV [measurable (raw)]: "A \<in> sets (count_space UNIV)"
-  by simp
-
-lemma measurable_card[measurable]:
-  fixes S :: "'a \<Rightarrow> nat set"
-  assumes [measurable]: "\<And>i. {x\<in>space M. i \<in> S x} \<in> sets M"
-  shows "(\<lambda>x. card (S x)) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_count_space_eq2_countable
-proof safe
-  fix n show "(\<lambda>x. card (S x)) -` {n} \<inter> space M \<in> sets M"
-  proof (cases n)
-    case 0
-    then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M = {x\<in>space M. infinite (S x) \<or> (\<forall>i. i \<notin> S x)}"
-      by auto
-    also have "\<dots> \<in> sets M"
-      by measurable
-    finally show ?thesis .
-  next
-    case (Suc i)
-    then have "(\<lambda>x. card (S x)) -` {n} \<inter> space M =
-      (\<Union>F\<in>{A\<in>{A. finite A}. card A = n}. {x\<in>space M. (\<forall>i. i \<in> S x \<longleftrightarrow> i \<in> F)})"
-      unfolding set_eq_iff[symmetric] Collect_bex_eq[symmetric] by (auto intro: card_ge_0_finite)
-    also have "\<dots> \<in> sets M"
-      by (intro sets.countable_UN' countable_Collect countable_Collect_finite) auto
-    finally show ?thesis .
-  qed
-qed rule
-
-lemma measurable_pred_countable[measurable (raw)]:
-  assumes "countable X"
-  shows
-    "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<forall>i\<in>X. P x i)"
-    "(\<And>i. i \<in> X \<Longrightarrow> Measurable.pred M (\<lambda>x. P x i)) \<Longrightarrow> Measurable.pred M (\<lambda>x. \<exists>i\<in>X. P x i)"
-  unfolding pred_def
-  by (auto intro!: sets.sets_Collect_countable_All' sets.sets_Collect_countable_Ex' assms)
-
-subsection \<open>Measurability for (co)inductive predicates\<close>
-
-lemma measurable_bot[measurable]: "bot \<in> measurable M (count_space UNIV)"
-  by (simp add: bot_fun_def)
-
-lemma measurable_top[measurable]: "top \<in> measurable M (count_space UNIV)"
-  by (simp add: top_fun_def)
-
-lemma measurable_SUP[measurable]:
-  fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
-  assumes [simp]: "countable I"
-  assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)"
-  shows "(\<lambda>x. SUP i:I. F i x) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_count_space_eq2_countable
-proof (safe intro!: UNIV_I)
-  fix a
-  have "(\<lambda>x. SUP i:I. F i x) -` {a} \<inter> space M =
-    {x\<in>space M. (\<forall>i\<in>I. F i x \<le> a) \<and> (\<forall>b. (\<forall>i\<in>I. F i x \<le> b) \<longrightarrow> a \<le> b)}"
-    unfolding SUP_le_iff[symmetric] by auto
-  also have "\<dots> \<in> sets M"
-    by measurable
-  finally show "(\<lambda>x. SUP i:I. F i x) -` {a} \<inter> space M \<in> sets M" .
-qed
-
-lemma measurable_INF[measurable]:
-  fixes F :: "'i \<Rightarrow> 'a \<Rightarrow> 'b::{complete_lattice, countable}"
-  assumes [simp]: "countable I"
-  assumes [measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> measurable M (count_space UNIV)"
-  shows "(\<lambda>x. INF i:I. F i x) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_count_space_eq2_countable
-proof (safe intro!: UNIV_I)
-  fix a
-  have "(\<lambda>x. INF i:I. F i x) -` {a} \<inter> space M =
-    {x\<in>space M. (\<forall>i\<in>I. a \<le> F i x) \<and> (\<forall>b. (\<forall>i\<in>I. b \<le> F i x) \<longrightarrow> b \<le> a)}"
-    unfolding le_INF_iff[symmetric] by auto
-  also have "\<dots> \<in> sets M"
-    by measurable
-  finally show "(\<lambda>x. INF i:I. F i x) -` {a} \<inter> space M \<in> sets M" .
-qed
-
-lemma measurable_lfp_coinduct[consumes 1, case_names continuity step]:
-  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
-  assumes "P M"
-  assumes F: "sup_continuous F"
-  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
-  shows "lfp F \<in> measurable M (count_space UNIV)"
-proof -
-  { fix i from \<open>P M\<close> have "((F ^^ i) bot) \<in> measurable M (count_space UNIV)"
-      by (induct i arbitrary: M) (auto intro!: *) }
-  then have "(\<lambda>x. SUP i. (F ^^ i) bot x) \<in> measurable M (count_space UNIV)"
-    by measurable
-  also have "(\<lambda>x. SUP i. (F ^^ i) bot x) = lfp F"
-    by (subst sup_continuous_lfp) (auto intro: F)
-  finally show ?thesis .
-qed
-
-lemma measurable_lfp:
-  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
-  assumes F: "sup_continuous F"
-  assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
-  shows "lfp F \<in> measurable M (count_space UNIV)"
-  by (coinduction rule: measurable_lfp_coinduct[OF _ F]) (blast intro: *)
-
-lemma measurable_gfp_coinduct[consumes 1, case_names continuity step]:
-  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
-  assumes "P M"
-  assumes F: "inf_continuous F"
-  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> A \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
-  shows "gfp F \<in> measurable M (count_space UNIV)"
-proof -
-  { fix i from \<open>P M\<close> have "((F ^^ i) top) \<in> measurable M (count_space UNIV)"
-      by (induct i arbitrary: M) (auto intro!: *) }
-  then have "(\<lambda>x. INF i. (F ^^ i) top x) \<in> measurable M (count_space UNIV)"
-    by measurable
-  also have "(\<lambda>x. INF i. (F ^^ i) top x) = gfp F"
-    by (subst inf_continuous_gfp) (auto intro: F)
-  finally show ?thesis .
-qed
-
-lemma measurable_gfp:
-  fixes F :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'b::{complete_lattice, countable})"
-  assumes F: "inf_continuous F"
-  assumes *: "\<And>A. A \<in> measurable M (count_space UNIV) \<Longrightarrow> F A \<in> measurable M (count_space UNIV)"
-  shows "gfp F \<in> measurable M (count_space UNIV)"
-  by (coinduction rule: measurable_gfp_coinduct[OF _ F]) (blast intro: *)
-
-lemma measurable_lfp2_coinduct[consumes 1, case_names continuity step]:
-  fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
-  assumes "P M s"
-  assumes F: "sup_continuous F"
-  assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
-  shows "lfp F s \<in> measurable M (count_space UNIV)"
-proof -
-  { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
-      by (induct i arbitrary: M s) (auto intro!: *) }
-  then have "(\<lambda>x. SUP i. (F ^^ i) bot s x) \<in> measurable M (count_space UNIV)"
-    by measurable
-  also have "(\<lambda>x. SUP i. (F ^^ i) bot s x) = lfp F s"
-    by (subst sup_continuous_lfp) (auto simp: F)
-  finally show ?thesis .
-qed
-
-lemma measurable_gfp2_coinduct[consumes 1, case_names continuity step]:
-  fixes F :: "('a \<Rightarrow> 'c \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> 'c \<Rightarrow> 'b::{complete_lattice, countable})"
-  assumes "P M s"
-  assumes F: "inf_continuous F"
-  assumes *: "\<And>M A s. P M s \<Longrightarrow> (\<And>N t. P N t \<Longrightarrow> A t \<in> measurable N (count_space UNIV)) \<Longrightarrow> F A s \<in> measurable M (count_space UNIV)"
-  shows "gfp F s \<in> measurable M (count_space UNIV)"
-proof -
-  { fix i from \<open>P M s\<close> have "(\<lambda>x. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
-      by (induct i arbitrary: M s) (auto intro!: *) }
-  then have "(\<lambda>x. INF i. (F ^^ i) top s x) \<in> measurable M (count_space UNIV)"
-    by measurable
-  also have "(\<lambda>x. INF i. (F ^^ i) top s x) = gfp F s"
-    by (subst inf_continuous_gfp) (auto simp: F)
-  finally show ?thesis .
-qed
-
-lemma measurable_enat_coinduct:
-  fixes f :: "'a \<Rightarrow> enat"
-  assumes "R f"
-  assumes *: "\<And>f. R f \<Longrightarrow> \<exists>g h i P. R g \<and> f = (\<lambda>x. if P x then h x else eSuc (g (i x))) \<and>
-    Measurable.pred M P \<and>
-    i \<in> measurable M M \<and>
-    h \<in> measurable M (count_space UNIV)"
-  shows "f \<in> measurable M (count_space UNIV)"
-proof (simp add: measurable_count_space_eq2_countable, rule )
-  fix a :: enat
-  have "f -` {a} \<inter> space M = {x\<in>space M. f x = a}"
-    by auto
-  { fix i :: nat
-    from \<open>R f\<close> have "Measurable.pred M (\<lambda>x. f x = enat i)"
-    proof (induction i arbitrary: f)
-      case 0
-      from *[OF this] obtain g h i P
-        where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and
-          [measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)"
-        by auto
-      have "Measurable.pred M (\<lambda>x. P x \<and> h x = 0)"
-        by measurable
-      also have "(\<lambda>x. P x \<and> h x = 0) = (\<lambda>x. f x = enat 0)"
-        by (auto simp: f zero_enat_def[symmetric])
-      finally show ?case .
-    next
-      case (Suc n)
-      from *[OF Suc.prems] obtain g h i P
-        where f: "f = (\<lambda>x. if P x then h x else eSuc (g (i x)))" and "R g" and
-          M[measurable]: "Measurable.pred M P" "i \<in> measurable M M" "h \<in> measurable M (count_space UNIV)"
-        by auto
-      have "(\<lambda>x. f x = enat (Suc n)) =
-        (\<lambda>x. (P x \<longrightarrow> h x = enat (Suc n)) \<and> (\<not> P x \<longrightarrow> g (i x) = enat n))"
-        by (auto simp: f zero_enat_def[symmetric] eSuc_enat[symmetric])
-      also have "Measurable.pred M \<dots>"
-        by (intro pred_intros_logic measurable_compose[OF M(2)] Suc \<open>R g\<close>) measurable
-      finally show ?case .
-    qed
-    then have "f -` {enat i} \<inter> space M \<in> sets M"
-      by (simp add: pred_def Int_def conj_commute) }
-  note fin = this
-  show "f -` {a} \<inter> space M \<in> sets M"
-  proof (cases a)
-    case infinity
-    then have "f -` {a} \<inter> space M = space M - (\<Union>n. f -` {enat n} \<inter> space M)"
-      by auto
-    also have "\<dots> \<in> sets M"
-      by (intro sets.Diff sets.top sets.Un sets.countable_UN) (auto intro!: fin)
-    finally show ?thesis .
-  qed (simp add: fin)
-qed
-
-lemma measurable_THE:
-  fixes P :: "'a \<Rightarrow> 'b \<Rightarrow> bool"
-  assumes [measurable]: "\<And>i. Measurable.pred M (P i)"
-  assumes I[simp]: "countable I" "\<And>i x. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> i \<in> I"
-  assumes unique: "\<And>x i j. x \<in> space M \<Longrightarrow> P i x \<Longrightarrow> P j x \<Longrightarrow> i = j"
-  shows "(\<lambda>x. THE i. P i x) \<in> measurable M (count_space UNIV)"
-  unfolding measurable_def
-proof safe
-  fix X
-  define f where "f x = (THE i. P i x)" for x
-  define undef where "undef = (THE i::'a. False)"
-  { fix i x assume "x \<in> space M" "P i x" then have "f x = i"
-      unfolding f_def using unique by auto }
-  note f_eq = this
-  { fix x assume "x \<in> space M" "\<forall>i\<in>I. \<not> P i x"
-    then have "\<And>i. \<not> P i x"
-      using I(2)[of x] by auto
-    then have "f x = undef"
-      by (auto simp: undef_def f_def) }
-  then have "f -` X \<inter> space M = (\<Union>i\<in>I \<inter> X. {x\<in>space M. P i x}) \<union>
-     (if undef \<in> X then space M - (\<Union>i\<in>I. {x\<in>space M. P i x}) else {})"
-    by (auto dest: f_eq)
-  also have "\<dots> \<in> sets M"
-    by (auto intro!: sets.Diff sets.countable_UN')
-  finally show "f -` X \<inter> space M \<in> sets M" .
-qed simp
-
-lemma measurable_Ex1[measurable (raw)]:
-  assumes [simp]: "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> Measurable.pred M (P i)"
-  shows "Measurable.pred M (\<lambda>x. \<exists>!i\<in>I. P i x)"
-  unfolding bex1_def by measurable
-
-lemma measurable_Sup_nat[measurable (raw)]:
-  fixes F :: "'a \<Rightarrow> nat set"
-  assumes [measurable]: "\<And>i. Measurable.pred M (\<lambda>x. i \<in> F x)"
-  shows "(\<lambda>x. Sup (F x)) \<in> M \<rightarrow>\<^sub>M count_space UNIV"
-proof (clarsimp simp add: measurable_count_space_eq2_countable)
-  fix a
-  have F_empty_iff: "F x = {} \<longleftrightarrow> (\<forall>i. i \<notin> F x)" for x
-    by auto
-  have "Measurable.pred M (\<lambda>x. if finite (F x) then if F x = {} then a = Max {}
-    else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None)"
-    unfolding finite_nat_set_iff_bounded Ball_def F_empty_iff by measurable
-  moreover have "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M =
-    {x\<in>space M. if finite (F x) then if F x = {} then a = Max {}
-      else a \<in> F x \<and> (\<forall>j. j \<in> F x \<longrightarrow> j \<le> a) else a = the None}"
-    by (intro set_eqI)
-       (auto simp: Sup_nat_def Max.infinite intro!: Max_in Max_eqI)
-  ultimately show "(\<lambda>x. Sup (F x)) -` {a} \<inter> space M \<in> sets M"
-    by auto
-qed
-
-lemma measurable_if_split[measurable (raw)]:
-  "(c \<Longrightarrow> Measurable.pred M f) \<Longrightarrow> (\<not> c \<Longrightarrow> Measurable.pred M g) \<Longrightarrow>
-   Measurable.pred M (if c then f else g)"
-  by simp
-
-lemma pred_restrict_space:
-  assumes "S \<in> sets M"
-  shows "Measurable.pred (restrict_space M S) P \<longleftrightarrow> Measurable.pred M (\<lambda>x. x \<in> S \<and> P x)"
-  unfolding pred_def sets_Collect_restrict_space_iff[OF assms] ..
-
-lemma measurable_predpow[measurable]:
-  assumes "Measurable.pred M T"
-  assumes "\<And>Q. Measurable.pred M Q \<Longrightarrow> Measurable.pred M (R Q)"
-  shows "Measurable.pred M ((R ^^ n) T)"
-  by (induct n) (auto intro: assms)
-
-hide_const (open) pred
-
-end
-
--- a/src/HOL/Probability/Measure_Space.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3283 +0,0 @@
-(*  Title:      HOL/Probability/Measure_Space.thy
-    Author:     Lawrence C Paulson
-    Author:     Johannes Hölzl, TU München
-    Author:     Armin Heller, TU München
-*)
-
-section \<open>Measure spaces and their properties\<close>
-
-theory Measure_Space
-imports
-  Measurable "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
-begin
-
-subsection "Relate extended reals and the indicator function"
-
-lemma suminf_cmult_indicator:
-  fixes f :: "nat \<Rightarrow> ennreal"
-  assumes "disjoint_family A" "x \<in> A i"
-  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
-proof -
-  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ennreal)"
-    using \<open>x \<in> A i\<close> assms unfolding disjoint_family_on_def indicator_def by auto
-  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ennreal)"
-    by (auto simp: setsum.If_cases)
-  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ennreal)"
-  proof (rule SUP_eqI)
-    fix y :: ennreal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
-    from this[of "Suc i"] show "f i \<le> y" by auto
-  qed (insert assms, simp)
-  ultimately show ?thesis using assms
-    by (subst suminf_eq_SUP) (auto simp: indicator_def)
-qed
-
-lemma suminf_indicator:
-  assumes "disjoint_family A"
-  shows "(\<Sum>n. indicator (A n) x :: ennreal) = indicator (\<Union>i. A i) x"
-proof cases
-  assume *: "x \<in> (\<Union>i. A i)"
-  then obtain i where "x \<in> A i" by auto
-  from suminf_cmult_indicator[OF assms(1), OF \<open>x \<in> A i\<close>, of "\<lambda>k. 1"]
-  show ?thesis using * by simp
-qed simp
-
-lemma setsum_indicator_disjoint_family:
-  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
-  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
-  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
-proof -
-  have "P \<inter> {i. x \<in> A i} = {j}"
-    using d \<open>x \<in> A j\<close> \<open>j \<in> P\<close> unfolding disjoint_family_on_def
-    by auto
-  thus ?thesis
-    unfolding indicator_def
-    by (simp add: if_distrib setsum.If_cases[OF \<open>finite P\<close>])
-qed
-
-text \<open>
-  The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
-  represent sigma algebras (with an arbitrary emeasure).
-\<close>
-
-subsection "Extend binary sets"
-
-lemma LIMSEQ_binaryset:
-  assumes f: "f {} = 0"
-  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
-proof -
-  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
-    proof
-      fix n
-      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
-        by (induct n)  (auto simp add: binaryset_def f)
-    qed
-  moreover
-  have "... \<longlonglongrightarrow> f A + f B" by (rule tendsto_const)
-  ultimately
-  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
-    by metis
-  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) \<longlonglongrightarrow> f A + f B"
-    by simp
-  thus ?thesis by (rule LIMSEQ_offset [where k=2])
-qed
-
-lemma binaryset_sums:
-  assumes f: "f {} = 0"
-  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
-    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
-
-lemma suminf_binaryset_eq:
-  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
-  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
-  by (metis binaryset_sums sums_unique)
-
-subsection \<open>Properties of a premeasure @{term \<mu>}\<close>
-
-text \<open>
-  The definitions for @{const positive} and @{const countably_additive} should be here, by they are
-  necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
-\<close>
-
-definition subadditive where
-  "subadditive M f \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
-
-lemma subadditiveD: "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
-  by (auto simp add: subadditive_def)
-
-definition countably_subadditive where
-  "countably_subadditive M f \<longleftrightarrow>
-    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
-
-lemma (in ring_of_sets) countably_subadditive_subadditive:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" and cs: "countably_subadditive M f"
-  shows  "subadditive M f"
-proof (auto simp add: subadditive_def)
-  fix x y
-  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
-  hence "disjoint_family (binaryset x y)"
-    by (auto simp add: disjoint_family_on_def binaryset_def)
-  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
-         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
-         f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
-    using cs by (auto simp add: countably_subadditive_def)
-  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
-         f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
-    by (simp add: range_binaryset_eq UN_binaryset_eq)
-  thus "f (x \<union> y) \<le>  f x + f y" using f x y
-    by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
-qed
-
-definition additive where
-  "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
-
-definition increasing where
-  "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
-
-lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
-
-lemma positiveD_empty:
-  "positive M f \<Longrightarrow> f {} = 0"
-  by (auto simp add: positive_def)
-
-lemma additiveD:
-  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
-  by (auto simp add: additive_def)
-
-lemma increasingD:
-  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
-  by (auto simp add: increasing_def)
-
-lemma countably_additiveI[case_names countably]:
-  "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
-  \<Longrightarrow> countably_additive M f"
-  by (simp add: countably_additive_def)
-
-lemma (in ring_of_sets) disjointed_additive:
-  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
-  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
-proof (induct n)
-  case (Suc n)
-  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
-    by simp
-  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
-    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_mono)
-  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
-    using \<open>incseq A\<close> by (auto dest: incseq_SucD simp: disjointed_mono)
-  finally show ?case .
-qed simp
-
-lemma (in ring_of_sets) additive_sum:
-  fixes A:: "'i \<Rightarrow> 'a set"
-  assumes f: "positive M f" and ad: "additive M f" and "finite S"
-      and A: "A`S \<subseteq> M"
-      and disj: "disjoint_family_on A S"
-  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
-  using \<open>finite S\<close> disj A
-proof induct
-  case empty show ?case using f by (simp add: positive_def)
-next
-  case (insert s S)
-  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
-    by (auto simp add: disjoint_family_on_def neq_iff)
-  moreover
-  have "A s \<in> M" using insert by blast
-  moreover have "(\<Union>i\<in>S. A i) \<in> M"
-    using insert \<open>finite S\<close> by auto
-  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
-    using ad UNION_in_sets A by (auto simp add: additive_def)
-  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
-    by (auto simp add: additive_def subset_insertI)
-qed
-
-lemma (in ring_of_sets) additive_increasing:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes posf: "positive M f" and addf: "additive M f"
-  shows "increasing M f"
-proof (auto simp add: increasing_def)
-  fix x y
-  assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
-  then have "y - x \<in> M" by auto
-  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono zero_le)
-  also have "... = f (x \<union> (y-x))" using addf
-    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
-  also have "... = f y"
-    by (metis Un_Diff_cancel Un_absorb1 xy(3))
-  finally show "f x \<le> f y" by simp
-qed
-
-lemma (in ring_of_sets) subadditive:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" "additive M f" and A: "A`S \<subseteq> M" and S: "finite S"
-  shows "f (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. f (A i))"
-using S A
-proof (induct S)
-  case empty thus ?case using f by (auto simp: positive_def)
-next
-  case (insert x F)
-  hence in_M: "A x \<in> M" "(\<Union>i\<in>F. A i) \<in> M" "(\<Union>i\<in>F. A i) - A x \<in> M" using A by force+
-  have subs: "(\<Union>i\<in>F. A i) - A x \<subseteq> (\<Union>i\<in>F. A i)" by auto
-  have "(\<Union>i\<in>(insert x F). A i) = A x \<union> ((\<Union>i\<in>F. A i) - A x)" by auto
-  hence "f (\<Union>i\<in>(insert x F). A i) = f (A x \<union> ((\<Union>i\<in>F. A i) - A x))"
-    by simp
-  also have "\<dots> = f (A x) + f ((\<Union>i\<in>F. A i) - A x)"
-    using f(2) by (rule additiveD) (insert in_M, auto)
-  also have "\<dots> \<le> f (A x) + f (\<Union>i\<in>F. A i)"
-    using additive_increasing[OF f] in_M subs by (auto simp: increasing_def intro: add_left_mono)
-  also have "\<dots> \<le> f (A x) + (\<Sum>i\<in>F. f (A i))" using insert by (auto intro: add_left_mono)
-  finally show "f (\<Union>i\<in>(insert x F). A i) \<le> (\<Sum>i\<in>(insert x F). f (A i))" using insert by simp
-qed
-
-lemma (in ring_of_sets) countably_additive_additive:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes posf: "positive M f" and ca: "countably_additive M f"
-  shows "additive M f"
-proof (auto simp add: additive_def)
-  fix x y
-  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
-  hence "disjoint_family (binaryset x y)"
-    by (auto simp add: disjoint_family_on_def binaryset_def)
-  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
-         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
-         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
-    using ca
-    by (simp add: countably_additive_def)
-  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
-         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
-    by (simp add: range_binaryset_eq UN_binaryset_eq)
-  thus "f (x \<union> y) = f x + f y" using posf x y
-    by (auto simp add: Un suminf_binaryset_eq positive_def)
-qed
-
-lemma (in algebra) increasing_additive_bound:
-  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" and ad: "additive M f"
-      and inc: "increasing M f"
-      and A: "range A \<subseteq> M"
-      and disj: "disjoint_family A"
-  shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
-proof (safe intro!: suminf_le_const)
-  fix N
-  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
-  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
-    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
-  also have "... \<le> f \<Omega>" using space_closed A
-    by (intro increasingD[OF inc] finite_UN) auto
-  finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
-qed (insert f A, auto simp: positive_def)
-
-lemma (in ring_of_sets) countably_additiveI_finite:
-  fixes \<mu> :: "'a set \<Rightarrow> ennreal"
-  assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
-  shows "countably_additive M \<mu>"
-proof (rule countably_additiveI)
-  fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
-
-  have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
-  from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
-
-  have inj_f: "inj_on f {i. F i \<noteq> {}}"
-  proof (rule inj_onI, simp)
-    fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
-    then have "f i \<in> F i" "f j \<in> F j" using f by force+
-    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
-  qed
-  have "finite (\<Union>i. F i)"
-    by (metis F(2) assms(1) infinite_super sets_into_space)
-
-  have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
-    by (auto simp: positiveD_empty[OF \<open>positive M \<mu>\<close>])
-  moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
-  proof (rule finite_imageD)
-    from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
-    then show "finite (f`{i. F i \<noteq> {}})"
-      by (rule finite_subset) fact
-  qed fact
-  ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
-    by (rule finite_subset)
-
-  have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
-    using disj by (auto simp: disjoint_family_on_def)
-
-  from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
-    by (rule suminf_finite) auto
-  also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
-    using fin_not_empty F_subset by (rule setsum.mono_neutral_left) auto
-  also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
-    using \<open>positive M \<mu>\<close> \<open>additive M \<mu>\<close> fin_not_empty disj_not_empty F by (intro additive_sum) auto
-  also have "\<dots> = \<mu> (\<Union>i. F i)"
-    by (rule arg_cong[where f=\<mu>]) auto
-  finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
-qed
-
-lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" "additive M f"
-  shows "countably_additive M f \<longleftrightarrow>
-    (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i))"
-  unfolding countably_additive_def
-proof safe
-  assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
-  then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets)
-  with count_sum[THEN spec, of "disjointed A"] A(3)
-  have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
-    by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
-  moreover have "(\<lambda>n. (\<Sum>i<n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
-    using f(1)[unfolded positive_def] dA
-    by (auto intro!: summable_LIMSEQ)
-  from LIMSEQ_Suc[OF this]
-  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) \<longlonglongrightarrow> (\<Sum>i. f (disjointed A i))"
-    unfolding lessThan_Suc_atMost .
-  moreover have "\<And>n. (\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
-    using disjointed_additive[OF f A(1,2)] .
-  ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)" by simp
-next
-  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M"
-  have *: "(\<Union>n. (\<Union>i<n. A i)) = (\<Union>i. A i)" by auto
-  have "(\<lambda>n. f (\<Union>i<n. A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
-  proof (unfold *[symmetric], intro cont[rule_format])
-    show "range (\<lambda>i. \<Union>i<i. A i) \<subseteq> M" "(\<Union>i. \<Union>i<i. A i) \<in> M"
-      using A * by auto
-  qed (force intro!: incseq_SucI)
-  moreover have "\<And>n. f (\<Union>i<n. A i) = (\<Sum>i<n. f (A i))"
-    using A
-    by (intro additive_sum[OF f, of _ A, symmetric])
-       (auto intro: disjoint_family_on_mono[where B=UNIV])
-  ultimately
-  have "(\<lambda>i. f (A i)) sums f (\<Union>i. A i)"
-    unfolding sums_def by simp
-  from sums_unique[OF this]
-  show "(\<Sum>i. f (A i)) = f (\<Union>i. A i)" by simp
-qed
-
-lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" "additive M f"
-  shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))
-     \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0)"
-proof safe
-  assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i))"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
-  with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
-    using \<open>positive M f\<close>[unfolded positive_def] by auto
-next
-  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
-
-  have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
-    using additive_increasing[OF f] unfolding increasing_def by simp
-
-  have decseq_fA: "decseq (\<lambda>i. f (A i))"
-    using A by (auto simp: decseq_def intro!: f_mono)
-  have decseq: "decseq (\<lambda>i. A i - (\<Inter>i. A i))"
-    using A by (auto simp: decseq_def)
-  then have decseq_f: "decseq (\<lambda>i. f (A i - (\<Inter>i. A i)))"
-    using A unfolding decseq_def by (auto intro!: f_mono Diff)
-  have "f (\<Inter>x. A x) \<le> f (A 0)"
-    using A by (auto intro!: f_mono)
-  then have f_Int_fin: "f (\<Inter>x. A x) \<noteq> \<infinity>"
-    using A by (auto simp: top_unique)
-  { fix i
-    have "f (A i - (\<Inter>i. A i)) \<le> f (A i)" using A by (auto intro!: f_mono)
-    then have "f (A i - (\<Inter>i. A i)) \<noteq> \<infinity>"
-      using A by (auto simp: top_unique) }
-  note f_fin = this
-  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) \<longlonglongrightarrow> 0"
-  proof (intro cont[rule_format, OF _ decseq _ f_fin])
-    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
-      using A by auto
-  qed
-  from INF_Lim_ereal[OF decseq_f this]
-  have "(INF n. f (A n - (\<Inter>i. A i))) = 0" .
-  moreover have "(INF n. f (\<Inter>i. A i)) = f (\<Inter>i. A i)"
-    by auto
-  ultimately have "(INF n. f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i)) = 0 + f (\<Inter>i. A i)"
-    using A(4) f_fin f_Int_fin
-    by (subst INF_ennreal_add_const) (auto simp: decseq_f)
-  moreover {
-    fix n
-    have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f ((A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i))"
-      using A by (subst f(2)[THEN additiveD]) auto
-    also have "(A n - (\<Inter>i. A i)) \<union> (\<Inter>i. A i) = A n"
-      by auto
-    finally have "f (A n - (\<Inter>i. A i)) + f (\<Inter>i. A i) = f (A n)" . }
-  ultimately have "(INF n. f (A n)) = f (\<Inter>i. A i)"
-    by simp
-  with LIMSEQ_INF[OF decseq_fA]
-  show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Inter>i. A i)" by simp
-qed
-
-lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>"
-  assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) \<longlonglongrightarrow> 0"
-  assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
-  shows "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
-proof -
-  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) \<longlonglongrightarrow> 0"
-    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
-  moreover
-  { fix i
-    have "f ((\<Union>i. A i) - A i \<union> A i) = f ((\<Union>i. A i) - A i) + f (A i)"
-      using A by (intro f(2)[THEN additiveD]) auto
-    also have "((\<Union>i. A i) - A i) \<union> A i = (\<Union>i. A i)"
-      by auto
-    finally have "f ((\<Union>i. A i) - A i) = f (\<Union>i. A i) - f (A i)"
-      using f(3)[rule_format, of "A i"] A by (auto simp: ennreal_add_diff_cancel subset_eq) }
-  moreover have "\<forall>\<^sub>F i in sequentially. f (A i) \<le> f (\<Union>i. A i)"
-    using increasingD[OF additive_increasing[OF f(1, 2)], of "A _" "\<Union>i. A i"] A
-    by (auto intro!: always_eventually simp: subset_eq)
-  ultimately show "(\<lambda>i. f (A i)) \<longlonglongrightarrow> f (\<Union>i. A i)"
-    by (auto intro: ennreal_tendsto_const_minus)
-qed
-
-lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
-  fixes f :: "'a set \<Rightarrow> ennreal"
-  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. 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 "countably_additive M f"
-  using countably_additive_iff_continuous_from_below[OF f]
-  using empty_continuous_imp_continuous_from_below[OF f fin] cont
-  by blast
-
-subsection \<open>Properties of @{const emeasure}\<close>
-
-lemma emeasure_positive: "positive (sets M) (emeasure M)"
-  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
-
-lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
-  using emeasure_positive[of M] by (simp add: positive_def)
-
-lemma emeasure_single_in_space: "emeasure M {x} \<noteq> 0 \<Longrightarrow> x \<in> space M"
-  using emeasure_notin_sets[of "{x}" M] by (auto dest: sets.sets_into_space zero_less_iff_neq_zero[THEN iffD2])
-
-lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
-  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
-
-lemma suminf_emeasure:
-  "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
-  using sets.countable_UN[of A UNIV M] emeasure_countably_additive[of M]
-  by (simp add: countably_additive_def)
-
-lemma sums_emeasure:
-  "disjoint_family F \<Longrightarrow> (\<And>i. F i \<in> sets M) \<Longrightarrow> (\<lambda>i. emeasure M (F i)) sums emeasure M (\<Union>i. F i)"
-  unfolding sums_iff by (intro conjI suminf_emeasure) auto
-
-lemma emeasure_additive: "additive (sets M) (emeasure M)"
-  by (metis sets.countably_additive_additive emeasure_positive emeasure_countably_additive)
-
-lemma plus_emeasure:
-  "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
-  using additiveD[OF emeasure_additive] ..
-
-lemma setsum_emeasure:
-  "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
-    (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
-  by (metis sets.additive_sum emeasure_positive emeasure_additive)
-
-lemma emeasure_mono:
-  "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
-  by (metis zero_le sets.additive_increasing emeasure_additive emeasure_notin_sets emeasure_positive increasingD)
-
-lemma emeasure_space:
-  "emeasure M A \<le> emeasure M (space M)"
-  by (metis emeasure_mono emeasure_notin_sets sets.sets_into_space sets.top zero_le)
-
-lemma emeasure_Diff:
-  assumes finite: "emeasure M B \<noteq> \<infinity>"
-  and [measurable]: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
-  shows "emeasure M (A - B) = emeasure M A - emeasure M B"
-proof -
-  have "(A - B) \<union> B = A" using \<open>B \<subseteq> A\<close> by auto
-  then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
-  also have "\<dots> = emeasure M (A - B) + emeasure M B"
-    by (subst plus_emeasure[symmetric]) auto
-  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
-    using finite by simp
-qed
-
-lemma emeasure_compl:
-  "s \<in> sets M \<Longrightarrow> emeasure M s \<noteq> \<infinity> \<Longrightarrow> emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
-  by (rule emeasure_Diff) (auto dest: sets.sets_into_space)
-
-lemma Lim_emeasure_incseq:
-  "range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) \<longlonglongrightarrow> emeasure M (\<Union>i. A i)"
-  using emeasure_countably_additive
-  by (auto simp add: sets.countably_additive_iff_continuous_from_below emeasure_positive
-    emeasure_additive)
-
-lemma incseq_emeasure:
-  assumes "range B \<subseteq> sets M" "incseq B"
-  shows "incseq (\<lambda>i. emeasure M (B i))"
-  using assms by (auto simp: incseq_def intro!: emeasure_mono)
-
-lemma SUP_emeasure_incseq:
-  assumes A: "range A \<subseteq> sets M" "incseq A"
-  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
-  using LIMSEQ_SUP[OF incseq_emeasure, OF A] Lim_emeasure_incseq[OF A]
-  by (simp add: LIMSEQ_unique)
-
-lemma decseq_emeasure:
-  assumes "range B \<subseteq> sets M" "decseq B"
-  shows "decseq (\<lambda>i. emeasure M (B i))"
-  using assms by (auto simp: decseq_def intro!: emeasure_mono)
-
-lemma INF_emeasure_decseq:
-  assumes A: "range A \<subseteq> sets M" and "decseq A"
-  and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
-  shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
-proof -
-  have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
-    using A by (auto intro!: emeasure_mono)
-  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by (auto simp: top_unique)
-
-  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = (SUP n. emeasure M (A 0) - emeasure M (A n))"
-    by (simp add: ennreal_INF_const_minus)
-  also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
-    using A finite \<open>decseq A\<close>[unfolded decseq_def] by (subst emeasure_Diff) auto
-  also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
-  proof (rule SUP_emeasure_incseq)
-    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
-      using A by auto
-    show "incseq (\<lambda>n. A 0 - A n)"
-      using \<open>decseq A\<close> by (auto simp add: incseq_def decseq_def)
-  qed
-  also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
-    using A finite * by (simp, subst emeasure_Diff) auto
-  finally show ?thesis
-    by (rule ennreal_minus_cancel[rotated 3])
-       (insert finite A, auto intro: INF_lower emeasure_mono)
-qed
-
-lemma emeasure_INT_decseq_subset:
-  fixes F :: "nat \<Rightarrow> 'a set"
-  assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
-  assumes F_sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M"
-    and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (F i) \<noteq> \<infinity>"
-  shows "emeasure M (\<Inter>i\<in>I. F i) = (INF i:I. emeasure M (F i))"
-proof cases
-  assume "finite I"
-  have "(\<Inter>i\<in>I. F i) = F (Max I)"
-    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F) auto
-  moreover have "(INF i:I. emeasure M (F i)) = emeasure M (F (Max I))"
-    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F emeasure_mono) auto
-  ultimately show ?thesis
-    by simp
-next
-  assume "infinite I"
-  define L where "L n = (LEAST i. i \<in> I \<and> i \<ge> n)" for n
-  have L: "L n \<in> I \<and> n \<le> L n" for n
-    unfolding L_def
-  proof (rule LeastI_ex)
-    show "\<exists>x. x \<in> I \<and> n \<le> x"
-      using \<open>infinite I\<close> finite_subset[of I "{..< n}"]
-      by (rule_tac ccontr) (auto simp: not_le)
-  qed
-  have L_eq[simp]: "i \<in> I \<Longrightarrow> L i = i" for i
-    unfolding L_def by (intro Least_equality) auto
-  have L_mono: "i \<le> j \<Longrightarrow> L i \<le> L j" for i j
-    using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def)
-
-  have "emeasure M (\<Inter>i. F (L i)) = (INF i. emeasure M (F (L i)))"
-  proof (intro INF_emeasure_decseq[symmetric])
-    show "decseq (\<lambda>i. F (L i))"
-      using L by (intro antimonoI F L_mono) auto
-  qed (insert L fin, auto)
-  also have "\<dots> = (INF i:I. emeasure M (F i))"
-  proof (intro antisym INF_greatest)
-    show "i \<in> I \<Longrightarrow> (INF i. emeasure M (F (L i))) \<le> emeasure M (F i)" for i
-      by (intro INF_lower2[of i]) auto
-  qed (insert L, auto intro: INF_lower)
-  also have "(\<Inter>i. F (L i)) = (\<Inter>i\<in>I. F i)"
-  proof (intro antisym INF_greatest)
-    show "i \<in> I \<Longrightarrow> (\<Inter>i. F (L i)) \<subseteq> F i" for i
-      by (intro INF_lower2[of i]) auto
-  qed (insert L, auto)
-  finally show ?thesis .
-qed
-
-lemma Lim_emeasure_decseq:
-  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. emeasure M (A i)) \<longlonglongrightarrow> emeasure M (\<Inter>i. A i)"
-  using LIMSEQ_INF[OF decseq_emeasure, OF A]
-  using INF_emeasure_decseq[OF A fin] by simp
-
-lemma emeasure_lfp'[consumes 1, case_names cont measurable]:
-  assumes "P M"
-  assumes cont: "sup_continuous F"
-  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
-  shows "emeasure M {x\<in>space M. lfp F x} = (SUP i. emeasure M {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
-proof -
-  have "emeasure M {x\<in>space M. lfp F x} = emeasure M (\<Union>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
-    using sup_continuous_lfp[OF cont] by (auto simp add: bot_fun_def intro!: arg_cong2[where f=emeasure])
-  moreover { fix i from \<open>P M\<close> have "{x\<in>space M. (F ^^ i) (\<lambda>x. False) x} \<in> sets M"
-    by (induct i arbitrary: M) (auto simp add: pred_def[symmetric] intro: *) }
-  moreover have "incseq (\<lambda>i. {x\<in>space M. (F ^^ i) (\<lambda>x. False) x})"
-  proof (rule incseq_SucI)
-    fix i
-    have "(F ^^ i) (\<lambda>x. False) \<le> (F ^^ (Suc i)) (\<lambda>x. False)"
-    proof (induct i)
-      case 0 show ?case by (simp add: le_fun_def)
-    next
-      case Suc thus ?case using monoD[OF sup_continuous_mono[OF cont] Suc] by auto
-    qed
-    then show "{x \<in> space M. (F ^^ i) (\<lambda>x. False) x} \<subseteq> {x \<in> space M. (F ^^ Suc i) (\<lambda>x. False) x}"
-      by auto
-  qed
-  ultimately show ?thesis
-    by (subst SUP_emeasure_incseq) auto
-qed
-
-lemma emeasure_lfp:
-  assumes [simp]: "\<And>s. sets (M s) = sets N"
-  assumes cont: "sup_continuous F" "sup_continuous f"
-  assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
-  assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> P \<le> lfp F \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
-  shows "emeasure (M s) {x\<in>space N. lfp F x} = lfp f s"
-proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and P="Measurable.pred N", symmetric])
-  fix C assume "incseq C" "\<And>i. Measurable.pred N (C i)"
-  then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (SUP i. C i) x}) = (SUP i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
-    unfolding SUP_apply[abs_def]
-    by (subst SUP_emeasure_incseq) (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
-qed (auto simp add: iter le_fun_def SUP_apply[abs_def] intro!: meas cont)
-
-lemma emeasure_subadditive_finite:
-  "finite I \<Longrightarrow> A ` I \<subseteq> sets M \<Longrightarrow> emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
-  by (rule sets.subadditive[OF emeasure_positive emeasure_additive]) auto
-
-lemma emeasure_subadditive:
-  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
-  using emeasure_subadditive_finite[of "{True, False}" "\<lambda>True \<Rightarrow> A | False \<Rightarrow> B" M] by simp
-
-lemma emeasure_subadditive_countably:
-  assumes "range f \<subseteq> sets M"
-  shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
-proof -
-  have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
-    unfolding UN_disjointed_eq ..
-  also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
-    using sets.range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
-    by (simp add:  disjoint_family_disjointed comp_def)
-  also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
-    using sets.range_disjointed_sets[OF assms] assms
-    by (auto intro!: suminf_le emeasure_mono disjointed_subset)
-  finally show ?thesis .
-qed
-
-lemma emeasure_insert:
-  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
-  shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
-proof -
-  have "{x} \<inter> A = {}" using \<open>x \<notin> A\<close> by auto
-  from plus_emeasure[OF sets this] show ?thesis by simp
-qed
-
-lemma emeasure_insert_ne:
-  "A \<noteq> {} \<Longrightarrow> {x} \<in> sets M \<Longrightarrow> A \<in> sets M \<Longrightarrow> x \<notin> A \<Longrightarrow> emeasure M (insert x A) = emeasure M {x} + emeasure M A"
-  by (rule emeasure_insert)
-
-lemma emeasure_eq_setsum_singleton:
-  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
-  using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
-  by (auto simp: disjoint_family_on_def subset_eq)
-
-lemma setsum_emeasure_cover:
-  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
-  assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
-  assumes disj: "disjoint_family_on B S"
-  shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
-proof -
-  have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
-  proof (rule setsum_emeasure)
-    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
-      using \<open>disjoint_family_on B S\<close>
-      unfolding disjoint_family_on_def by auto
-  qed (insert assms, auto)
-  also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
-    using A by auto
-  finally show ?thesis by simp
-qed
-
-lemma emeasure_eq_0:
-  "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
-  by (metis emeasure_mono order_eq_iff zero_le)
-
-lemma emeasure_UN_eq_0:
-  assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
-  shows "emeasure M (\<Union>i. N i) = 0"
-proof -
-  have "emeasure M (\<Union>i. N i) \<le> 0"
-    using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
-  then show ?thesis
-    by (auto intro: antisym zero_le)
-qed
-
-lemma measure_eqI_finite:
-  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
-  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
-  shows "M = N"
-proof (rule measure_eqI)
-  fix X assume "X \<in> sets M"
-  then have X: "X \<subseteq> A" by auto
-  then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
-    using \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
-  also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
-    using X eq by (auto intro!: setsum.cong)
-  also have "\<dots> = emeasure N X"
-    using X \<open>finite A\<close> by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
-  finally show "emeasure M X = emeasure N X" .
-qed simp
-
-lemma measure_eqI_generator_eq:
-  fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
-  assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
-  and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
-  and M: "sets M = sigma_sets \<Omega> E"
-  and N: "sets N = sigma_sets \<Omega> E"
-  and A: "range A \<subseteq> E" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
-  shows "M = N"
-proof -
-  let ?\<mu>  = "emeasure M" and ?\<nu> = "emeasure N"
-  interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
-  have "space M = \<Omega>"
-    using sets.top[of M] sets.space_closed[of M] S.top S.space_closed \<open>sets M = sigma_sets \<Omega> E\<close>
-    by blast
-
-  { fix F D assume "F \<in> E" and "?\<mu> F \<noteq> \<infinity>"
-    then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
-    have "?\<nu> F \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> \<open>F \<in> E\<close> eq by simp
-    assume "D \<in> sets M"
-    with \<open>Int_stable E\<close> \<open>E \<subseteq> Pow \<Omega>\<close> have "emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
-      unfolding M
-    proof (induct rule: sigma_sets_induct_disjoint)
-      case (basic A)
-      then have "F \<inter> A \<in> E" using \<open>Int_stable E\<close> \<open>F \<in> E\<close> by (auto simp: Int_stable_def)
-      then show ?case using eq by auto
-    next
-      case empty then show ?case by simp
-    next
-      case (compl A)
-      then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
-        and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
-        using \<open>F \<in> E\<close> S.sets_into_space by (auto simp: M)
-      have "?\<nu> (F \<inter> A) \<le> ?\<nu> F" by (auto intro!: emeasure_mono simp: M N)
-      then have "?\<nu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<nu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
-      have "?\<mu> (F \<inter> A) \<le> ?\<mu> F" by (auto intro!: emeasure_mono simp: M N)
-      then have "?\<mu> (F \<inter> A) \<noteq> \<infinity>" using \<open>?\<mu> F \<noteq> \<infinity>\<close> by (auto simp: top_unique)
-      then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
-        using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> by (auto intro!: emeasure_Diff simp: M N)
-      also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq \<open>F \<in> E\<close> compl by simp
-      also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding **
-        using \<open>F \<inter> A \<in> sigma_sets \<Omega> E\<close> \<open>?\<nu> (F \<inter> A) \<noteq> \<infinity>\<close>
-        by (auto intro!: emeasure_Diff[symmetric] simp: M N)
-      finally show ?case
-        using \<open>space M = \<Omega>\<close> by auto
-    next
-      case (union A)
-      then have "?\<mu> (\<Union>x. F \<inter> A x) = ?\<nu> (\<Union>x. F \<inter> A x)"
-        by (subst (1 2) suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def subset_eq M N)
-      with A show ?case
-        by auto
-    qed }
-  note * = this
-  show "M = N"
-  proof (rule measure_eqI)
-    show "sets M = sets N"
-      using M N by simp
-    have [simp, intro]: "\<And>i. A i \<in> sets M"
-      using A(1) by (auto simp: subset_eq M)
-    fix F assume "F \<in> sets M"
-    let ?D = "disjointed (\<lambda>i. F \<inter> A i)"
-    from \<open>space M = \<Omega>\<close> have F_eq: "F = (\<Union>i. ?D i)"
-      using \<open>F \<in> sets M\<close>[THEN sets.sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
-    have [simp, intro]: "\<And>i. ?D i \<in> sets M"
-      using sets.range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] \<open>F \<in> sets M\<close>
-      by (auto simp: subset_eq)
-    have "disjoint_family ?D"
-      by (auto simp: disjoint_family_disjointed)
-    moreover
-    have "(\<Sum>i. emeasure M (?D i)) = (\<Sum>i. emeasure N (?D i))"
-    proof (intro arg_cong[where f=suminf] ext)
-      fix i
-      have "A i \<inter> ?D i = ?D i"
-        by (auto simp: disjointed_def)
-      then show "emeasure M (?D i) = emeasure N (?D i)"
-        using *[of "A i" "?D i", OF _ A(3)] A(1) by auto
-    qed
-    ultimately show "emeasure M F = emeasure N F"
-      by (simp add: image_subset_iff \<open>sets M = sets N\<close>[symmetric] F_eq[symmetric] suminf_emeasure)
-  qed
-qed
-
-lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
-proof (intro measure_eqI emeasure_measure_of_sigma)
-  show "sigma_algebra (space M) (sets M)" ..
-  show "positive (sets M) (emeasure M)"
-    by (simp add: positive_def)
-  show "countably_additive (sets M) (emeasure M)"
-    by (simp add: emeasure_countably_additive)
-qed simp_all
-
-subsection \<open>\<open>\<mu>\<close>-null sets\<close>
-
-definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
-  "null_sets M = {N\<in>sets M. emeasure M N = 0}"
-
-lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
-  by (simp add: null_sets_def)
-
-lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
-  unfolding null_sets_def by simp
-
-lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
-  unfolding null_sets_def by simp
-
-interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
-proof (rule ring_of_setsI)
-  show "null_sets M \<subseteq> Pow (space M)"
-    using sets.sets_into_space by auto
-  show "{} \<in> null_sets M"
-    by auto
-  fix A B assume null_sets: "A \<in> null_sets M" "B \<in> null_sets M"
-  then have sets: "A \<in> sets M" "B \<in> sets M"
-    by auto
-  then have *: "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
-    "emeasure M (A - B) \<le> emeasure M A"
-    by (auto intro!: emeasure_subadditive emeasure_mono)
-  then have "emeasure M B = 0" "emeasure M A = 0"
-    using null_sets by auto
-  with sets * show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
-    by (auto intro!: antisym zero_le)
-qed
-
-lemma UN_from_nat_into:
-  assumes I: "countable I" "I \<noteq> {}"
-  shows "(\<Union>i\<in>I. N i) = (\<Union>i. N (from_nat_into I i))"
-proof -
-  have "(\<Union>i\<in>I. N i) = \<Union>(N ` range (from_nat_into I))"
-    using I by simp
-  also have "\<dots> = (\<Union>i. (N \<circ> from_nat_into I) i)"
-    by simp
-  finally show ?thesis by simp
-qed
-
-lemma null_sets_UN':
-  assumes "countable I"
-  assumes "\<And>i. i \<in> I \<Longrightarrow> N i \<in> null_sets M"
-  shows "(\<Union>i\<in>I. N i) \<in> null_sets M"
-proof cases
-  assume "I = {}" then show ?thesis by simp
-next
-  assume "I \<noteq> {}"
-  show ?thesis
-  proof (intro conjI CollectI null_setsI)
-    show "(\<Union>i\<in>I. N i) \<in> sets M"
-      using assms by (intro sets.countable_UN') auto
-    have "emeasure M (\<Union>i\<in>I. N i) \<le> (\<Sum>n. emeasure M (N (from_nat_into I n)))"
-      unfolding UN_from_nat_into[OF \<open>countable I\<close> \<open>I \<noteq> {}\<close>]
-      using assms \<open>I \<noteq> {}\<close> by (intro emeasure_subadditive_countably) (auto intro: from_nat_into)
-    also have "(\<lambda>n. emeasure M (N (from_nat_into I n))) = (\<lambda>_. 0)"
-      using assms \<open>I \<noteq> {}\<close> by (auto intro: from_nat_into)
-    finally show "emeasure M (\<Union>i\<in>I. N i) = 0"
-      by (intro antisym zero_le) simp
-  qed
-qed
-
-lemma null_sets_UN[intro]:
-  "(\<And>i::'i::countable. N i \<in> null_sets M) \<Longrightarrow> (\<Union>i. N i) \<in> null_sets M"
-  by (rule null_sets_UN') auto
-
-lemma null_set_Int1:
-  assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
-proof (intro CollectI conjI null_setsI)
-  show "emeasure M (A \<inter> B) = 0" using assms
-    by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
-qed (insert assms, auto)
-
-lemma null_set_Int2:
-  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
-  using assms by (subst Int_commute) (rule null_set_Int1)
-
-lemma emeasure_Diff_null_set:
-  assumes "B \<in> null_sets M" "A \<in> sets M"
-  shows "emeasure M (A - B) = emeasure M A"
-proof -
-  have *: "A - B = (A - (A \<inter> B))" by auto
-  have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
-  then show ?thesis
-    unfolding * using assms
-    by (subst emeasure_Diff) auto
-qed
-
-lemma null_set_Diff:
-  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
-proof (intro CollectI conjI null_setsI)
-  show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
-qed (insert assms, auto)
-
-lemma emeasure_Un_null_set:
-  assumes "A \<in> sets M" "B \<in> null_sets M"
-  shows "emeasure M (A \<union> B) = emeasure M A"
-proof -
-  have *: "A \<union> B = A \<union> (B - A)" by auto
-  have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
-  then show ?thesis
-    unfolding * using assms
-    by (subst plus_emeasure[symmetric]) auto
-qed
-
-subsection \<open>The almost everywhere filter (i.e.\ quantifier)\<close>
-
-definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
-  "ae_filter M = (INF N:null_sets M. principal (space M - N))"
-
-abbreviation almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
-  "almost_everywhere M P \<equiv> eventually P (ae_filter M)"
-
-syntax
-  "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
-
-translations
-  "AE x in M. P" \<rightleftharpoons> "CONST almost_everywhere M (\<lambda>x. P)"
-
-lemma eventually_ae_filter: "eventually P (ae_filter M) \<longleftrightarrow> (\<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
-  unfolding ae_filter_def by (subst eventually_INF_base) (auto simp: eventually_principal subset_eq)
-
-lemma AE_I':
-  "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
-  unfolding eventually_ae_filter by auto
-
-lemma AE_iff_null:
-  assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
-  shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
-proof
-  assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
-    unfolding eventually_ae_filter by auto
-  have "emeasure M ?P \<le> emeasure M N"
-    using assms N(1,2) by (auto intro: emeasure_mono)
-  then have "emeasure M ?P = 0"
-    unfolding \<open>emeasure M N = 0\<close> by auto
-  then show "?P \<in> null_sets M" using assms by auto
-next
-  assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
-qed
-
-lemma AE_iff_null_sets:
-  "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
-  using Int_absorb1[OF sets.sets_into_space, of N M]
-  by (subst AE_iff_null) (auto simp: Int_def[symmetric])
-
-lemma AE_not_in:
-  "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
-  by (metis AE_iff_null_sets null_setsD2)
-
-lemma AE_iff_measurable:
-  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
-  using AE_iff_null[of _ P] by auto
-
-lemma AE_E[consumes 1]:
-  assumes "AE x in M. P x"
-  obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
-  using assms unfolding eventually_ae_filter by auto
-
-lemma AE_E2:
-  assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
-  shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
-proof -
-  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
-  with AE_iff_null[of M P] assms show ?thesis by auto
-qed
-
-lemma AE_I:
-  assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
-  shows "AE x in M. P x"
-  using assms unfolding eventually_ae_filter by auto
-
-lemma AE_mp[elim!]:
-  assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
-  shows "AE x in M. Q x"
-proof -
-  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
-    and A: "A \<in> sets M" "emeasure M A = 0"
-    by (auto elim!: AE_E)
-
-  from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
-    and B: "B \<in> sets M" "emeasure M B = 0"
-    by (auto elim!: AE_E)
-
-  show ?thesis
-  proof (intro AE_I)
-    have "emeasure M (A \<union> B) \<le> 0"
-      using emeasure_subadditive[of A M B] A B by auto
-    then show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0"
-      using A B by auto
-    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
-      using P imp by auto
-  qed
-qed
-
-(* depricated replace by laws about eventually *)
-lemma
-  shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
-    and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
-    and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
-    and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x"
-    and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)"
-  by auto
-
-lemma AE_impI:
-  "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
-  by (cases P) auto
-
-lemma AE_measure:
-  assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
-  shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
-proof -
-  from AE_E[OF AE] guess N . note N = this
-  with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
-    by (intro emeasure_mono) auto
-  also have "\<dots> \<le> emeasure M ?P + emeasure M N"
-    using sets N by (intro emeasure_subadditive) auto
-  also have "\<dots> = emeasure M ?P" using N by simp
-  finally show "emeasure M ?P = emeasure M (space M)"
-    using emeasure_space[of M "?P"] by auto
-qed
-
-lemma AE_space: "AE x in M. x \<in> space M"
-  by (rule AE_I[where N="{}"]) auto
-
-lemma AE_I2[simp, intro]:
-  "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
-  using AE_space by force
-
-lemma AE_Ball_mp:
-  "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
-  by auto
-
-lemma AE_cong[cong]:
-  "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)"
-  by auto
-
-lemma AE_all_countable:
-  "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
-proof
-  assume "\<forall>i. AE x in M. P i x"
-  from this[unfolded eventually_ae_filter Bex_def, THEN choice]
-  obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
-  have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
-  also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
-  finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
-  moreover from N have "(\<Union>i. N i) \<in> null_sets M"
-    by (intro null_sets_UN) auto
-  ultimately show "AE x in M. \<forall>i. P i x"
-    unfolding eventually_ae_filter by auto
-qed auto
-
-lemma AE_ball_countable:
-  assumes [intro]: "countable X"
-  shows "(AE x in M. \<forall>y\<in>X. P x y) \<longleftrightarrow> (\<forall>y\<in>X. AE x in M. P x y)"
-proof
-  assume "\<forall>y\<in>X. AE x in M. P x y"
-  from this[unfolded eventually_ae_filter Bex_def, THEN bchoice]
-  obtain N where N: "\<And>y. y \<in> X \<Longrightarrow> N y \<in> null_sets M" "\<And>y. y \<in> X \<Longrightarrow> {x\<in>space M. \<not> P x y} \<subseteq> N y"
-    by auto
-  have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. {x\<in>space M. \<not> P x y})"
-    by auto
-  also have "\<dots> \<subseteq> (\<Union>y\<in>X. N y)"
-    using N by auto
-  finally have "{x\<in>space M. \<not> (\<forall>y\<in>X. P x y)} \<subseteq> (\<Union>y\<in>X. N y)" .
-  moreover from N have "(\<Union>y\<in>X. N y) \<in> null_sets M"
-    by (intro null_sets_UN') auto
-  ultimately show "AE x in M. \<forall>y\<in>X. P x y"
-    unfolding eventually_ae_filter by auto
-qed auto
-
-lemma AE_discrete_difference:
-  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"
-  shows "AE x in M. x \<notin> X"
-proof -
-  have "(\<Union>x\<in>X. {x}) \<in> null_sets M"
-    using assms by (intro null_sets_UN') auto
-  from AE_not_in[OF this] show "AE x in M. x \<notin> X"
-    by auto
-qed
-
-lemma AE_finite_all:
-  assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)"
-  using f by induct auto
-
-lemma AE_finite_allI:
-  assumes "finite S"
-  shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
-  using AE_finite_all[OF \<open>finite S\<close>] by auto
-
-lemma emeasure_mono_AE:
-  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
-    and B: "B \<in> sets M"
-  shows "emeasure M A \<le> emeasure M B"
-proof cases
-  assume A: "A \<in> sets M"
-  from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
-    by (auto simp: eventually_ae_filter)
-  have "emeasure M A = emeasure M (A - N)"
-    using N A by (subst emeasure_Diff_null_set) auto
-  also have "emeasure M (A - N) \<le> emeasure M (B - N)"
-    using N A B sets.sets_into_space by (auto intro!: emeasure_mono)
-  also have "emeasure M (B - N) = emeasure M B"
-    using N B by (subst emeasure_Diff_null_set) auto
-  finally show ?thesis .
-qed (simp add: emeasure_notin_sets)
-
-lemma emeasure_eq_AE:
-  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
-  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
-  shows "emeasure M A = emeasure M B"
-  using assms by (safe intro!: antisym emeasure_mono_AE) auto
-
-lemma emeasure_Collect_eq_AE:
-  "AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> Measurable.pred M Q \<Longrightarrow> Measurable.pred M P \<Longrightarrow>
-   emeasure M {x\<in>space M. P x} = emeasure M {x\<in>space M. Q x}"
-   by (intro emeasure_eq_AE) auto
-
-lemma emeasure_eq_0_AE: "AE x in M. \<not> P x \<Longrightarrow> emeasure M {x\<in>space M. P x} = 0"
-  using AE_iff_measurable[OF _ refl, of M "\<lambda>x. \<not> P x"]
-  by (cases "{x\<in>space M. P x} \<in> sets M") (simp_all add: emeasure_notin_sets)
-
-lemma emeasure_add_AE:
-  assumes [measurable]: "A \<in> sets M" "B \<in> sets M" "C \<in> sets M"
-  assumes 1: "AE x in M. x \<in> C \<longleftrightarrow> x \<in> A \<or> x \<in> B"
-  assumes 2: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)"
-  shows "emeasure M C = emeasure M A + emeasure M B"
-proof -
-  have "emeasure M C = emeasure M (A \<union> B)"
-    by (rule emeasure_eq_AE) (insert 1, auto)
-  also have "\<dots> = emeasure M A + emeasure M (B - A)"
-    by (subst plus_emeasure) auto
-  also have "emeasure M (B - A) = emeasure M B"
-    by (rule emeasure_eq_AE) (insert 2, auto)
-  finally show ?thesis .
-qed
-
-subsection \<open>\<open>\<sigma>\<close>-finite Measures\<close>
-
-locale sigma_finite_measure =
-  fixes M :: "'a measure"
-  assumes sigma_finite_countable:
-    "\<exists>A::'a set set. countable A \<and> A \<subseteq> sets M \<and> (\<Union>A) = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"
-
-lemma (in sigma_finite_measure) sigma_finite:
-  obtains A :: "nat \<Rightarrow> 'a set"
-  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
-proof -
-  obtain A :: "'a set set" where
-    [simp]: "countable A" and
-    A: "A \<subseteq> sets M" "(\<Union>A) = space M" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
-    using sigma_finite_countable by metis
-  show thesis
-  proof cases
-    assume "A = {}" with \<open>(\<Union>A) = space M\<close> show thesis
-      by (intro that[of "\<lambda>_. {}"]) auto
-  next
-    assume "A \<noteq> {}"
-    show thesis
-    proof
-      show "range (from_nat_into A) \<subseteq> sets M"
-        using \<open>A \<noteq> {}\<close> A by auto
-      have "(\<Union>i. from_nat_into A i) = \<Union>A"
-        using range_from_nat_into[OF \<open>A \<noteq> {}\<close> \<open>countable A\<close>] by auto
-      with A show "(\<Union>i. from_nat_into A i) = space M"
-        by auto
-    qed (intro A from_nat_into \<open>A \<noteq> {}\<close>)
-  qed
-qed
-
-lemma (in sigma_finite_measure) sigma_finite_disjoint:
-  obtains A :: "nat \<Rightarrow> 'a set"
-  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
-proof -
-  obtain A :: "nat \<Rightarrow> 'a set" where
-    range: "range A \<subseteq> sets M" and
-    space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
-    using sigma_finite by blast
-  show thesis
-  proof (rule that[of "disjointed A"])
-    show "range (disjointed A) \<subseteq> sets M"
-      by (rule sets.range_disjointed_sets[OF range])
-    show "(\<Union>i. disjointed A i) = space M"
-      and "disjoint_family (disjointed A)"
-      using disjoint_family_disjointed UN_disjointed_eq[of A] space range
-      by auto
-    show "emeasure M (disjointed A i) \<noteq> \<infinity>" for i
-    proof -
-      have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
-        using range disjointed_subset[of A i] by (auto intro!: emeasure_mono)
-      then show ?thesis using measure[of i] by (auto simp: top_unique)
-    qed
-  qed
-qed
-
-lemma (in sigma_finite_measure) sigma_finite_incseq:
-  obtains A :: "nat \<Rightarrow> 'a set"
-  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
-proof -
-  obtain F :: "nat \<Rightarrow> 'a set" where
-    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
-    using sigma_finite by blast
-  show thesis
-  proof (rule that[of "\<lambda>n. \<Union>i\<le>n. F i"])
-    show "range (\<lambda>n. \<Union>i\<le>n. F i) \<subseteq> sets M"
-      using F by (force simp: incseq_def)
-    show "(\<Union>n. \<Union>i\<le>n. F i) = space M"
-    proof -
-      from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
-      with F show ?thesis by fastforce
-    qed
-    show "emeasure M (\<Union>i\<le>n. F i) \<noteq> \<infinity>" for n
-    proof -
-      have "emeasure M (\<Union>i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))"
-        using F by (auto intro!: emeasure_subadditive_finite)
-      also have "\<dots> < \<infinity>"
-        using F by (auto simp: setsum_Pinfty less_top)
-      finally show ?thesis by simp
-    qed
-    show "incseq (\<lambda>n. \<Union>i\<le>n. F i)"
-      by (force simp: incseq_def)
-  qed
-qed
-
-subsection \<open>Measure space induced by distribution of @{const measurable}-functions\<close>
-
-definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
-  "distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
-
-lemma
-  shows sets_distr[simp, measurable_cong]: "sets (distr M N f) = sets N"
-    and space_distr[simp]: "space (distr M N f) = space N"
-  by (auto simp: distr_def)
-
-lemma
-  shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
-    and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
-  by (auto simp: measurable_def)
-
-lemma distr_cong:
-  "M = K \<Longrightarrow> sets N = sets L \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> distr M N f = distr K L g"
-  using sets_eq_imp_space_eq[of N L] by (simp add: distr_def Int_def cong: rev_conj_cong)
-
-lemma emeasure_distr:
-  fixes f :: "'a \<Rightarrow> 'b"
-  assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
-  shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
-  unfolding distr_def
-proof (rule emeasure_measure_of_sigma)
-  show "positive (sets N) ?\<mu>"
-    by (auto simp: positive_def)
-
-  show "countably_additive (sets N) ?\<mu>"
-  proof (intro countably_additiveI)
-    fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
-    then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
-    then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
-      using f by (auto simp: measurable_def)
-    moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
-      using * by blast
-    moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
-      using \<open>disjoint_family A\<close> by (auto simp: disjoint_family_on_def)
-    ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
-      using suminf_emeasure[OF _ **] A f
-      by (auto simp: comp_def vimage_UN)
-  qed
-  show "sigma_algebra (space N) (sets N)" ..
-qed fact
-
-lemma emeasure_Collect_distr:
-  assumes X[measurable]: "X \<in> measurable M N" "Measurable.pred N P"
-  shows "emeasure (distr M N X) {x\<in>space N. P x} = emeasure M {x\<in>space M. P (X x)}"
-  by (subst emeasure_distr)
-     (auto intro!: arg_cong2[where f=emeasure] X(1)[THEN measurable_space])
-
-lemma emeasure_lfp2[consumes 1, case_names cont f measurable]:
-  assumes "P M"
-  assumes cont: "sup_continuous F"
-  assumes f: "\<And>M. P M \<Longrightarrow> f \<in> measurable M' M"
-  assumes *: "\<And>M A. P M \<Longrightarrow> (\<And>N. P N \<Longrightarrow> Measurable.pred N A) \<Longrightarrow> Measurable.pred M (F A)"
-  shows "emeasure M' {x\<in>space M'. lfp F (f x)} = (SUP i. emeasure M' {x\<in>space M'. (F ^^ i) (\<lambda>x. False) (f x)})"
-proof (subst (1 2) emeasure_Collect_distr[symmetric, where X=f])
-  show "f \<in> measurable M' M"  "f \<in> measurable M' M"
-    using f[OF \<open>P M\<close>] by auto
-  { fix i show "Measurable.pred M ((F ^^ i) (\<lambda>x. False))"
-    using \<open>P M\<close> by (induction i arbitrary: M) (auto intro!: *) }
-  show "Measurable.pred M (lfp F)"
-    using \<open>P M\<close> cont * by (rule measurable_lfp_coinduct[of P])
-
-  have "emeasure (distr M' M f) {x \<in> space (distr M' M f). lfp F x} =
-    (SUP i. emeasure (distr M' M f) {x \<in> space (distr M' M f). (F ^^ i) (\<lambda>x. False) x})"
-    using \<open>P M\<close>
-  proof (coinduction arbitrary: M rule: emeasure_lfp')
-    case (measurable A N) then have "\<And>N. P N \<Longrightarrow> Measurable.pred (distr M' N f) A"
-      by metis
-    then have "\<And>N. P N \<Longrightarrow> Measurable.pred N A"
-      by simp
-    with \<open>P N\<close>[THEN *] show ?case
-      by auto
-  qed fact
-  then show "emeasure (distr M' M f) {x \<in> space M. lfp F x} =
-    (SUP i. emeasure (distr M' M f) {x \<in> space M. (F ^^ i) (\<lambda>x. False) x})"
-   by simp
-qed
-
-lemma distr_id[simp]: "distr N N (\<lambda>x. x) = N"
-  by (rule measure_eqI) (auto simp: emeasure_distr)
-
-lemma measure_distr:
-  "f \<in> measurable M N \<Longrightarrow> S \<in> sets N \<Longrightarrow> measure (distr M N f) S = measure M (f -` S \<inter> space M)"
-  by (simp add: emeasure_distr measure_def)
-
-lemma distr_cong_AE:
-  assumes 1: "M = K" "sets N = sets L" and
-    2: "(AE x in M. f x = g x)" and "f \<in> measurable M N" and "g \<in> measurable K L"
-  shows "distr M N f = distr K L g"
-proof (rule measure_eqI)
-  fix A assume "A \<in> sets (distr M N f)"
-  with assms show "emeasure (distr M N f) A = emeasure (distr K L g) A"
-    by (auto simp add: emeasure_distr intro!: emeasure_eq_AE measurable_sets)
-qed (insert 1, simp)
-
-lemma AE_distrD:
-  assumes f: "f \<in> measurable M M'"
-    and AE: "AE x in distr M M' f. P x"
-  shows "AE x in M. P (f x)"
-proof -
-  from AE[THEN AE_E] guess N .
-  with f show ?thesis
-    unfolding eventually_ae_filter
-    by (intro bexI[of _ "f -` N \<inter> space M"])
-       (auto simp: emeasure_distr measurable_def)
-qed
-
-lemma AE_distr_iff:
-  assumes f[measurable]: "f \<in> measurable M N" and P[measurable]: "{x \<in> space N. P x} \<in> sets N"
-  shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
-proof (subst (1 2) AE_iff_measurable[OF _ refl])
-  have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
-    using f[THEN measurable_space] by auto
-  then show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
-    (emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
-    by (simp add: emeasure_distr)
-qed auto
-
-lemma null_sets_distr_iff:
-  "f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N"
-  by (auto simp add: null_sets_def emeasure_distr)
-
-lemma distr_distr:
-  "g \<in> measurable N L \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> distr (distr M N f) L g = distr M L (g \<circ> f)"
-  by (auto simp add: emeasure_distr measurable_space
-           intro!: arg_cong[where f="emeasure M"] measure_eqI)
-
-subsection \<open>Real measure values\<close>
-
-lemma ring_of_finite_sets: "ring_of_sets (space M) {A\<in>sets M. emeasure M A \<noteq> top}"
-proof (rule ring_of_setsI)
-  show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
-    a \<union> b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
-    using emeasure_subadditive[of a M b] by (auto simp: top_unique)
-
-  show "a \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow> b \<in> {A \<in> sets M. emeasure M A \<noteq> top} \<Longrightarrow>
-    a - b \<in> {A \<in> sets M. emeasure M A \<noteq> top}" for a b
-    using emeasure_mono[of "a - b" a M] by (auto simp: Diff_subset top_unique)
-qed (auto dest: sets.sets_into_space)
-
-lemma measure_nonneg[simp]: "0 \<le> measure M A"
-  unfolding measure_def by auto
-
-lemma zero_less_measure_iff: "0 < measure M A \<longleftrightarrow> measure M A \<noteq> 0"
-  using measure_nonneg[of M A] by (auto simp add: le_less)
-
-lemma measure_le_0_iff: "measure M X \<le> 0 \<longleftrightarrow> measure M X = 0"
-  using measure_nonneg[of M X] by linarith
-
-lemma measure_empty[simp]: "measure M {} = 0"
-  unfolding measure_def by (simp add: zero_ennreal.rep_eq)
-
-lemma emeasure_eq_ennreal_measure:
-  "emeasure M A \<noteq> top \<Longrightarrow> emeasure M A = ennreal (measure M A)"
-  by (cases "emeasure M A" rule: ennreal_cases) (auto simp: measure_def)
-
-lemma measure_zero_top: "emeasure M A = top \<Longrightarrow> measure M A = 0"
-  by (simp add: measure_def enn2ereal_top)
-
-lemma measure_eq_emeasure_eq_ennreal: "0 \<le> x \<Longrightarrow> emeasure M A = ennreal x \<Longrightarrow> measure M A = x"
-  using emeasure_eq_ennreal_measure[of M A]
-  by (cases "A \<in> M") (auto simp: measure_notin_sets emeasure_notin_sets)
-
-lemma enn2real_plus:"a < top \<Longrightarrow> b < top \<Longrightarrow> enn2real (a + b) = enn2real a + enn2real b"
-  by (simp add: enn2real_def plus_ennreal.rep_eq real_of_ereal_add less_top
-           del: real_of_ereal_enn2ereal)
-
-lemma measure_Union:
-  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M B \<noteq> \<infinity> \<Longrightarrow> A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow>
-    measure M (A \<union> B) = measure M A + measure M B"
-  by (simp add: measure_def plus_emeasure[symmetric] enn2real_plus less_top)
-
-lemma disjoint_family_on_insert:
-  "i \<notin> I \<Longrightarrow> disjoint_family_on A (insert i I) \<longleftrightarrow> A i \<inter> (\<Union>i\<in>I. A i) = {} \<and> disjoint_family_on A I"
-  by (fastforce simp: disjoint_family_on_def)
-
-lemma measure_finite_Union:
-  "finite S \<Longrightarrow> A`S \<subseteq> sets M \<Longrightarrow> disjoint_family_on A S \<Longrightarrow> (\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>) \<Longrightarrow>
-    measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
-  by (induction S rule: finite_induct)
-     (auto simp: disjoint_family_on_insert measure_Union setsum_emeasure[symmetric] sets.countable_UN'[OF countable_finite])
-
-lemma measure_Diff:
-  assumes finite: "emeasure M A \<noteq> \<infinity>"
-  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
-  shows "measure M (A - B) = measure M A - measure M B"
-proof -
-  have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
-    using measurable by (auto intro!: emeasure_mono)
-  hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
-    using measurable finite by (rule_tac measure_Union) (auto simp: top_unique)
-  thus ?thesis using \<open>B \<subseteq> A\<close> by (auto simp: Un_absorb2)
-qed
-
-lemma measure_UNION:
-  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
-  assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
-proof -
-  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))"
-    unfolding suminf_emeasure[OF measurable, symmetric] by (simp add: summable_sums)
-  moreover
-  { fix i
-    have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
-      using measurable by (auto intro!: emeasure_mono)
-    then have "emeasure M (A i) = ennreal ((measure M (A i)))"
-      using finite by (intro emeasure_eq_ennreal_measure) (auto simp: top_unique) }
-  ultimately show ?thesis using finite
-    by (subst (asm) (2) emeasure_eq_ennreal_measure) simp_all
-qed
-
-lemma measure_subadditive:
-  assumes measurable: "A \<in> sets M" "B \<in> sets M"
-  and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
-  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
-proof -
-  have "emeasure M (A \<union> B) \<noteq> \<infinity>"
-    using emeasure_subadditive[OF measurable] fin by (auto simp: top_unique)
-  then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
-    using emeasure_subadditive[OF measurable] fin
-    apply simp
-    apply (subst (asm) (2 3 4) emeasure_eq_ennreal_measure)
-    apply (auto simp add: ennreal_plus[symmetric] simp del: ennreal_plus)
-    done
-qed
-
-lemma measure_subadditive_finite:
-  assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
-  shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
-proof -
-  { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
-      using emeasure_subadditive_finite[OF A] .
-    also have "\<dots> < \<infinity>"
-      using fin by (simp add: less_top A)
-    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> top" by simp }
-  note * = this
-  show ?thesis
-    using emeasure_subadditive_finite[OF A] fin
-    unfolding emeasure_eq_ennreal_measure[OF *]
-    by (simp_all add: setsum_nonneg emeasure_eq_ennreal_measure)
-qed
-
-lemma measure_subadditive_countably:
-  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
-  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
-proof -
-  from fin have **: "\<And>i. emeasure M (A i) \<noteq> top"
-    using ennreal_suminf_lessD[of "\<lambda>i. emeasure M (A i)"] by (simp add: less_top)
-  { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
-      using emeasure_subadditive_countably[OF A] .
-    also have "\<dots> < \<infinity>"
-      using fin by (simp add: less_top)
-    finally have "emeasure M (\<Union>i. A i) \<noteq> top" by simp }
-  then have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
-    by (rule emeasure_eq_ennreal_measure[symmetric])
-  also have "\<dots> \<le> (\<Sum>i. emeasure M (A i))"
-    using emeasure_subadditive_countably[OF A] .
-  also have "\<dots> = ennreal (\<Sum>i. measure M (A i))"
-    using fin unfolding emeasure_eq_ennreal_measure[OF **]
-    by (subst suminf_ennreal) (auto simp: **)
-  finally show ?thesis
-    apply (rule ennreal_le_iff[THEN iffD1, rotated])
-    apply (intro suminf_nonneg allI measure_nonneg summable_suminf_not_top)
-    using fin
-    apply (simp add: emeasure_eq_ennreal_measure[OF **])
-    done
-qed
-
-lemma measure_eq_setsum_singleton:
-  "finite S \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M) \<Longrightarrow> (\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>) \<Longrightarrow>
-    measure M S = (\<Sum>x\<in>S. measure M {x})"
-  using emeasure_eq_setsum_singleton[of S M]
-  by (intro measure_eq_emeasure_eq_ennreal) (auto simp: setsum_nonneg emeasure_eq_ennreal_measure)
-
-lemma Lim_measure_incseq:
-  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
-proof (rule tendsto_ennrealD)
-  have "ennreal (measure M (\<Union>i. A i)) = emeasure M (\<Union>i. A i)"
-    using fin by (auto simp: emeasure_eq_ennreal_measure)
-  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
-    using assms emeasure_mono[of "A _" "\<Union>i. A i" M]
-    by (intro emeasure_eq_ennreal_measure[symmetric]) (auto simp: less_top UN_upper intro: le_less_trans)
-  ultimately show "(\<lambda>x. ennreal (Sigma_Algebra.measure M (A x))) \<longlonglongrightarrow> ennreal (Sigma_Algebra.measure M (\<Union>i. A i))"
-    using A by (auto intro!: Lim_emeasure_incseq)
-qed auto
-
-lemma Lim_measure_decseq:
-  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
-  shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
-proof (rule tendsto_ennrealD)
-  have "ennreal (measure M (\<Inter>i. A i)) = emeasure M (\<Inter>i. A i)"
-    using fin[of 0] A emeasure_mono[of "\<Inter>i. A i" "A 0" M]
-    by (auto intro!: emeasure_eq_ennreal_measure[symmetric] simp: INT_lower less_top intro: le_less_trans)
-  moreover have "ennreal (measure M (A i)) = emeasure M (A i)" for i
-    using A fin[of i] by (intro emeasure_eq_ennreal_measure[symmetric]) auto
-  ultimately show "(\<lambda>x. ennreal (Sigma_Algebra.measure M (A x))) \<longlonglongrightarrow> ennreal (Sigma_Algebra.measure M (\<Inter>i. A i))"
-    using fin A by (auto intro!: Lim_emeasure_decseq)
-qed auto
-
-subsection \<open>Measure spaces with @{term "emeasure M (space M) < \<infinity>"}\<close>
-
-locale finite_measure = sigma_finite_measure M for M +
-  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> top"
-
-lemma finite_measureI[Pure.intro!]:
-  "emeasure M (space M) \<noteq> \<infinity> \<Longrightarrow> finite_measure M"
-  proof qed (auto intro!: exI[of _ "{space M}"])
-
-lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> top"
-  using finite_emeasure_space emeasure_space[of M A] by (auto simp: top_unique)
-
-lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ennreal (measure M A)"
-  by (intro emeasure_eq_ennreal_measure) simp
-
-lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ennreal r"
-  using emeasure_finite[of A] by (cases "emeasure M A" rule: ennreal_cases) auto
-
-lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
-  using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
-
-lemma (in finite_measure) finite_measure_Diff:
-  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
-  shows "measure M (A - B) = measure M A - measure M B"
-  using measure_Diff[OF _ assms] by simp
-
-lemma (in finite_measure) finite_measure_Union:
-  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
-  shows "measure M (A \<union> B) = measure M A + measure M B"
-  using measure_Union[OF _ _ assms] by simp
-
-lemma (in finite_measure) finite_measure_finite_Union:
-  assumes measurable: "finite S" "A`S \<subseteq> sets M" "disjoint_family_on A S"
-  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
-  using measure_finite_Union[OF assms] by simp
-
-lemma (in finite_measure) finite_measure_UNION:
-  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
-  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
-  using measure_UNION[OF A] by simp
-
-lemma (in finite_measure) finite_measure_mono:
-  assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
-  using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
-
-lemma (in finite_measure) finite_measure_subadditive:
-  assumes m: "A \<in> sets M" "B \<in> sets M"
-  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
-  using measure_subadditive[OF m] by simp
-
-lemma (in finite_measure) finite_measure_subadditive_finite:
-  assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
-  using measure_subadditive_finite[OF assms] by simp
-
-lemma (in finite_measure) finite_measure_subadditive_countably:
-  "range A \<subseteq> sets M \<Longrightarrow> summable (\<lambda>i. measure M (A i)) \<Longrightarrow> measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
-  by (rule measure_subadditive_countably)
-     (simp_all add: ennreal_suminf_neq_top emeasure_eq_measure)
-
-lemma (in finite_measure) finite_measure_eq_setsum_singleton:
-  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
-  using measure_eq_setsum_singleton[OF assms] by simp
-
-lemma (in finite_measure) finite_Lim_measure_incseq:
-  assumes A: "range A \<subseteq> sets M" "incseq A"
-  shows "(\<lambda>i. measure M (A i)) \<longlonglongrightarrow> measure M (\<Union>i. A i)"
-  using Lim_measure_incseq[OF A] by simp
-
-lemma (in finite_measure) finite_Lim_measure_decseq:
-  assumes A: "range A \<subseteq> sets M" "decseq A"
-  shows "(\<lambda>n. measure M (A n)) \<longlonglongrightarrow> measure M (\<Inter>i. A i)"
-  using Lim_measure_decseq[OF A] by simp
-
-lemma (in finite_measure) finite_measure_compl:
-  assumes S: "S \<in> sets M"
-  shows "measure M (space M - S) = measure M (space M) - measure M S"
-  using measure_Diff[OF _ sets.top S sets.sets_into_space] S by simp
-
-lemma (in finite_measure) finite_measure_mono_AE:
-  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
-  shows "measure M A \<le> measure M B"
-  using assms emeasure_mono_AE[OF imp B]
-  by (simp add: emeasure_eq_measure)
-
-lemma (in finite_measure) finite_measure_eq_AE:
-  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
-  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
-  shows "measure M A = measure M B"
-  using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
-
-lemma (in finite_measure) measure_increasing: "increasing M (measure M)"
-  by (auto intro!: finite_measure_mono simp: increasing_def)
-
-lemma (in finite_measure) measure_zero_union:
-  assumes "s \<in> sets M" "t \<in> sets M" "measure M t = 0"
-  shows "measure M (s \<union> t) = measure M s"
-using assms
-proof -
-  have "measure M (s \<union> t) \<le> measure M s"
-    using finite_measure_subadditive[of s t] assms by auto
-  moreover have "measure M (s \<union> t) \<ge> measure M s"
-    using assms by (blast intro: finite_measure_mono)
-  ultimately show ?thesis by simp
-qed
-
-lemma (in finite_measure) measure_eq_compl:
-  assumes "s \<in> sets M" "t \<in> sets M"
-  assumes "measure M (space M - s) = measure M (space M - t)"
-  shows "measure M s = measure M t"
-  using assms finite_measure_compl by auto
-
-lemma (in finite_measure) measure_eq_bigunion_image:
-  assumes "range f \<subseteq> sets M" "range g \<subseteq> sets M"
-  assumes "disjoint_family f" "disjoint_family g"
-  assumes "\<And> n :: nat. measure M (f n) = measure M (g n)"
-  shows "measure M (\<Union>i. f i) = measure M (\<Union>i. g i)"
-using assms
-proof -
-  have a: "(\<lambda> i. measure M (f i)) sums (measure M (\<Union>i. f i))"
-    by (rule finite_measure_UNION[OF assms(1,3)])
-  have b: "(\<lambda> i. measure M (g i)) sums (measure M (\<Union>i. g i))"
-    by (rule finite_measure_UNION[OF assms(2,4)])
-  show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
-qed
-
-lemma (in finite_measure) measure_countably_zero:
-  assumes "range c \<subseteq> sets M"
-  assumes "\<And> i. measure M (c i) = 0"
-  shows "measure M (\<Union>i :: nat. c i) = 0"
-proof (rule antisym)
-  show "measure M (\<Union>i :: nat. c i) \<le> 0"
-    using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
-qed simp
-
-lemma (in finite_measure) measure_space_inter:
-  assumes events:"s \<in> sets M" "t \<in> sets M"
-  assumes "measure M t = measure M (space M)"
-  shows "measure M (s \<inter> t) = measure M s"
-proof -
-  have "measure M ((space M - s) \<union> (space M - t)) = measure M (space M - s)"
-    using events assms finite_measure_compl[of "t"] by (auto intro!: measure_zero_union)
-  also have "(space M - s) \<union> (space M - t) = space M - (s \<inter> t)"
-    by blast
-  finally show "measure M (s \<inter> t) = measure M s"
-    using events by (auto intro!: measure_eq_compl[of "s \<inter> t" s])
-qed
-
-lemma (in finite_measure) measure_equiprobable_finite_unions:
-  assumes s: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> sets M"
-  assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> measure M {x} = measure M {y}"
-  shows "measure M s = real (card s) * measure M {SOME x. x \<in> s}"
-proof cases
-  assume "s \<noteq> {}"
-  then have "\<exists> x. x \<in> s" by blast
-  from someI_ex[OF this] assms
-  have prob_some: "\<And> x. x \<in> s \<Longrightarrow> measure M {x} = measure M {SOME y. y \<in> s}" by blast
-  have "measure M s = (\<Sum> x \<in> s. measure M {x})"
-    using finite_measure_eq_setsum_singleton[OF s] by simp
-  also have "\<dots> = (\<Sum> x \<in> s. measure M {SOME y. y \<in> s})" using prob_some by auto
-  also have "\<dots> = real (card s) * measure M {(SOME x. x \<in> s)}"
-    using setsum_constant assms by simp
-  finally show ?thesis by simp
-qed simp
-
-lemma (in finite_measure) measure_real_sum_image_fn:
-  assumes "e \<in> sets M"
-  assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> sets M"
-  assumes "finite s"
-  assumes disjoint: "\<And> x y. \<lbrakk>x \<in> s ; y \<in> s ; x \<noteq> y\<rbrakk> \<Longrightarrow> f x \<inter> f y = {}"
-  assumes upper: "space M \<subseteq> (\<Union>i \<in> s. f i)"
-  shows "measure M e = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
-proof -
-  have "e \<subseteq> (\<Union>i\<in>s. f i)"
-    using \<open>e \<in> sets M\<close> sets.sets_into_space upper by blast
-  then have e: "e = (\<Union>i \<in> s. e \<inter> f i)"
-    by auto
-  hence "measure M e = measure M (\<Union>i \<in> s. e \<inter> f i)" by simp
-  also have "\<dots> = (\<Sum> x \<in> s. measure M (e \<inter> f x))"
-  proof (rule finite_measure_finite_Union)
-    show "finite s" by fact
-    show "(\<lambda>i. e \<inter> f i)`s \<subseteq> sets M" using assms(2) by auto
-    show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
-      using disjoint by (auto simp: disjoint_family_on_def)
-  qed
-  finally show ?thesis .
-qed
-
-lemma (in finite_measure) measure_exclude:
-  assumes "A \<in> sets M" "B \<in> sets M"
-  assumes "measure M A = measure M (space M)" "A \<inter> B = {}"
-  shows "measure M B = 0"
-  using measure_space_inter[of B A] assms by (auto simp: ac_simps)
-lemma (in finite_measure) finite_measure_distr:
-  assumes f: "f \<in> measurable M M'"
-  shows "finite_measure (distr M M' f)"
-proof (rule finite_measureI)
-  have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
-  with f show "emeasure (distr M M' f) (space (distr M M' f)) \<noteq> \<infinity>" by (auto simp: emeasure_distr)
-qed
-
-lemma emeasure_gfp[consumes 1, case_names cont measurable]:
-  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
-  assumes "\<And>s. finite_measure (M s)"
-  assumes cont: "inf_continuous F" "inf_continuous f"
-  assumes meas: "\<And>P. Measurable.pred N P \<Longrightarrow> Measurable.pred N (F P)"
-  assumes iter: "\<And>P s. Measurable.pred N P \<Longrightarrow> emeasure (M s) {x\<in>space N. F P x} = f (\<lambda>s. emeasure (M s) {x\<in>space N. P x}) s"
-  assumes bound: "\<And>P. f P \<le> f (\<lambda>s. emeasure (M s) (space (M s)))"
-  shows "emeasure (M s) {x\<in>space N. gfp F x} = gfp f s"
-proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. emeasure (M s) {x\<in>space N. F x}" and g=f and f=F and
-    P="Measurable.pred N", symmetric])
-  interpret finite_measure "M s" for s by fact
-  fix C assume "decseq C" "\<And>i. Measurable.pred N (C i)"
-  then show "(\<lambda>s. emeasure (M s) {x \<in> space N. (INF i. C i) x}) = (INF i. (\<lambda>s. emeasure (M s) {x \<in> space N. C i x}))"
-    unfolding INF_apply[abs_def]
-    by (subst INF_emeasure_decseq) (auto simp: antimono_def fun_eq_iff intro!: arg_cong2[where f=emeasure])
-next
-  show "f x \<le> (\<lambda>s. emeasure (M s) {x \<in> space N. F top x})" for x
-    using bound[of x] sets_eq_imp_space_eq[OF sets] by (simp add: iter)
-qed (auto simp add: iter le_fun_def INF_apply[abs_def] intro!: meas cont)
-
-subsection \<open>Counting space\<close>
-
-lemma strict_monoI_Suc:
-  assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
-  unfolding strict_mono_def
-proof safe
-  fix n m :: nat assume "n < m" then show "f n < f m"
-    by (induct m) (auto simp: less_Suc_eq intro: less_trans ord)
-qed
-
-lemma emeasure_count_space:
-  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then of_nat (card X) else \<infinity>)"
-    (is "_ = ?M X")
-  unfolding count_space_def
-proof (rule emeasure_measure_of_sigma)
-  show "X \<in> Pow A" using \<open>X \<subseteq> A\<close> by auto
-  show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
-  show positive: "positive (Pow A) ?M"
-    by (auto simp: positive_def)
-  have additive: "additive (Pow A) ?M"
-    by (auto simp: card_Un_disjoint additive_def)
-
-  interpret ring_of_sets A "Pow A"
-    by (rule ring_of_setsI) auto
-  show "countably_additive (Pow A) ?M"
-    unfolding countably_additive_iff_continuous_from_below[OF positive additive]
-  proof safe
-    fix F :: "nat \<Rightarrow> 'a set" assume "incseq F"
-    show "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> ?M (\<Union>i. F i)"
-    proof cases
-      assume "\<exists>i. \<forall>j\<ge>i. F i = F j"
-      then guess i .. note i = this
-      { fix j from i \<open>incseq F\<close> have "F j \<subseteq> F i"
-          by (cases "i \<le> j") (auto simp: incseq_def) }
-      then have eq: "(\<Union>i. F i) = F i"
-        by auto
-      with i show ?thesis
-        by (auto intro!: Lim_eventually eventually_sequentiallyI[where c=i])
-    next
-      assume "\<not> (\<exists>i. \<forall>j\<ge>i. F i = F j)"
-      then obtain f where f: "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis
-      then have "\<And>i. F i \<subseteq> F (f i)" using \<open>incseq F\<close> by (auto simp: incseq_def)
-      with f have *: "\<And>i. F i \<subset> F (f i)" by auto
-
-      have "incseq (\<lambda>i. ?M (F i))"
-        using \<open>incseq F\<close> unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
-      then have "(\<lambda>i. ?M (F i)) \<longlonglongrightarrow> (SUP n. ?M (F n))"
-        by (rule LIMSEQ_SUP)
-
-      moreover have "(SUP n. ?M (F n)) = top"
-      proof (rule ennreal_SUP_eq_top)
-        fix n :: nat show "\<exists>k::nat\<in>UNIV. of_nat n \<le> ?M (F k)"
-        proof (induct n)
-          case (Suc n)
-          then guess k .. note k = this
-          moreover have "finite (F k) \<Longrightarrow> finite (F (f k)) \<Longrightarrow> card (F k) < card (F (f k))"
-            using \<open>F k \<subset> F (f k)\<close> by (simp add: psubset_card_mono)
-          moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
-            using \<open>k \<le> f k\<close> \<open>incseq F\<close> by (auto simp: incseq_def dest: finite_subset)
-          ultimately show ?case
-            by (auto intro!: exI[of _ "f k"] simp del: of_nat_Suc)
-        qed auto
-      qed
-
-      moreover
-      have "inj (\<lambda>n. F ((f ^^ n) 0))"
-        by (intro strict_mono_imp_inj_on strict_monoI_Suc) (simp add: *)
-      then have 1: "infinite (range (\<lambda>i. F ((f ^^ i) 0)))"
-        by (rule range_inj_infinite)
-      have "infinite (Pow (\<Union>i. F i))"
-        by (rule infinite_super[OF _ 1]) auto
-      then have "infinite (\<Union>i. F i)"
-        by auto
-
-      ultimately show ?thesis by auto
-    qed
-  qed
-qed
-
-lemma distr_bij_count_space:
-  assumes f: "bij_betw f A B"
-  shows "distr (count_space A) (count_space B) f = count_space B"
-proof (rule measure_eqI)
-  have f': "f \<in> measurable (count_space A) (count_space B)"
-    using f unfolding Pi_def bij_betw_def by auto
-  fix X assume "X \<in> sets (distr (count_space A) (count_space B) f)"
-  then have X: "X \<in> sets (count_space B)" by auto
-  moreover from X have "f -` X \<inter> A = the_inv_into A f ` X"
-    using f by (auto simp: bij_betw_def subset_image_iff image_iff the_inv_into_f_f intro: the_inv_into_f_f[symmetric])
-  moreover have "inj_on (the_inv_into A f) B"
-    using X f by (auto simp: bij_betw_def inj_on_the_inv_into)
-  with X have "inj_on (the_inv_into A f) X"
-    by (auto intro: subset_inj_on)
-  ultimately show "emeasure (distr (count_space A) (count_space B) f) X = emeasure (count_space B) X"
-    using f unfolding emeasure_distr[OF f' X]
-    by (subst (1 2) emeasure_count_space) (auto simp: card_image dest: finite_imageD)
-qed simp
-
-lemma emeasure_count_space_finite[simp]:
-  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = of_nat (card X)"
-  using emeasure_count_space[of X A] by simp
-
-lemma emeasure_count_space_infinite[simp]:
-  "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
-  using emeasure_count_space[of X A] by simp
-
-lemma measure_count_space: "measure (count_space A) X = (if X \<subseteq> A then of_nat (card X) else 0)"
-  by (cases "finite X") (auto simp: measure_notin_sets ennreal_of_nat_eq_real_of_nat
-                                    measure_zero_top measure_eq_emeasure_eq_ennreal)
-
-lemma emeasure_count_space_eq_0:
-  "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
-proof cases
-  assume X: "X \<subseteq> A"
-  then show ?thesis
-  proof (intro iffI impI)
-    assume "emeasure (count_space A) X = 0"
-    with X show "X = {}"
-      by (subst (asm) emeasure_count_space) (auto split: if_split_asm)
-  qed simp
-qed (simp add: emeasure_notin_sets)
-
-lemma space_empty: "space M = {} \<Longrightarrow> M = count_space {}"
-  by (rule measure_eqI) (simp_all add: space_empty_iff)
-
-lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
-  unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
-
-lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
-  unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
-
-lemma sigma_finite_measure_count_space_countable:
-  assumes A: "countable A"
-  shows "sigma_finite_measure (count_space A)"
-  proof qed (insert A, auto intro!: exI[of _ "(\<lambda>a. {a}) ` A"])
-
-lemma sigma_finite_measure_count_space:
-  fixes A :: "'a::countable set" shows "sigma_finite_measure (count_space A)"
-  by (rule sigma_finite_measure_count_space_countable) auto
-
-lemma finite_measure_count_space:
-  assumes [simp]: "finite A"
-  shows "finite_measure (count_space A)"
-  by rule simp
-
-lemma sigma_finite_measure_count_space_finite:
-  assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
-proof -
-  interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
-  show "sigma_finite_measure (count_space A)" ..
-qed
-
-subsection \<open>Measure restricted to space\<close>
-
-lemma emeasure_restrict_space:
-  assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
-  shows "emeasure (restrict_space M \<Omega>) A = emeasure M A"
-proof (cases "A \<in> sets M")
-  case True
-  show ?thesis
-  proof (rule emeasure_measure_of[OF restrict_space_def])
-    show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)" "A \<in> sets (restrict_space M \<Omega>)"
-      using \<open>A \<subseteq> \<Omega>\<close> \<open>A \<in> sets M\<close> sets.space_closed by (auto simp: sets_restrict_space)
-    show "positive (sets (restrict_space M \<Omega>)) (emeasure M)"
-      by (auto simp: positive_def)
-    show "countably_additive (sets (restrict_space M \<Omega>)) (emeasure M)"
-    proof (rule countably_additiveI)
-      fix A :: "nat \<Rightarrow> _" assume "range A \<subseteq> sets (restrict_space M \<Omega>)" "disjoint_family A"
-      with assms have "\<And>i. A i \<in> sets M" "\<And>i. A i \<subseteq> space M" "disjoint_family A"
-        by (fastforce simp: sets_restrict_space_iff[OF assms(1)] image_subset_iff
-                      dest: sets.sets_into_space)+
-      then show "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
-        by (subst suminf_emeasure) (auto simp: disjoint_family_subset)
-    qed
-  qed
-next
-  case False
-  with assms have "A \<notin> sets (restrict_space M \<Omega>)"
-    by (simp add: sets_restrict_space_iff)
-  with False show ?thesis
-    by (simp add: emeasure_notin_sets)
-qed
-
-lemma measure_restrict_space:
-  assumes "\<Omega> \<inter> space M \<in> sets M" "A \<subseteq> \<Omega>"
-  shows "measure (restrict_space M \<Omega>) A = measure M A"
-  using emeasure_restrict_space[OF assms] by (simp add: measure_def)
-
-lemma AE_restrict_space_iff:
-  assumes "\<Omega> \<inter> space M \<in> sets M"
-  shows "(AE x in restrict_space M \<Omega>. P x) \<longleftrightarrow> (AE x in M. x \<in> \<Omega> \<longrightarrow> P x)"
-proof -
-  have ex_cong: "\<And>P Q f. (\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> (\<And>x. Q x \<Longrightarrow> P (f x)) \<Longrightarrow> (\<exists>x. P x) \<longleftrightarrow> (\<exists>x. Q x)"
-    by auto
-  { fix X assume X: "X \<in> sets M" "emeasure M X = 0"
-    then have "emeasure M (\<Omega> \<inter> space M \<inter> X) \<le> emeasure M X"
-      by (intro emeasure_mono) auto
-    then have "emeasure M (\<Omega> \<inter> space M \<inter> X) = 0"
-      using X by (auto intro!: antisym) }
-  with assms show ?thesis
-    unfolding eventually_ae_filter
-    by (auto simp add: space_restrict_space null_sets_def sets_restrict_space_iff
-                       emeasure_restrict_space cong: conj_cong
-             intro!: ex_cong[where f="\<lambda>X. (\<Omega> \<inter> space M) \<inter> X"])
-qed
-
-lemma restrict_restrict_space:
-  assumes "A \<inter> space M \<in> sets M" "B \<inter> space M \<in> sets M"
-  shows "restrict_space (restrict_space M A) B = restrict_space M (A \<inter> B)" (is "?l = ?r")
-proof (rule measure_eqI[symmetric])
-  show "sets ?r = sets ?l"
-    unfolding sets_restrict_space image_comp by (intro image_cong) auto
-next
-  fix X assume "X \<in> sets (restrict_space M (A \<inter> B))"
-  then obtain Y where "Y \<in> sets M" "X = Y \<inter> A \<inter> B"
-    by (auto simp: sets_restrict_space)
-  with assms sets.Int[OF assms] show "emeasure ?r X = emeasure ?l X"
-    by (subst (1 2) emeasure_restrict_space)
-       (auto simp: space_restrict_space sets_restrict_space_iff emeasure_restrict_space ac_simps)
-qed
-
-lemma restrict_count_space: "restrict_space (count_space B) A = count_space (A \<inter> B)"
-proof (rule measure_eqI)
-  show "sets (restrict_space (count_space B) A) = sets (count_space (A \<inter> B))"
-    by (subst sets_restrict_space) auto
-  moreover fix X assume "X \<in> sets (restrict_space (count_space B) A)"
-  ultimately have "X \<subseteq> A \<inter> B" by auto
-  then show "emeasure (restrict_space (count_space B) A) X = emeasure (count_space (A \<inter> B)) X"
-    by (cases "finite X") (auto simp add: emeasure_restrict_space)
-qed
-
-lemma sigma_finite_measure_restrict_space:
-  assumes "sigma_finite_measure M"
-  and A: "A \<in> sets M"
-  shows "sigma_finite_measure (restrict_space M A)"
-proof -
-  interpret sigma_finite_measure M by fact
-  from sigma_finite_countable obtain C
-    where C: "countable C" "C \<subseteq> sets M" "(\<Union>C) = space M" "\<forall>a\<in>C. emeasure M a \<noteq> \<infinity>"
-    by blast
-  let ?C = "op \<inter> A ` C"
-  from C have "countable ?C" "?C \<subseteq> sets (restrict_space M A)" "(\<Union>?C) = space (restrict_space M A)"
-    by(auto simp add: sets_restrict_space space_restrict_space)
-  moreover {
-    fix a
-    assume "a \<in> ?C"
-    then obtain a' where "a = A \<inter> a'" "a' \<in> C" ..
-    then have "emeasure (restrict_space M A) a \<le> emeasure M a'"
-      using A C by(auto simp add: emeasure_restrict_space intro: emeasure_mono)
-    also have "\<dots> < \<infinity>" using C(4)[rule_format, of a'] \<open>a' \<in> C\<close> by (simp add: less_top)
-    finally have "emeasure (restrict_space M A) a \<noteq> \<infinity>" by simp }
-  ultimately show ?thesis
-    by unfold_locales (rule exI conjI|assumption|blast)+
-qed
-
-lemma finite_measure_restrict_space:
-  assumes "finite_measure M"
-  and A: "A \<in> sets M"
-  shows "finite_measure (restrict_space M A)"
-proof -
-  interpret finite_measure M by fact
-  show ?thesis
-    by(rule finite_measureI)(simp add: emeasure_restrict_space A space_restrict_space)
-qed
-
-lemma restrict_distr:
-  assumes [measurable]: "f \<in> measurable M N"
-  assumes [simp]: "\<Omega> \<inter> space N \<in> sets N" and restrict: "f \<in> space M \<rightarrow> \<Omega>"
-  shows "restrict_space (distr M N f) \<Omega> = distr M (restrict_space N \<Omega>) f"
-  (is "?l = ?r")
-proof (rule measure_eqI)
-  fix A assume "A \<in> sets (restrict_space (distr M N f) \<Omega>)"
-  with restrict show "emeasure ?l A = emeasure ?r A"
-    by (subst emeasure_distr)
-       (auto simp: sets_restrict_space_iff emeasure_restrict_space emeasure_distr
-             intro!: measurable_restrict_space2)
-qed (simp add: sets_restrict_space)
-
-lemma measure_eqI_restrict_generator:
-  assumes E: "Int_stable E" "E \<subseteq> Pow \<Omega>" "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
-  assumes sets_eq: "sets M = sets N" and \<Omega>: "\<Omega> \<in> sets M"
-  assumes "sets (restrict_space M \<Omega>) = sigma_sets \<Omega> E"
-  assumes "sets (restrict_space N \<Omega>) = sigma_sets \<Omega> E"
-  assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>"
-  assumes A: "countable A" "A \<noteq> {}" "A \<subseteq> E" "\<Union>A = \<Omega>" "\<And>a. a \<in> A \<Longrightarrow> emeasure M a \<noteq> \<infinity>"
-  shows "M = N"
-proof (rule measure_eqI)
-  fix X assume X: "X \<in> sets M"
-  then have "emeasure M X = emeasure (restrict_space M \<Omega>) (X \<inter> \<Omega>)"
-    using ae \<Omega> by (auto simp add: emeasure_restrict_space intro!: emeasure_eq_AE)
-  also have "restrict_space M \<Omega> = restrict_space N \<Omega>"
-  proof (rule measure_eqI_generator_eq)
-    fix X assume "X \<in> E"
-    then show "emeasure (restrict_space M \<Omega>) X = emeasure (restrict_space N \<Omega>) X"
-      using E \<Omega> by (subst (1 2) emeasure_restrict_space) (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq])
-  next
-    show "range (from_nat_into A) \<subseteq> E" "(\<Union>i. from_nat_into A i) = \<Omega>"
-      using A by (auto cong del: strong_SUP_cong)
-  next
-    fix i
-    have "emeasure (restrict_space M \<Omega>) (from_nat_into A i) = emeasure M (from_nat_into A i)"
-      using A \<Omega> by (subst emeasure_restrict_space)
-                   (auto simp: sets_eq sets_eq[THEN sets_eq_imp_space_eq] intro: from_nat_into)
-    with A show "emeasure (restrict_space M \<Omega>) (from_nat_into A i) \<noteq> \<infinity>"
-      by (auto intro: from_nat_into)
-  qed fact+
-  also have "emeasure (restrict_space N \<Omega>) (X \<inter> \<Omega>) = emeasure N X"
-    using X ae \<Omega> by (auto simp add: emeasure_restrict_space sets_eq intro!: emeasure_eq_AE)
-  finally show "emeasure M X = emeasure N X" .
-qed fact
-
-subsection \<open>Null measure\<close>
-
-definition "null_measure M = sigma (space M) (sets M)"
-
-lemma space_null_measure[simp]: "space (null_measure M) = space M"
-  by (simp add: null_measure_def)
-
-lemma sets_null_measure[simp, measurable_cong]: "sets (null_measure M) = sets M"
-  by (simp add: null_measure_def)
-
-lemma emeasure_null_measure[simp]: "emeasure (null_measure M) X = 0"
-  by (cases "X \<in> sets M", rule emeasure_measure_of)
-     (auto simp: positive_def countably_additive_def emeasure_notin_sets null_measure_def
-           dest: sets.sets_into_space)
-
-lemma measure_null_measure[simp]: "measure (null_measure M) X = 0"
-  by (intro measure_eq_emeasure_eq_ennreal) auto
-
-lemma null_measure_idem [simp]: "null_measure (null_measure M) = null_measure M"
-  by(rule measure_eqI) simp_all
-
-subsection \<open>Scaling a measure\<close>
-
-definition scale_measure :: "ennreal \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
-where
-  "scale_measure r M = measure_of (space M) (sets M) (\<lambda>A. r * emeasure M A)"
-
-lemma space_scale_measure: "space (scale_measure r M) = space M"
-  by (simp add: scale_measure_def)
-
-lemma sets_scale_measure [simp, measurable_cong]: "sets (scale_measure r M) = sets M"
-  by (simp add: scale_measure_def)
-
-lemma emeasure_scale_measure [simp]:
-  "emeasure (scale_measure r M) A = r * emeasure M A"
-  (is "_ = ?\<mu> A")
-proof(cases "A \<in> sets M")
-  case True
-  show ?thesis unfolding scale_measure_def
-  proof(rule emeasure_measure_of_sigma)
-    show "sigma_algebra (space M) (sets M)" ..
-    show "positive (sets M) ?\<mu>" by (simp add: positive_def)
-    show "countably_additive (sets M) ?\<mu>"
-    proof (rule countably_additiveI)
-      fix A :: "nat \<Rightarrow> _"  assume *: "range A \<subseteq> sets M" "disjoint_family A"
-      have "(\<Sum>i. ?\<mu> (A i)) = r * (\<Sum>i. emeasure M (A i))"
-        by simp
-      also have "\<dots> = ?\<mu> (\<Union>i. A i)" using * by(simp add: suminf_emeasure)
-      finally show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)" .
-    qed
-  qed(fact True)
-qed(simp add: emeasure_notin_sets)
-
-lemma scale_measure_1 [simp]: "scale_measure 1 M = M"
-  by(rule measure_eqI) simp_all
-
-lemma scale_measure_0[simp]: "scale_measure 0 M = null_measure M"
-  by(rule measure_eqI) simp_all
-
-lemma measure_scale_measure [simp]: "0 \<le> r \<Longrightarrow> measure (scale_measure r M) A = r * measure M A"
-  using emeasure_scale_measure[of r M A]
-    emeasure_eq_ennreal_measure[of M A]
-    measure_eq_emeasure_eq_ennreal[of _ "scale_measure r M" A]
-  by (cases "emeasure (scale_measure r M) A = top")
-     (auto simp del: emeasure_scale_measure
-           simp: ennreal_top_eq_mult_iff ennreal_mult_eq_top_iff measure_zero_top ennreal_mult[symmetric])
-
-lemma scale_scale_measure [simp]:
-  "scale_measure r (scale_measure r' M) = scale_measure (r * r') M"
-  by (rule measure_eqI) (simp_all add: max_def mult.assoc)
-
-lemma scale_null_measure [simp]: "scale_measure r (null_measure M) = null_measure M"
-  by (rule measure_eqI) simp_all
-
-
-subsection \<open>Complete lattice structure on measures\<close>
-
-lemma (in finite_measure) finite_measure_Diff':
-  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A - B) = measure M A - measure M (A \<inter> B)"
-  using finite_measure_Diff[of A "A \<inter> B"] by (auto simp: Diff_Int)
-
-lemma (in finite_measure) finite_measure_Union':
-  "A \<in> sets M \<Longrightarrow> B \<in> sets M \<Longrightarrow> measure M (A \<union> B) = measure M A + measure M (B - A)"
-  using finite_measure_Union[of A "B - A"] by auto
-
-lemma finite_unsigned_Hahn_decomposition:
-  assumes "finite_measure M" "finite_measure N" and [simp]: "sets N = sets M"
-  shows "\<exists>Y\<in>sets M. (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
-proof -
-  interpret M: finite_measure M by fact
-  interpret N: finite_measure N by fact
-
-  define d where "d X = measure M X - measure N X" for X
-
-  have [intro]: "bdd_above (d`sets M)"
-    using sets.sets_into_space[of _ M]
-    by (intro bdd_aboveI[where M="measure M (space M)"])
-       (auto simp: d_def field_simps subset_eq intro!: add_increasing M.finite_measure_mono)
-
-  define \<gamma> where "\<gamma> = (SUP X:sets M. d X)"
-  have le_\<gamma>[intro]: "X \<in> sets M \<Longrightarrow> d X \<le> \<gamma>" for X
-    by (auto simp: \<gamma>_def intro!: cSUP_upper)
-
-  have "\<exists>f. \<forall>n. f n \<in> sets M \<and> d (f n) > \<gamma> - 1 / 2^n"
-  proof (intro choice_iff[THEN iffD1] allI)
-    fix n
-    have "\<exists>X\<in>sets M. \<gamma> - 1 / 2^n < d X"
-      unfolding \<gamma>_def by (intro less_cSUP_iff[THEN iffD1]) auto
-    then show "\<exists>y. y \<in> sets M \<and> \<gamma> - 1 / 2 ^ n < d y"
-      by auto
-  qed
-  then obtain E where [measurable]: "E n \<in> sets M" and E: "d (E n) > \<gamma> - 1 / 2^n" for n
-    by auto
-
-  define F where "F m n = (if m \<le> n then \<Inter>i\<in>{m..n}. E i else space M)" for m n
-
-  have [measurable]: "m \<le> n \<Longrightarrow> F m n \<in> sets M" for m n
-    by (auto simp: F_def)
-
-  have 1: "\<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)" if "m \<le> n" for m n
-    using that
-  proof (induct rule: dec_induct)
-    case base with E[of m] show ?case
-      by (simp add: F_def field_simps)
-  next
-    case (step i)
-    have F_Suc: "F m (Suc i) = F m i \<inter> E (Suc i)"
-      using \<open>m \<le> i\<close> by (auto simp: F_def le_Suc_eq)
-
-    have "\<gamma> + (\<gamma> - 2 / 2^m + 1 / 2 ^ Suc i) \<le> (\<gamma> - 1 / 2^Suc i) + (\<gamma> - 2 / 2^m + 1 / 2^i)"
-      by (simp add: field_simps)
-    also have "\<dots> \<le> d (E (Suc i)) + d (F m i)"
-      using E[of "Suc i"] by (intro add_mono step) auto
-    also have "\<dots> = d (E (Suc i)) + d (F m i - E (Suc i)) + d (F m (Suc i))"
-      using \<open>m \<le> i\<close> by (simp add: d_def field_simps F_Suc M.finite_measure_Diff' N.finite_measure_Diff')
-    also have "\<dots> = d (E (Suc i) \<union> F m i) + d (F m (Suc i))"
-      using \<open>m \<le> i\<close> by (simp add: d_def field_simps M.finite_measure_Union' N.finite_measure_Union')
-    also have "\<dots> \<le> \<gamma> + d (F m (Suc i))"
-      using \<open>m \<le> i\<close> by auto
-    finally show ?case
-      by auto
-  qed
-
-  define F' where "F' m = (\<Inter>i\<in>{m..}. E i)" for m
-  have F'_eq: "F' m = (\<Inter>i. F m (i + m))" for m
-    by (fastforce simp: le_iff_add[of m] F'_def F_def)
-
-  have [measurable]: "F' m \<in> sets M" for m
-    by (auto simp: F'_def)
-
-  have \<gamma>_le: "\<gamma> - 0 \<le> d (\<Union>m. F' m)"
-  proof (rule LIMSEQ_le)
-    show "(\<lambda>n. \<gamma> - 2 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 0"
-      by (intro tendsto_intros LIMSEQ_divide_realpow_zero) auto
-    have "incseq F'"
-      by (auto simp: incseq_def F'_def)
-    then show "(\<lambda>m. d (F' m)) \<longlonglongrightarrow> d (\<Union>m. F' m)"
-      unfolding d_def
-      by (intro tendsto_diff M.finite_Lim_measure_incseq N.finite_Lim_measure_incseq) auto
-
-    have "\<gamma> - 2 / 2 ^ m + 0 \<le> d (F' m)" for m
-    proof (rule LIMSEQ_le)
-      have *: "decseq (\<lambda>n. F m (n + m))"
-        by (auto simp: decseq_def F_def)
-      show "(\<lambda>n. d (F m n)) \<longlonglongrightarrow> d (F' m)"
-        unfolding d_def F'_eq
-        by (rule LIMSEQ_offset[where k=m])
-           (auto intro!: tendsto_diff M.finite_Lim_measure_decseq N.finite_Lim_measure_decseq *)
-      show "(\<lambda>n. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n) \<longlonglongrightarrow> \<gamma> - 2 / 2 ^ m + 0"
-        by (intro tendsto_add LIMSEQ_divide_realpow_zero tendsto_const) auto
-      show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ m + 1 / 2 ^ n \<le> d (F m n)"
-        using 1[of m] by (intro exI[of _ m]) auto
-    qed
-    then show "\<exists>N. \<forall>n\<ge>N. \<gamma> - 2 / 2 ^ n \<le> d (F' n)"
-      by auto
-  qed
-
-  show ?thesis
-  proof (safe intro!: bexI[of _ "\<Union>m. F' m"])
-    fix X assume [measurable]: "X \<in> sets M" and X: "X \<subseteq> (\<Union>m. F' m)"
-    have "d (\<Union>m. F' m) - d X = d ((\<Union>m. F' m) - X)"
-      using X by (auto simp: d_def M.finite_measure_Diff N.finite_measure_Diff)
-    also have "\<dots> \<le> \<gamma>"
-      by auto
-    finally have "0 \<le> d X"
-      using \<gamma>_le by auto
-    then show "emeasure N X \<le> emeasure M X"
-      by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
-  next
-    fix X assume [measurable]: "X \<in> sets M" and X: "X \<inter> (\<Union>m. F' m) = {}"
-    then have "d (\<Union>m. F' m) + d X = d (X \<union> (\<Union>m. F' m))"
-      by (auto simp: d_def M.finite_measure_Union N.finite_measure_Union)
-    also have "\<dots> \<le> \<gamma>"
-      by auto
-    finally have "d X \<le> 0"
-      using \<gamma>_le by auto
-    then show "emeasure M X \<le> emeasure N X"
-      by (auto simp: d_def M.emeasure_eq_measure N.emeasure_eq_measure)
-  qed auto
-qed
-
-lemma unsigned_Hahn_decomposition:
-  assumes [simp]: "sets N = sets M" and [measurable]: "A \<in> sets M"
-    and [simp]: "emeasure M A \<noteq> top" "emeasure N A \<noteq> top"
-  shows "\<exists>Y\<in>sets M. Y \<subseteq> A \<and> (\<forall>X\<in>sets M. X \<subseteq> Y \<longrightarrow> N X \<le> M X) \<and> (\<forall>X\<in>sets M. X \<subseteq> A \<longrightarrow> X \<inter> Y = {} \<longrightarrow> M X \<le> N X)"
-proof -
-  have "\<exists>Y\<in>sets (restrict_space M A).
-    (\<forall>X\<in>sets (restrict_space M A). X \<subseteq> Y \<longrightarrow> (restrict_space N A) X \<le> (restrict_space M A) X) \<and>
-    (\<forall>X\<in>sets (restrict_space M A). X \<inter> Y = {} \<longrightarrow> (restrict_space M A) X \<le> (restrict_space N A) X)"
-  proof (rule finite_unsigned_Hahn_decomposition)
-    show "finite_measure (restrict_space M A)" "finite_measure (restrict_space N A)"
-      by (auto simp: space_restrict_space emeasure_restrict_space less_top intro!: finite_measureI)
-  qed (simp add: sets_restrict_space)
-  then guess Y ..
-  then show ?thesis
-    apply (intro bexI[of _ Y] conjI ballI conjI)
-    apply (simp_all add: sets_restrict_space emeasure_restrict_space)
-    apply safe
-    subgoal for X Z
-      by (erule ballE[of _ _ X]) (auto simp add: Int_absorb1)
-    subgoal for X Z
-      by (erule ballE[of _ _ X])  (auto simp add: Int_absorb1 ac_simps)
-    apply auto
-    done
-qed
-
-text \<open>
-  Define a lexicographical order on @{type measure}, in the order space, sets and measure. The parts
-  of the lexicographical order are point-wise ordered.
-\<close>
-
-instantiation measure :: (type) order_bot
-begin
-
-inductive less_eq_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
-  "space M \<subset> space N \<Longrightarrow> less_eq_measure M N"
-| "space M = space N \<Longrightarrow> sets M \<subset> sets N \<Longrightarrow> less_eq_measure M N"
-| "space M = space N \<Longrightarrow> sets M = sets N \<Longrightarrow> emeasure M \<le> emeasure N \<Longrightarrow> less_eq_measure M N"
-
-lemma le_measure_iff:
-  "M \<le> N \<longleftrightarrow> (if space M = space N then
-    if sets M = sets N then emeasure M \<le> emeasure N else sets M \<subseteq> sets N else space M \<subseteq> space N)"
-  by (auto elim: less_eq_measure.cases intro: less_eq_measure.intros)
-
-definition less_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
-  "less_measure M N \<longleftrightarrow> (M \<le> N \<and> \<not> N \<le> M)"
-
-definition bot_measure :: "'a measure" where
-  "bot_measure = sigma {} {}"
-
-lemma
-  shows space_bot[simp]: "space bot = {}"
-    and sets_bot[simp]: "sets bot = {{}}"
-    and emeasure_bot[simp]: "emeasure bot X = 0"
-  by (auto simp: bot_measure_def sigma_sets_empty_eq emeasure_sigma)
-
-instance
-proof standard
-  show "bot \<le> a" for a :: "'a measure"
-    by (simp add: le_measure_iff bot_measure_def sigma_sets_empty_eq emeasure_sigma le_fun_def)
-qed (auto simp: le_measure_iff less_measure_def split: if_split_asm intro: measure_eqI)
-
-end
-
-lemma le_measure: "sets M = sets N \<Longrightarrow> M \<le> N \<longleftrightarrow> (\<forall>A\<in>sets M. emeasure M A \<le> emeasure N A)"
-  apply (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq)
-  subgoal for X
-    by (cases "X \<in> sets M") (auto simp: emeasure_notin_sets)
-  done
-
-definition sup_measure' :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
-where
-  "sup_measure' A B = measure_of (space A) (sets A) (\<lambda>X. SUP Y:sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
-
-lemma assumes [simp]: "sets B = sets A"
-  shows space_sup_measure'[simp]: "space (sup_measure' A B) = space A"
-    and sets_sup_measure'[simp]: "sets (sup_measure' A B) = sets A"
-  using sets_eq_imp_space_eq[OF assms] by (simp_all add: sup_measure'_def)
-
-lemma emeasure_sup_measure':
-  assumes sets_eq[simp]: "sets B = sets A" and [simp, intro]: "X \<in> sets A"
-  shows "emeasure (sup_measure' A B) X = (SUP Y:sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
-    (is "_ = ?S X")
-proof -
-  note sets_eq_imp_space_eq[OF sets_eq, simp]
-  show ?thesis
-    using sup_measure'_def
-  proof (rule emeasure_measure_of)
-    let ?d = "\<lambda>X Y. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)"
-    show "countably_additive (sets (sup_measure' A B)) (\<lambda>X. SUP Y : sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y))"
-    proof (rule countably_additiveI, goal_cases)
-      case (1 X)
-      then have [measurable]: "\<And>i. X i \<in> sets A" and "disjoint_family X"
-        by auto
-      have "(\<Sum>i. ?S (X i)) = (SUP Y:sets A. \<Sum>i. ?d (X i) Y)"
-      proof (rule ennreal_suminf_SUP_eq_directed)
-        fix J :: "nat set" and a b assume "finite J" and [measurable]: "a \<in> sets A" "b \<in> sets A"
-        have "\<exists>c\<in>sets A. c \<subseteq> X i \<and> (\<forall>a\<in>sets A. ?d (X i) a \<le> ?d (X i) c)" for i
-        proof cases
-          assume "emeasure A (X i) = top \<or> emeasure B (X i) = top"
-          then show ?thesis
-          proof
-            assume "emeasure A (X i) = top" then show ?thesis
-              by (intro bexI[of _ "X i"]) auto
-          next
-            assume "emeasure B (X i) = top" then show ?thesis
-              by (intro bexI[of _ "{}"]) auto
-          qed
-        next
-          assume finite: "\<not> (emeasure A (X i) = top \<or> emeasure B (X i) = top)"
-          then have "\<exists>Y\<in>sets A. Y \<subseteq> X i \<and> (\<forall>C\<in>sets A. C \<subseteq> Y \<longrightarrow> B C \<le> A C) \<and> (\<forall>C\<in>sets A. C \<subseteq> X i \<longrightarrow> C \<inter> Y = {} \<longrightarrow> A C \<le> B C)"
-            using unsigned_Hahn_decomposition[of B A "X i"] by simp
-          then obtain Y where [measurable]: "Y \<in> sets A" and [simp]: "Y \<subseteq> X i"
-            and B_le_A: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> B C \<le> A C"
-            and A_le_B: "\<And>C. C \<in> sets A \<Longrightarrow> C \<subseteq> X i \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> A C \<le> B C"
-            by auto
-
-          show ?thesis
-          proof (intro bexI[of _ Y] ballI conjI)
-            fix a assume [measurable]: "a \<in> sets A"
-            have *: "(X i \<inter> a \<inter> Y \<union> (X i \<inter> a - Y)) = X i \<inter> a" "(X i - a) \<inter> Y \<union> (X i - a - Y) = X i \<inter> - a"
-              for a Y by auto
-            then have "?d (X i) a =
-              (A (X i \<inter> a \<inter> Y) + A (X i \<inter> a \<inter> - Y)) + (B (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))"
-              by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric])
-            also have "\<dots> \<le> (A (X i \<inter> a \<inter> Y) + B (X i \<inter> a \<inter> - Y)) + (A (X i \<inter> - a \<inter> Y) + B (X i \<inter> - a \<inter> - Y))"
-              by (intro add_mono order_refl B_le_A A_le_B) (auto simp: Diff_eq[symmetric])
-            also have "\<dots> \<le> (A (X i \<inter> Y \<inter> a) + A (X i \<inter> Y \<inter> - a)) + (B (X i \<inter> - Y \<inter> a) + B (X i \<inter> - Y \<inter> - a))"
-              by (simp add: ac_simps)
-            also have "\<dots> \<le> A (X i \<inter> Y) + B (X i \<inter> - Y)"
-              by (subst (1 2) plus_emeasure) (auto simp: Diff_eq[symmetric] *)
-            finally show "?d (X i) a \<le> ?d (X i) Y" .
-          qed auto
-        qed
-        then obtain C where [measurable]: "C i \<in> sets A" and "C i \<subseteq> X i"
-          and C: "\<And>a. a \<in> sets A \<Longrightarrow> ?d (X i) a \<le> ?d (X i) (C i)" for i
-          by metis
-        have *: "X i \<inter> (\<Union>i. C i) = X i \<inter> C i" for i
-        proof safe
-          fix x j assume "x \<in> X i" "x \<in> C j"
-          moreover have "i = j \<or> X i \<inter> X j = {}"
-            using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
-          ultimately show "x \<in> C i"
-            using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
-        qed auto
-        have **: "X i \<inter> - (\<Union>i. C i) = X i \<inter> - C i" for i
-        proof safe
-          fix x j assume "x \<in> X i" "x \<notin> C i" "x \<in> C j"
-          moreover have "i = j \<or> X i \<inter> X j = {}"
-            using \<open>disjoint_family X\<close> by (auto simp: disjoint_family_on_def)
-          ultimately show False
-            using \<open>C i \<subseteq> X i\<close> \<open>C j \<subseteq> X j\<close> by auto
-        qed auto
-        show "\<exists>c\<in>sets A. \<forall>i\<in>J. ?d (X i) a \<le> ?d (X i) c \<and> ?d (X i) b \<le> ?d (X i) c"
-          apply (intro bexI[of _ "\<Union>i. C i"])
-          unfolding * **
-          apply (intro C ballI conjI)
-          apply auto
-          done
-      qed
-      also have "\<dots> = ?S (\<Union>i. X i)"
-        unfolding UN_extend_simps(4)
-        by (auto simp add: suminf_add[symmetric] Diff_eq[symmetric] simp del: UN_simps
-                 intro!: SUP_cong arg_cong2[where f="op +"] suminf_emeasure
-                         disjoint_family_on_bisimulation[OF \<open>disjoint_family X\<close>])
-      finally show "(\<Sum>i. ?S (X i)) = ?S (\<Union>i. X i)" .
-    qed
-  qed (auto dest: sets.sets_into_space simp: positive_def intro!: SUP_const)
-qed
-
-lemma le_emeasure_sup_measure'1:
-  assumes "sets B = sets A" "X \<in> sets A" shows "emeasure A X \<le> emeasure (sup_measure' A B) X"
-  by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "X"] assms)
-
-lemma le_emeasure_sup_measure'2:
-  assumes "sets B = sets A" "X \<in> sets A" shows "emeasure B X \<le> emeasure (sup_measure' A B) X"
-  by (subst emeasure_sup_measure'[OF assms]) (auto intro!: SUP_upper2[of "{}"] assms)
-
-lemma emeasure_sup_measure'_le2:
-  assumes [simp]: "sets B = sets C" "sets A = sets C" and [measurable]: "X \<in> sets C"
-  assumes A: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure A Y \<le> emeasure C Y"
-  assumes B: "\<And>Y. Y \<subseteq> X \<Longrightarrow> Y \<in> sets A \<Longrightarrow> emeasure B Y \<le> emeasure C Y"
-  shows "emeasure (sup_measure' A B) X \<le> emeasure C X"
-proof (subst emeasure_sup_measure')
-  show "(SUP Y:sets A. emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y)) \<le> emeasure C X"
-    unfolding \<open>sets A = sets C\<close>
-  proof (intro SUP_least)
-    fix Y assume [measurable]: "Y \<in> sets C"
-    have [simp]: "X \<inter> Y \<union> (X - Y) = X"
-      by auto
-    have "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C (X \<inter> Y) + emeasure C (X \<inter> - Y)"
-      by (intro add_mono A B) (auto simp: Diff_eq[symmetric])
-    also have "\<dots> = emeasure C X"
-      by (subst plus_emeasure) (auto simp: Diff_eq[symmetric])
-    finally show "emeasure A (X \<inter> Y) + emeasure B (X \<inter> - Y) \<le> emeasure C X" .
-  qed
-qed simp_all
-
-definition sup_lexord :: "'a \<Rightarrow> 'a \<Rightarrow> ('a \<Rightarrow> 'b::order) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
-where
-  "sup_lexord A B k s c =
-    (if k A = k B then c else if \<not> k A \<le> k B \<and> \<not> k B \<le> k A then s else if k B \<le> k A then A else B)"
-
-lemma sup_lexord:
-  "(k A < k B \<Longrightarrow> P B) \<Longrightarrow> (k B < k A \<Longrightarrow> P A) \<Longrightarrow> (k A = k B \<Longrightarrow> P c) \<Longrightarrow>
-    (\<not> k B \<le> k A \<Longrightarrow> \<not> k A \<le> k B \<Longrightarrow> P s) \<Longrightarrow> P (sup_lexord A B k s c)"
-  by (auto simp: sup_lexord_def)
-
-lemmas le_sup_lexord = sup_lexord[where P="\<lambda>a. c \<le> a" for c]
-
-lemma sup_lexord1: "k A = k B \<Longrightarrow> sup_lexord A B k s c = c"
-  by (simp add: sup_lexord_def)
-
-lemma sup_lexord_commute: "sup_lexord A B k s c = sup_lexord B A k s c"
-  by (auto simp: sup_lexord_def)
-
-lemma sigma_sets_le_sets_iff: "(sigma_sets (space x) \<A> \<subseteq> sets x) = (\<A> \<subseteq> sets x)"
-  using sets.sigma_sets_subset[of \<A> x] by auto
-
-lemma sigma_le_iff: "\<A> \<subseteq> Pow \<Omega> \<Longrightarrow> sigma \<Omega> \<A> \<le> x \<longleftrightarrow> (\<Omega> \<subseteq> space x \<and> (space x = \<Omega> \<longrightarrow> \<A> \<subseteq> sets x))"
-  by (cases "\<Omega> = space x")
-     (simp_all add: eq_commute[of _ "sets x"] le_measure_iff emeasure_sigma le_fun_def
-                    sigma_sets_superset_generator sigma_sets_le_sets_iff)
-
-instantiation measure :: (type) semilattice_sup
-begin
-
-definition sup_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
-where
-  "sup_measure A B =
-    sup_lexord A B space (sigma (space A \<union> space B) {})
-      (sup_lexord A B sets (sigma (space A) (sets A \<union> sets B)) (sup_measure' A B))"
-
-instance
-proof
-  fix x y z :: "'a measure"
-  show "x \<le> sup x y"
-    unfolding sup_measure_def
-  proof (intro le_sup_lexord)
-    assume "space x = space y"
-    then have *: "sets x \<union> sets y \<subseteq> Pow (space x)"
-      using sets.space_closed by auto
-    assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y"
-    then have "sets x \<subset> sets x \<union> sets y"
-      by auto
-    also have "\<dots> \<le> sigma (space x) (sets x \<union> sets y)"
-      by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
-    finally show "x \<le> sigma (space x) (sets x \<union> sets y)"
-      by (simp add: space_measure_of[OF *] less_eq_measure.intros(2))
-  next
-    assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y"
-    then show "x \<le> sigma (space x \<union> space y) {}"
-      by (intro less_eq_measure.intros) auto
-  next
-    assume "sets x = sets y" then show "x \<le> sup_measure' x y"
-      by (simp add: le_measure le_emeasure_sup_measure'1)
-  qed (auto intro: less_eq_measure.intros)
-  show "y \<le> sup x y"
-    unfolding sup_measure_def
-  proof (intro le_sup_lexord)
-    assume **: "space x = space y"
-    then have *: "sets x \<union> sets y \<subseteq> Pow (space y)"
-      using sets.space_closed by auto
-    assume "\<not> sets y \<subseteq> sets x" "\<not> sets x \<subseteq> sets y"
-    then have "sets y \<subset> sets x \<union> sets y"
-      by auto
-    also have "\<dots> \<le> sigma (space y) (sets x \<union> sets y)"
-      by (subst sets_measure_of[OF *]) (rule sigma_sets_superset_generator)
-    finally show "y \<le> sigma (space x) (sets x \<union> sets y)"
-      by (simp add: ** space_measure_of[OF *] less_eq_measure.intros(2))
-  next
-    assume "\<not> space y \<subseteq> space x" "\<not> space x \<subseteq> space y"
-    then show "y \<le> sigma (space x \<union> space y) {}"
-      by (intro less_eq_measure.intros) auto
-  next
-    assume "sets x = sets y" then show "y \<le> sup_measure' x y"
-      by (simp add: le_measure le_emeasure_sup_measure'2)
-  qed (auto intro: less_eq_measure.intros)
-  show "x \<le> y \<Longrightarrow> z \<le> y \<Longrightarrow> sup x z \<le> y"
-    unfolding sup_measure_def
-  proof (intro sup_lexord[where P="\<lambda>x. x \<le> y"])
-    assume "x \<le> y" "z \<le> y" and [simp]: "space x = space z" "sets x = sets z"
-    from \<open>x \<le> y\<close> show "sup_measure' x z \<le> y"
-    proof cases
-      case 1 then show ?thesis
-        by (intro less_eq_measure.intros(1)) simp
-    next
-      case 2 then show ?thesis
-        by (intro less_eq_measure.intros(2)) simp_all
-    next
-      case 3 with \<open>z \<le> y\<close> \<open>x \<le> y\<close> show ?thesis
-        by (auto simp add: le_measure intro!: emeasure_sup_measure'_le2)
-    qed
-  next
-    assume **: "x \<le> y" "z \<le> y" "space x = space z" "\<not> sets z \<subseteq> sets x" "\<not> sets x \<subseteq> sets z"
-    then have *: "sets x \<union> sets z \<subseteq> Pow (space x)"
-      using sets.space_closed by auto
-    show "sigma (space x) (sets x \<union> sets z) \<le> y"
-      unfolding sigma_le_iff[OF *] using ** by (auto simp: le_measure_iff split: if_split_asm)
-  next
-    assume "x \<le> y" "z \<le> y" "\<not> space z \<subseteq> space x" "\<not> space x \<subseteq> space z"
-    then have "space x \<subseteq> space y" "space z \<subseteq> space y"
-      by (auto simp: le_measure_iff split: if_split_asm)
-    then show "sigma (space x \<union> space z) {} \<le> y"
-      by (simp add: sigma_le_iff)
-  qed
-qed
-
-end
-
-lemma space_empty_eq_bot: "space a = {} \<longleftrightarrow> a = bot"
-  using space_empty[of a] by (auto intro!: measure_eqI)
-
-interpretation sup_measure: comm_monoid_set sup "bot :: 'a measure"
-  proof qed (auto intro!: antisym)
-
-lemma sup_measure_F_mono':
-  "finite J \<Longrightarrow> finite I \<Longrightarrow> sup_measure.F id I \<le> sup_measure.F id (I \<union> J)"
-proof (induction J rule: finite_induct)
-  case empty then show ?case
-    by simp
-next
-  case (insert i J)
-  show ?case
-  proof cases
-    assume "i \<in> I" with insert show ?thesis
-      by (auto simp: insert_absorb)
-  next
-    assume "i \<notin> I"
-    have "sup_measure.F id I \<le> sup_measure.F id (I \<union> J)"
-      by (intro insert)
-    also have "\<dots> \<le> sup_measure.F id (insert i (I \<union> J))"
-      using insert \<open>i \<notin> I\<close> by (subst sup_measure.insert) auto
-    finally show ?thesis
-      by auto
-  qed
-qed
-
-lemma sup_measure_F_mono: "finite I \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sup_measure.F id J \<le> sup_measure.F id I"
-  using sup_measure_F_mono'[of I J] by (auto simp: finite_subset Un_absorb1)
-
-lemma sets_eq_iff_bounded: "A \<le> B \<Longrightarrow> B \<le> C \<Longrightarrow> sets A = sets C \<Longrightarrow> sets B = sets A"
-  by (auto dest: sets_eq_imp_space_eq simp add: le_measure_iff split: if_split_asm)
-
-lemma sets_sup: "sets A = sets M \<Longrightarrow> sets B = sets M \<Longrightarrow> sets (sup A B) = sets M"
-  by (auto simp add: sup_measure_def sup_lexord_def dest: sets_eq_imp_space_eq)
-
-lemma sets_sup_measure_F:
-  "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> sets i = sets M) \<Longrightarrow> sets (sup_measure.F id I) = sets M"
-  by (induction I rule: finite_ne_induct) (simp_all add: sets_sup)
-
-lemma le_measureD1: "A \<le> B \<Longrightarrow> space A \<le> space B"
-  by (auto simp: le_measure_iff split: if_split_asm)
-
-lemma le_measureD2: "A \<le> B \<Longrightarrow> space A = space B \<Longrightarrow> sets A \<le> sets B"
-  by (auto simp: le_measure_iff split: if_split_asm)
-
-lemma le_measureD3: "A \<le> B \<Longrightarrow> sets A = sets B \<Longrightarrow> emeasure A X \<le> emeasure B X"
-  by (auto simp: le_measure_iff le_fun_def dest: sets_eq_imp_space_eq split: if_split_asm)
-
-definition Sup_measure' :: "'a measure set \<Rightarrow> 'a measure"
-where
-  "Sup_measure' M = measure_of (\<Union>a\<in>M. space a) (\<Union>a\<in>M. sets a)
-    (\<lambda>X. (SUP P:{P. finite P \<and> P \<subseteq> M }. sup_measure.F id P X))"
-
-lemma UN_space_closed: "UNION S sets \<subseteq> Pow (UNION S space)"
-  using sets.space_closed by auto
-
-lemma space_Sup_measure'2: "space (Sup_measure' M) = (\<Union>m\<in>M. space m)"
-  unfolding Sup_measure'_def by (intro space_measure_of[OF UN_space_closed])
-
-lemma sets_Sup_measure'2: "sets (Sup_measure' M) = sigma_sets (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)"
-  unfolding Sup_measure'_def by (intro sets_measure_of[OF UN_space_closed])
-
-lemma sets_Sup_measure':
-  assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
-  shows "sets (Sup_measure' M) = sets A"
-  using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close> by (simp add: Sup_measure'_def)
-
-lemma space_Sup_measure':
-  assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "M \<noteq> {}"
-  shows "space (Sup_measure' M) = space A"
-  using sets_eq[THEN sets_eq_imp_space_eq, simp] \<open>M \<noteq> {}\<close>
-  by (simp add: Sup_measure'_def )
-
-lemma emeasure_Sup_measure':
-  assumes sets_eq[simp]: "\<And>m. m \<in> M \<Longrightarrow> sets m = sets A" and "X \<in> sets A" "M \<noteq> {}"
-  shows "emeasure (Sup_measure' M) X = (SUP P:{P. finite P \<and> P \<subseteq> M}. sup_measure.F id P X)"
-    (is "_ = ?S X")
-  using Sup_measure'_def
-proof (rule emeasure_measure_of)
-  note sets_eq[THEN sets_eq_imp_space_eq, simp]
-  have *: "sets (Sup_measure' M) = sets A" "space (Sup_measure' M) = space A"
-    using \<open>M \<noteq> {}\<close> by (simp_all add: Sup_measure'_def)
-  let ?\<mu> = "sup_measure.F id"
-  show "countably_additive (sets (Sup_measure' M)) ?S"
-  proof (rule countably_additiveI, goal_cases)
-    case (1 F)
-    then have **: "range F \<subseteq> sets A"
-      by (auto simp: *)
-    show "(\<Sum>i. ?S (F i)) = ?S (\<Union>i. F i)"
-    proof (subst ennreal_suminf_SUP_eq_directed)
-      fix i j and N :: "nat set" assume ij: "i \<in> {P. finite P \<and> P \<subseteq> M}" "j \<in> {P. finite P \<and> P \<subseteq> M}"
-      have "(i \<noteq> {} \<longrightarrow> sets (?\<mu> i) = sets A) \<and> (j \<noteq> {} \<longrightarrow> sets (?\<mu> j) = sets A) \<and>
-        (i \<noteq> {} \<or> j \<noteq> {} \<longrightarrow> sets (?\<mu> (i \<union> j)) = sets A)"
-        using ij by (intro impI sets_sup_measure_F conjI) auto
-      then have "?\<mu> j (F n) \<le> ?\<mu> (i \<union> j) (F n) \<and> ?\<mu> i (F n) \<le> ?\<mu> (i \<union> j) (F n)" for n
-        using ij
-        by (cases "i = {}"; cases "j = {}")
-           (auto intro!: le_measureD3 sup_measure_F_mono simp: sets_sup_measure_F
-                 simp del: id_apply)
-      with ij show "\<exists>k\<in>{P. finite P \<and> P \<subseteq> M}. \<forall>n\<in>N. ?\<mu> i (F n) \<le> ?\<mu> k (F n) \<and> ?\<mu> j (F n) \<le> ?\<mu> k (F n)"
-        by (safe intro!: bexI[of _ "i \<union> j"]) auto
-    next
-      show "(SUP P : {P. finite P \<and> P \<subseteq> M}. \<Sum>n. ?\<mu> P (F n)) = (SUP P : {P. finite P \<and> P \<subseteq> M}. ?\<mu> P (UNION UNIV F))"
-      proof (intro SUP_cong refl)
-        fix i assume i: "i \<in> {P. finite P \<and> P \<subseteq> M}"
-        show "(\<Sum>n. ?\<mu> i (F n)) = ?\<mu> i (UNION UNIV F)"
-        proof cases
-          assume "i \<noteq> {}" with i ** show ?thesis
-            apply (intro suminf_emeasure \<open>disjoint_family F\<close>)
-            apply (subst sets_sup_measure_F[OF _ _ sets_eq])
-            apply auto
-            done
-        qed simp
-      qed
-    qed
-  qed
-  show "positive (sets (Sup_measure' M)) ?S"
-    by (auto simp: positive_def bot_ennreal[symmetric])
-  show "X \<in> sets (Sup_measure' M)"
-    using assms * by auto
-qed (rule UN_space_closed)
-
-definition Sup_lexord :: "('a \<Rightarrow> 'b::complete_lattice) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> ('a set \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> 'a"
-where
-  "Sup_lexord k c s A = (let U = (SUP a:A. k a) in if \<exists>a\<in>A. k a = U then c {a\<in>A. k a = U} else s A)"
-
-lemma Sup_lexord:
-  "(\<And>a S. a \<in> A \<Longrightarrow> k a = (SUP a:A. k a) \<Longrightarrow> S = {a'\<in>A. k a' = k a} \<Longrightarrow> P (c S)) \<Longrightarrow> ((\<And>a. a \<in> A \<Longrightarrow> k a \<noteq> (SUP a:A. k a)) \<Longrightarrow> P (s A)) \<Longrightarrow>
-    P (Sup_lexord k c s A)"
-  by (auto simp: Sup_lexord_def Let_def)
-
-lemma Sup_lexord1:
-  assumes A: "A \<noteq> {}" "(\<And>a. a \<in> A \<Longrightarrow> k a = (\<Union>a\<in>A. k a))" "P (c A)"
-  shows "P (Sup_lexord k c s A)"
-  unfolding Sup_lexord_def Let_def
-proof (clarsimp, safe)
-  show "\<forall>a\<in>A. k a \<noteq> (\<Union>x\<in>A. k x) \<Longrightarrow> P (s A)"
-    by (metis assms(1,2) ex_in_conv)
-next
-  fix a assume "a \<in> A" "k a = (\<Union>x\<in>A. k x)"
-  then have "{a \<in> A. k a = (\<Union>x\<in>A. k x)} = {a \<in> A. k a = k a}"
-    by (metis A(2)[symmetric])
-  then show "P (c {a \<in> A. k a = (\<Union>x\<in>A. k x)})"
-    by (simp add: A(3))
-qed
-
-instantiation measure :: (type) complete_lattice
-begin
-
-definition Sup_measure :: "'a measure set \<Rightarrow> 'a measure"
-where
-  "Sup_measure = Sup_lexord space (Sup_lexord sets Sup_measure'
-    (\<lambda>U. sigma (\<Union>u\<in>U. space u) (\<Union>u\<in>U. sets u))) (\<lambda>U. sigma (\<Union>u\<in>U. space u) {})"
-
-definition Inf_measure :: "'a measure set \<Rightarrow> 'a measure"
-where
-  "Inf_measure A = Sup {x. \<forall>a\<in>A. x \<le> a}"
-
-definition inf_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
-where
-  "inf_measure a b = Inf {a, b}"
-
-definition top_measure :: "'a measure"
-where
-  "top_measure = Inf {}"
-
-instance
-proof
-  note UN_space_closed [simp]
-  show upper: "x \<le> Sup A" if x: "x \<in> A" for x :: "'a measure" and A
-    unfolding Sup_measure_def
-  proof (intro Sup_lexord[where P="\<lambda>y. x \<le> y"])
-    assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)"
-    from this[OF \<open>x \<in> A\<close>] \<open>x \<in> A\<close> show "x \<le> sigma (\<Union>a\<in>A. space a) {}"
-      by (intro less_eq_measure.intros) auto
-  next
-    fix a S assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
-      and neq: "\<And>aa. aa \<in> S \<Longrightarrow> sets aa \<noteq> (\<Union>a\<in>S. sets a)"
-    have sp_a: "space a = (UNION S space)"
-      using \<open>a\<in>A\<close> by (auto simp: S)
-    show "x \<le> sigma (UNION S space) (UNION S sets)"
-    proof cases
-      assume [simp]: "space x = space a"
-      have "sets x \<subset> (\<Union>a\<in>S. sets a)"
-        using \<open>x\<in>A\<close> neq[of x] by (auto simp: S)
-      also have "\<dots> \<subseteq> sigma_sets (\<Union>x\<in>S. space x) (\<Union>x\<in>S. sets x)"
-        by (rule sigma_sets_superset_generator)
-      finally show ?thesis
-        by (intro less_eq_measure.intros(2)) (simp_all add: sp_a)
-    next
-      assume "space x \<noteq> space a"
-      moreover have "space x \<le> space a"
-        unfolding a using \<open>x\<in>A\<close> by auto
-      ultimately show ?thesis
-        by (intro less_eq_measure.intros) (simp add: less_le sp_a)
-    qed
-  next
-    fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
-      and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
-    then have "S' \<noteq> {}" "space b = space a"
-      by auto
-    have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b"
-      by (auto simp: S')
-    note sets_eq[THEN sets_eq_imp_space_eq, simp]
-    have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
-      using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
-    show "x \<le> Sup_measure' S'"
-    proof cases
-      assume "x \<in> S"
-      with \<open>b \<in> S\<close> have "space x = space b"
-        by (simp add: S)
-      show ?thesis
-      proof cases
-        assume "x \<in> S'"
-        show "x \<le> Sup_measure' S'"
-        proof (intro le_measure[THEN iffD2] ballI)
-          show "sets x = sets (Sup_measure' S')"
-            using \<open>x\<in>S'\<close> * by (simp add: S')
-          fix X assume "X \<in> sets x"
-          show "emeasure x X \<le> emeasure (Sup_measure' S') X"
-          proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets x\<close>])
-            show "emeasure x X \<le> (SUP P : {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X)"
-              using \<open>x\<in>S'\<close> by (intro SUP_upper2[where i="{x}"]) auto
-          qed (insert \<open>x\<in>S'\<close> S', auto)
-        qed
-      next
-        assume "x \<notin> S'"
-        then have "sets x \<noteq> sets b"
-          using \<open>x\<in>S\<close> by (auto simp: S')
-        moreover have "sets x \<le> sets b"
-          using \<open>x\<in>S\<close> unfolding b by auto
-        ultimately show ?thesis
-          using * \<open>x \<in> S\<close>
-          by (intro less_eq_measure.intros(2))
-             (simp_all add: * \<open>space x = space b\<close> less_le)
-      qed
-    next
-      assume "x \<notin> S"
-      with \<open>x\<in>A\<close> \<open>x \<notin> S\<close> \<open>space b = space a\<close> show ?thesis
-        by (intro less_eq_measure.intros)
-           (simp_all add: * less_le a SUP_upper S)
-    qed
-  qed
-  show least: "Sup A \<le> x" if x: "\<And>z. z \<in> A \<Longrightarrow> z \<le> x" for x :: "'a measure" and A
-    unfolding Sup_measure_def
-  proof (intro Sup_lexord[where P="\<lambda>y. y \<le> x"])
-    assume "\<And>a. a \<in> A \<Longrightarrow> space a \<noteq> (\<Union>a\<in>A. space a)"
-    show "sigma (UNION A space) {} \<le> x"
-      using x[THEN le_measureD1] by (subst sigma_le_iff) auto
-  next
-    fix a S assume "a \<in> A" "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
-      "\<And>a. a \<in> S \<Longrightarrow> sets a \<noteq> (\<Union>a\<in>S. sets a)"
-    have "UNION S space \<subseteq> space x"
-      using S le_measureD1[OF x] by auto
-    moreover
-    have "UNION S space = space a"
-      using \<open>a\<in>A\<close> S by auto
-    then have "space x = UNION S space \<Longrightarrow> UNION S sets \<subseteq> sets x"
-      using \<open>a \<in> A\<close> le_measureD2[OF x] by (auto simp: S)
-    ultimately show "sigma (UNION S space) (UNION S sets) \<le> x"
-      by (subst sigma_le_iff) simp_all
-  next
-    fix a b S S' assume "a \<in> A" and a: "space a = (\<Union>a\<in>A. space a)" and S: "S = {a' \<in> A. space a' = space a}"
-      and "b \<in> S" and b: "sets b = (\<Union>a\<in>S. sets a)" and S': "S' = {a' \<in> S. sets a' = sets b}"
-    then have "S' \<noteq> {}" "space b = space a"
-      by auto
-    have sets_eq: "\<And>x. x \<in> S' \<Longrightarrow> sets x = sets b"
-      by (auto simp: S')
-    note sets_eq[THEN sets_eq_imp_space_eq, simp]
-    have *: "sets (Sup_measure' S') = sets b" "space (Sup_measure' S') = space b"
-      using \<open>S' \<noteq> {}\<close> by (simp_all add: Sup_measure'_def sets_eq)
-    show "Sup_measure' S' \<le> x"
-    proof cases
-      assume "space x = space a"
-      show ?thesis
-      proof cases
-        assume **: "sets x = sets b"
-        show ?thesis
-        proof (intro le_measure[THEN iffD2] ballI)
-          show ***: "sets (Sup_measure' S') = sets x"
-            by (simp add: * **)
-          fix X assume "X \<in> sets (Sup_measure' S')"
-          show "emeasure (Sup_measure' S') X \<le> emeasure x X"
-            unfolding ***
-          proof (subst emeasure_Sup_measure'[OF _ \<open>X \<in> sets (Sup_measure' S')\<close>])
-            show "(SUP P : {P. finite P \<and> P \<subseteq> S'}. emeasure (sup_measure.F id P) X) \<le> emeasure x X"
-            proof (safe intro!: SUP_least)
-              fix P assume P: "finite P" "P \<subseteq> S'"
-              show "emeasure (sup_measure.F id P) X \<le> emeasure x X"
-              proof cases
-                assume "P = {}" then show ?thesis
-                  by auto
-              next
-                assume "P \<noteq> {}"
-                from P have "finite P" "P \<subseteq> A"
-                  unfolding S' S by (simp_all add: subset_eq)
-                then have "sup_measure.F id P \<le> x"
-                  by (induction P) (auto simp: x)
-                moreover have "sets (sup_measure.F id P) = sets x"
-                  using \<open>finite P\<close> \<open>P \<noteq> {}\<close> \<open>P \<subseteq> S'\<close> \<open>sets x = sets b\<close>
-                  by (intro sets_sup_measure_F) (auto simp: S')
-                ultimately show "emeasure (sup_measure.F id P) X \<le> emeasure x X"
-                  by (rule le_measureD3)
-              qed
-            qed
-            show "m \<in> S' \<Longrightarrow> sets m = sets (Sup_measure' S')" for m
-              unfolding * by (simp add: S')
-          qed fact
-        qed
-      next
-        assume "sets x \<noteq> sets b"
-        moreover have "sets b \<le> sets x"
-          unfolding b S using x[THEN le_measureD2] \<open>space x = space a\<close> by auto
-        ultimately show "Sup_measure' S' \<le> x"
-          using \<open>space x = space a\<close> \<open>b \<in> S\<close>
-          by (intro less_eq_measure.intros(2)) (simp_all add: * S)
-      qed
-    next
-      assume "space x \<noteq> space a"
-      then have "space a < space x"
-        using le_measureD1[OF x[OF \<open>a\<in>A\<close>]] by auto
-      then show "Sup_measure' S' \<le> x"
-        by (intro less_eq_measure.intros) (simp add: * \<open>space b = space a\<close>)
-    qed
-  qed
-  show "Sup {} = (bot::'a measure)" "Inf {} = (top::'a measure)"
-    by (auto intro!: antisym least simp: top_measure_def)
-  show lower: "x \<in> A \<Longrightarrow> Inf A \<le> x" for x :: "'a measure" and A
-    unfolding Inf_measure_def by (intro least) auto
-  show greatest: "(\<And>z. z \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> x \<le> Inf A" for x :: "'a measure" and A
-    unfolding Inf_measure_def by (intro upper) auto
-  show "inf x y \<le> x" "inf x y \<le> y" "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> inf y z" for x y z :: "'a measure"
-    by (auto simp: inf_measure_def intro!: lower greatest)
-qed
-
-end
-
-lemma sets_SUP:
-  assumes "\<And>x. x \<in> I \<Longrightarrow> sets (M x) = sets N"
-  shows "I \<noteq> {} \<Longrightarrow> sets (SUP i:I. M i) = sets N"
-  unfolding Sup_measure_def
-  using assms assms[THEN sets_eq_imp_space_eq]
-    sets_Sup_measure'[where A=N and M="M`I"]
-  by (intro Sup_lexord1[where P="\<lambda>x. sets x = sets N"]) auto
-
-lemma emeasure_SUP:
-  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N" "I \<noteq> {}"
-  shows "emeasure (SUP i:I. M i) X = (SUP J:{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emeasure (SUP i:J. M i) X)"
-proof -
-  have eq: "finite J \<Longrightarrow> sup_measure.F id J = (SUP i:J. i)" for J :: "'b measure set"
-    by (induction J rule: finite_induct) auto
-  have 1: "J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sets (SUP x:J. M x) = sets N" for J
-    by (intro sets_SUP sets) (auto )
-
-  from \<open>I \<noteq> {}\<close> obtain i where "i\<in>I" by auto
-  have "Sup_measure' (M`I) X = (SUP P:{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X)"
-    using sets by (intro emeasure_Sup_measure') auto
-  also have "Sup_measure' (M`I) = (SUP i:I. M i)"
-    unfolding Sup_measure_def using \<open>I \<noteq> {}\<close> sets sets(1)[THEN sets_eq_imp_space_eq]
-    by (intro Sup_lexord1[where P="\<lambda>x. _ = x"]) auto
-  also have "(SUP P:{P. finite P \<and> P \<subseteq> M`I}. sup_measure.F id P X) =
-    (SUP J:{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. (SUP i:J. M i) X)"
-  proof (intro SUP_eq)
-    fix J assume "J \<in> {P. finite P \<and> P \<subseteq> M`I}"
-    then obtain J' where J': "J' \<subseteq> I" "finite J'" and J: "J = M`J'" and "finite J"
-      using finite_subset_image[of J M I] by auto
-    show "\<exists>j\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. sup_measure.F id J X \<le> (SUP i:j. M i) X"
-    proof cases
-      assume "J' = {}" with \<open>i \<in> I\<close> show ?thesis
-        by (auto simp add: J)
-    next
-      assume "J' \<noteq> {}" with J J' show ?thesis
-        by (intro bexI[of _ "J'"]) (auto simp add: eq simp del: id_apply)
-    qed
-  next
-    fix J assume J: "J \<in> {P. P \<noteq> {} \<and> finite P \<and> P \<subseteq> I}"
-    show "\<exists>J'\<in>{J. finite J \<and> J \<subseteq> M`I}. (SUP i:J. M i) X \<le> sup_measure.F id J' X"
-      using J by (intro bexI[of _ "M`J"]) (auto simp add: eq simp del: id_apply)
-  qed
-  finally show ?thesis .
-qed
-
-lemma emeasure_SUP_chain:
-  assumes sets: "\<And>i. i \<in> A \<Longrightarrow> sets (M i) = sets N" "X \<in> sets N"
-  assumes ch: "Complete_Partial_Order.chain op \<le> (M ` A)" and "A \<noteq> {}"
-  shows "emeasure (SUP i:A. M i) X = (SUP i:A. emeasure (M i) X)"
-proof (subst emeasure_SUP[OF sets \<open>A \<noteq> {}\<close>])
-  show "(SUP J:{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (SUPREMUM J M) X) = (SUP i:A. emeasure (M i) X)"
-  proof (rule SUP_eq)
-    fix J assume "J \<in> {J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}"
-    then have J: "Complete_Partial_Order.chain op \<le> (M ` J)" "finite J" "J \<noteq> {}" and "J \<subseteq> A"
-      using ch[THEN chain_subset, of "M`J"] by auto
-    with in_chain_finite[OF J(1)] obtain j where "j \<in> J" "(SUP j:J. M j) = M j"
-      by auto
-    with \<open>J \<subseteq> A\<close> show "\<exists>j\<in>A. emeasure (SUPREMUM J M) X \<le> emeasure (M j) X"
-      by auto
-  next
-    fix j assume "j\<in>A" then show "\<exists>i\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> A}. emeasure (M j) X \<le> emeasure (SUPREMUM i M) X"
-      by (intro bexI[of _ "{j}"]) auto
-  qed
-qed
-
-subsubsection \<open>Supremum of a set of $\sigma$-algebras\<close>
-
-lemma space_Sup_eq_UN: "space (Sup M) = (\<Union>x\<in>M. space x)"
-  unfolding Sup_measure_def
-  apply (intro Sup_lexord[where P="\<lambda>x. space x = _"])
-  apply (subst space_Sup_measure'2)
-  apply auto []
-  apply (subst space_measure_of[OF UN_space_closed])
-  apply auto
-  done
-
-lemma sets_Sup_eq:
-  assumes *: "\<And>m. m \<in> M \<Longrightarrow> space m = X" and "M \<noteq> {}"
-  shows "sets (Sup M) = sigma_sets X (\<Union>x\<in>M. sets x)"
-  unfolding Sup_measure_def
-  apply (rule Sup_lexord1)
-  apply fact
-  apply (simp add: assms)
-  apply (rule Sup_lexord)
-  subgoal premises that for a S
-    unfolding that(3) that(2)[symmetric]
-    using that(1)
-    apply (subst sets_Sup_measure'2)
-    apply (intro arg_cong2[where f=sigma_sets])
-    apply (auto simp: *)
-    done
-  apply (subst sets_measure_of[OF UN_space_closed])
-  apply (simp add:  assms)
-  done
-
-lemma in_sets_Sup: "(\<And>m. m \<in> M \<Longrightarrow> space m = X) \<Longrightarrow> m \<in> M \<Longrightarrow> A \<in> sets m \<Longrightarrow> A \<in> sets (Sup M)"
-  by (subst sets_Sup_eq[where X=X]) auto
-
-lemma Sup_lexord_rel:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> k (A i) = k (B i)"
-    "R (c (A ` {a \<in> I. k (B a) = (SUP x:I. k (B x))})) (c (B ` {a \<in> I. k (B a) = (SUP x:I. k (B x))}))"
-    "R (s (A`I)) (s (B`I))"
-  shows "R (Sup_lexord k c s (A`I)) (Sup_lexord k c s (B`I))"
-proof -
-  have "A ` {a \<in> I. k (B a) = (SUP x:I. k (B x))} =  {a \<in> A ` I. k a = (SUP x:I. k (B x))}"
-    using assms(1) by auto
-  moreover have "B ` {a \<in> I. k (B a) = (SUP x:I. k (B x))} =  {a \<in> B ` I. k a = (SUP x:I. k (B x))}"
-    by auto
-  ultimately show ?thesis
-    using assms by (auto simp: Sup_lexord_def Let_def)
-qed
-
-lemma sets_SUP_cong:
-  assumes eq: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i)" shows "sets (SUP i:I. M i) = sets (SUP i:I. N i)"
-  unfolding Sup_measure_def
-  using eq eq[THEN sets_eq_imp_space_eq]
-  apply (intro Sup_lexord_rel[where R="\<lambda>x y. sets x = sets y"])
-  apply simp
-  apply simp
-  apply (simp add: sets_Sup_measure'2)
-  apply (intro arg_cong2[where f="\<lambda>x y. sets (sigma x y)"])
-  apply auto
-  done
-
-lemma sets_Sup_in_sets:
-  assumes "M \<noteq> {}"
-  assumes "\<And>m. m \<in> M \<Longrightarrow> space m = space N"
-  assumes "\<And>m. m \<in> M \<Longrightarrow> sets m \<subseteq> sets N"
-  shows "sets (Sup M) \<subseteq> sets N"
-proof -
-  have *: "UNION M space = space N"
-    using assms by auto
-  show ?thesis
-    unfolding * using assms by (subst sets_Sup_eq[of M "space N"]) (auto intro!: sets.sigma_sets_subset)
-qed
-
-lemma measurable_Sup1:
-  assumes m: "m \<in> M" and f: "f \<in> measurable m N"
-    and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n"
-  shows "f \<in> measurable (Sup M) N"
-proof -
-  have "space (Sup M) = space m"
-    using m by (auto simp add: space_Sup_eq_UN dest: const_space)
-  then show ?thesis
-    using m f unfolding measurable_def by (auto intro: in_sets_Sup[OF const_space])
-qed
-
-lemma measurable_Sup2:
-  assumes M: "M \<noteq> {}"
-  assumes f: "\<And>m. m \<in> M \<Longrightarrow> f \<in> measurable N m"
-    and const_space: "\<And>m n. m \<in> M \<Longrightarrow> n \<in> M \<Longrightarrow> space m = space n"
-  shows "f \<in> measurable N (Sup M)"
-proof -
-  from M obtain m where "m \<in> M" by auto
-  have space_eq: "\<And>n. n \<in> M \<Longrightarrow> space n = space m"
-    by (intro const_space \<open>m \<in> M\<close>)
-  have "f \<in> measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m))"
-  proof (rule measurable_measure_of)
-    show "f \<in> space N \<rightarrow> UNION M space"
-      using measurable_space[OF f] M by auto
-  qed (auto intro: measurable_sets f dest: sets.sets_into_space)
-  also have "measurable N (sigma (\<Union>m\<in>M. space m) (\<Union>m\<in>M. sets m)) = measurable N (Sup M)"
-    apply (intro measurable_cong_sets refl)
-    apply (subst sets_Sup_eq[OF space_eq M])
-    apply simp
-    apply (subst sets_measure_of[OF UN_space_closed])
-    apply (simp add: space_eq M)
-    done
-  finally show ?thesis .
-qed
-
-lemma sets_Sup_sigma:
-  assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
-  shows "sets (SUP m:M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
-proof -
-  { fix a m assume "a \<in> sigma_sets \<Omega> m" "m \<in> M"
-    then have "a \<in> sigma_sets \<Omega> (\<Union>M)"
-     by induction (auto intro: sigma_sets.intros) }
-  then show "sets (SUP m:M. sigma \<Omega> m) = sets (sigma \<Omega> (\<Union>M))"
-    apply (subst sets_Sup_eq[where X="\<Omega>"])
-    apply (auto simp add: M) []
-    apply auto []
-    apply (simp add: space_measure_of_conv M Union_least)
-    apply (rule sigma_sets_eqI)
-    apply auto
-    done
-qed
-
-lemma Sup_sigma:
-  assumes [simp]: "M \<noteq> {}" and M: "\<And>m. m \<in> M \<Longrightarrow> m \<subseteq> Pow \<Omega>"
-  shows "(SUP m:M. sigma \<Omega> m) = (sigma \<Omega> (\<Union>M))"
-proof (intro antisym SUP_least)
-  have *: "\<Union>M \<subseteq> Pow \<Omega>"
-    using M by auto
-  show "sigma \<Omega> (\<Union>M) \<le> (SUP m:M. sigma \<Omega> m)"
-  proof (intro less_eq_measure.intros(3))
-    show "space (sigma \<Omega> (\<Union>M)) = space (SUP m:M. sigma \<Omega> m)"
-      "sets (sigma \<Omega> (\<Union>M)) = sets (SUP m:M. sigma \<Omega> m)"
-      using sets_Sup_sigma[OF assms] sets_Sup_sigma[OF assms, THEN sets_eq_imp_space_eq]
-      by auto
-  qed (simp add: emeasure_sigma le_fun_def)
-  fix m assume "m \<in> M" then show "sigma \<Omega> m \<le> sigma \<Omega> (\<Union>M)"
-    by (subst sigma_le_iff) (auto simp add: M *)
-qed
-
-lemma SUP_sigma_sigma:
-  "M \<noteq> {} \<Longrightarrow> (\<And>m. m \<in> M \<Longrightarrow> f m \<subseteq> Pow \<Omega>) \<Longrightarrow> (SUP m:M. sigma \<Omega> (f m)) = sigma \<Omega> (\<Union>m\<in>M. f m)"
-  using Sup_sigma[of "f`M" \<Omega>] by auto
-
-lemma sets_vimage_Sup_eq:
-  assumes *: "M \<noteq> {}" "f \<in> X \<rightarrow> Y" "\<And>m. m \<in> M \<Longrightarrow> space m = Y"
-  shows "sets (vimage_algebra X f (Sup M)) = sets (SUP m : M. vimage_algebra X f m)"
-  (is "?IS = ?SI")
-proof
-  show "?IS \<subseteq> ?SI"
-    apply (intro sets_image_in_sets measurable_Sup2)
-    apply (simp add: space_Sup_eq_UN *)
-    apply (simp add: *)
-    apply (intro measurable_Sup1)
-    apply (rule imageI)
-    apply assumption
-    apply (rule measurable_vimage_algebra1)
-    apply (auto simp: *)
-    done
-  show "?SI \<subseteq> ?IS"
-    apply (intro sets_Sup_in_sets)
-    apply (auto simp: *) []
-    apply (auto simp: *) []
-    apply (elim imageE)
-    apply simp
-    apply (rule sets_image_in_sets)
-    apply simp
-    apply (simp add: measurable_def)
-    apply (simp add: * space_Sup_eq_UN sets_vimage_algebra2)
-    apply (auto intro: in_sets_Sup[OF *(3)])
-    done
-qed
-
-lemma restrict_space_eq_vimage_algebra':
-  "sets (restrict_space M \<Omega>) = sets (vimage_algebra (\<Omega> \<inter> space M) (\<lambda>x. x) M)"
-proof -
-  have *: "{A \<inter> (\<Omega> \<inter> space M) |A. A \<in> sets M} = {A \<inter> \<Omega> |A. A \<in> sets M}"
-    using sets.sets_into_space[of _ M] by blast
-
-  show ?thesis
-    unfolding restrict_space_def
-    by (subst sets_measure_of)
-       (auto simp add: image_subset_iff sets_vimage_algebra * dest: sets.sets_into_space intro!: arg_cong2[where f=sigma_sets])
-qed
-
-lemma sigma_le_sets:
-  assumes [simp]: "A \<subseteq> Pow X" shows "sets (sigma X A) \<subseteq> sets N \<longleftrightarrow> X \<in> sets N \<and> A \<subseteq> sets N"
-proof
-  have "X \<in> sigma_sets X A" "A \<subseteq> sigma_sets X A"
-    by (auto intro: sigma_sets_top)
-  moreover assume "sets (sigma X A) \<subseteq> sets N"
-  ultimately show "X \<in> sets N \<and> A \<subseteq> sets N"
-    by auto
-next
-  assume *: "X \<in> sets N \<and> A \<subseteq> sets N"
-  { fix Y assume "Y \<in> sigma_sets X A" from this * have "Y \<in> sets N"
-      by induction auto }
-  then show "sets (sigma X A) \<subseteq> sets N"
-    by auto
-qed
-
-lemma measurable_iff_sets:
-  "f \<in> measurable M N \<longleftrightarrow> (f \<in> space M \<rightarrow> space N \<and> sets (vimage_algebra (space M) f N) \<subseteq> sets M)"
-proof -
-  have *: "{f -` A \<inter> space M |A. A \<in> sets N} \<subseteq> Pow (space M)"
-    by auto
-  show ?thesis
-    unfolding measurable_def
-    by (auto simp add: vimage_algebra_def sigma_le_sets[OF *])
-qed
-
-lemma sets_vimage_algebra_space: "X \<in> sets (vimage_algebra X f M)"
-  using sets.top[of "vimage_algebra X f M"] by simp
-
-lemma measurable_mono:
-  assumes N: "sets N' \<le> sets N" "space N = space N'"
-  assumes M: "sets M \<le> sets M'" "space M = space M'"
-  shows "measurable M N \<subseteq> measurable M' N'"
-  unfolding measurable_def
-proof safe
-  fix f A assume "f \<in> space M \<rightarrow> space N" "A \<in> sets N'"
-  moreover assume "\<forall>y\<in>sets N. f -` y \<inter> space M \<in> sets M" note this[THEN bspec, of A]
-  ultimately show "f -` A \<inter> space M' \<in> sets M'"
-    using assms by auto
-qed (insert N M, auto)
-
-lemma measurable_Sup_measurable:
-  assumes f: "f \<in> space N \<rightarrow> A"
-  shows "f \<in> measurable N (Sup {M. space M = A \<and> f \<in> measurable N M})"
-proof (rule measurable_Sup2)
-  show "{M. space M = A \<and> f \<in> measurable N M} \<noteq> {}"
-    using f unfolding ex_in_conv[symmetric]
-    by (intro exI[of _ "sigma A {}"]) (auto intro!: measurable_measure_of)
-qed auto
-
-lemma (in sigma_algebra) sigma_sets_subset':
-  assumes a: "a \<subseteq> M" "\<Omega>' \<in> M"
-  shows "sigma_sets \<Omega>' a \<subseteq> M"
-proof
-  show "x \<in> M" if x: "x \<in> sigma_sets \<Omega>' a" for x
-    using x by (induct rule: sigma_sets.induct) (insert a, auto)
-qed
-
-lemma in_sets_SUP: "i \<in> I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> space (M i) = Y) \<Longrightarrow> X \<in> sets (M i) \<Longrightarrow> X \<in> sets (SUP i:I. M i)"
-  by (intro in_sets_Sup[where X=Y]) auto
-
-lemma measurable_SUP1:
-  "i \<in> I \<Longrightarrow> f \<in> measurable (M i) N \<Longrightarrow> (\<And>m n. m \<in> I \<Longrightarrow> n \<in> I \<Longrightarrow> space (M m) = space (M n)) \<Longrightarrow>
-    f \<in> measurable (SUP i:I. M i) N"
-  by (auto intro: measurable_Sup1)
-
-lemma sets_image_in_sets':
-  assumes X: "X \<in> sets N"
-  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> f -` A \<inter> X \<in> sets N"
-  shows "sets (vimage_algebra X f M) \<subseteq> sets N"
-  unfolding sets_vimage_algebra
-  by (rule sets.sigma_sets_subset') (auto intro!: measurable_sets X f)
-
-lemma mono_vimage_algebra:
-  "sets M \<le> sets N \<Longrightarrow> sets (vimage_algebra X f M) \<subseteq> sets (vimage_algebra X f N)"
-  using sets.top[of "sigma X {f -` A \<inter> X |A. A \<in> sets N}"]
-  unfolding vimage_algebra_def
-  apply (subst (asm) space_measure_of)
-  apply auto []
-  apply (subst sigma_le_sets)
-  apply auto
-  done
-
-lemma mono_restrict_space: "sets M \<le> sets N \<Longrightarrow> sets (restrict_space M X) \<subseteq> sets (restrict_space N X)"
-  unfolding sets_restrict_space by (rule image_mono)
-
-lemma sets_eq_bot: "sets M = {{}} \<longleftrightarrow> M = bot"
-  apply safe
-  apply (intro measure_eqI)
-  apply auto
-  done
-
-lemma sets_eq_bot2: "{{}} = sets M \<longleftrightarrow> M = bot"
-  using sets_eq_bot[of M] by blast
-
-end
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2401 +0,0 @@
-(*  Title:      HOL/Probability/Nonnegative_Lebesgue_Integration.thy
-    Author:     Johannes Hölzl, TU München
-    Author:     Armin Heller, TU München
-*)
-
-section \<open>Lebesgue Integration for Nonnegative Functions\<close>
-
-theory Nonnegative_Lebesgue_Integration
-  imports Measure_Space Borel_Space
-begin
-
-subsection "Simple function"
-
-text \<open>
-
-Our simple functions are not restricted to nonnegative real numbers. Instead
-they are just functions with a finite range and are measurable when singleton
-sets are measurable.
-
-\<close>
-
-definition "simple_function M g \<longleftrightarrow>
-    finite (g ` space M) \<and>
-    (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
-
-lemma simple_functionD:
-  assumes "simple_function M g"
-  shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M"
-proof -
-  show "finite (g ` space M)"
-    using assms unfolding simple_function_def by auto
-  have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto
-  also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
-  finally show "g -` X \<inter> space M \<in> sets M" using assms
-    by (auto simp del: UN_simps simp: simple_function_def)
-qed
-
-lemma measurable_simple_function[measurable_dest]:
-  "simple_function M f \<Longrightarrow> f \<in> measurable M (count_space UNIV)"
-  unfolding simple_function_def measurable_def
-proof safe
-  fix A assume "finite (f ` space M)" "\<forall>x\<in>f ` space M. f -` {x} \<inter> space M \<in> sets M"
-  then have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) \<in> sets M"
-    by (intro sets.finite_UN) auto
-  also have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) = f -` A \<inter> space M"
-    by (auto split: if_split_asm)
-  finally show "f -` A \<inter> space M \<in> sets M" .
-qed simp
-
-lemma borel_measurable_simple_function:
-  "simple_function M f \<Longrightarrow> f \<in> borel_measurable M"
-  by (auto dest!: measurable_simple_function simp: measurable_def)
-
-lemma simple_function_measurable2[intro]:
-  assumes "simple_function M f" "simple_function M g"
-  shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
-proof -
-  have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"
-    by auto
-  then show ?thesis using assms[THEN simple_functionD(2)] by auto
-qed
-
-lemma simple_function_indicator_representation:
-  fixes f ::"'a \<Rightarrow> ennreal"
-  assumes f: "simple_function M f" and x: "x \<in> space M"
-  shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
-  (is "?l = ?r")
-proof -
-  have "?r = (\<Sum>y \<in> f ` space M.
-    (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
-    by (auto intro!: setsum.cong)
-  also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
-    using assms by (auto dest: simple_functionD simp: setsum.delta)
-  also have "... = f x" using x by (auto simp: indicator_def)
-  finally show ?thesis by auto
-qed
-
-lemma simple_function_notspace:
-  "simple_function M (\<lambda>x. h x * indicator (- space M) x::ennreal)" (is "simple_function M ?h")
-proof -
-  have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
-  hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset)
-  have "?h -` {0} \<inter> space M = space M" by auto
-  thus ?thesis unfolding simple_function_def by auto
-qed
-
-lemma simple_function_cong:
-  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
-  shows "simple_function M f \<longleftrightarrow> simple_function M g"
-proof -
-  have "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
-    using assms by auto
-  with assms show ?thesis
-    by (simp add: simple_function_def cong: image_cong)
-qed
-
-lemma simple_function_cong_algebra:
-  assumes "sets N = sets M" "space N = space M"
-  shows "simple_function M f \<longleftrightarrow> simple_function N f"
-  unfolding simple_function_def assms ..
-
-lemma simple_function_borel_measurable:
-  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
-  assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
-  shows "simple_function M f"
-  using assms unfolding simple_function_def
-  by (auto intro: borel_measurable_vimage)
-
-lemma simple_function_iff_borel_measurable:
-  fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
-  shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M"
-  by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable)
-
-lemma simple_function_eq_measurable:
-  "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)"
-  using measurable_simple_function[of M f] by (fastforce simp: simple_function_def)
-
-lemma simple_function_const[intro, simp]:
-  "simple_function M (\<lambda>x. c)"
-  by (auto intro: finite_subset simp: simple_function_def)
-lemma simple_function_compose[intro, simp]:
-  assumes "simple_function M f"
-  shows "simple_function M (g \<circ> f)"
-  unfolding simple_function_def
-proof safe
-  show "finite ((g \<circ> f) ` space M)"
-    using assms unfolding simple_function_def by (auto simp: image_comp [symmetric])
-next
-  fix x assume "x \<in> space M"
-  let ?G = "g -` {g (f x)} \<inter> (f`space M)"
-  have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
-    (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
-  show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
-    using assms unfolding simple_function_def *
-    by (rule_tac sets.finite_UN) auto
-qed
-
-lemma simple_function_indicator[intro, simp]:
-  assumes "A \<in> sets M"
-  shows "simple_function M (indicator A)"
-proof -
-  have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
-    by (auto simp: indicator_def)
-  hence "finite ?S" by (rule finite_subset) simp
-  moreover have "- A \<inter> space M = space M - A" by auto
-  ultimately show ?thesis unfolding simple_function_def
-    using assms by (auto simp: indicator_def [abs_def])
-qed
-
-lemma simple_function_Pair[intro, simp]:
-  assumes "simple_function M f"
-  assumes "simple_function M g"
-  shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p")
-  unfolding simple_function_def
-proof safe
-  show "finite (?p ` space M)"
-    using assms unfolding simple_function_def
-    by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto
-next
-  fix x assume "x \<in> space M"
-  have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
-      (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
-    by auto
-  with \<open>x \<in> space M\<close> show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
-    using assms unfolding simple_function_def by auto
-qed
-
-lemma simple_function_compose1:
-  assumes "simple_function M f"
-  shows "simple_function M (\<lambda>x. g (f x))"
-  using simple_function_compose[OF assms, of g]
-  by (simp add: comp_def)
-
-lemma simple_function_compose2:
-  assumes "simple_function M f" and "simple_function M g"
-  shows "simple_function M (\<lambda>x. h (f x) (g x))"
-proof -
-  have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))"
-    using assms by auto
-  thus ?thesis by (simp_all add: comp_def)
-qed
-
-lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"]
-  and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
-  and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
-  and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
-  and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"]
-  and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
-  and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
-
-lemma simple_function_setsum[intro, simp]:
-  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
-  shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
-proof cases
-  assume "finite P" from this assms show ?thesis by induct auto
-qed auto
-
-lemma simple_function_ennreal[intro, simp]:
-  fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
-  shows "simple_function M (\<lambda>x. ennreal (f x))"
-  by (rule simple_function_compose1[OF sf])
-
-lemma simple_function_real_of_nat[intro, simp]:
-  fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
-  shows "simple_function M (\<lambda>x. real (f x))"
-  by (rule simple_function_compose1[OF sf])
-
-lemma borel_measurable_implies_simple_function_sequence:
-  fixes u :: "'a \<Rightarrow> ennreal"
-  assumes u[measurable]: "u \<in> borel_measurable M"
-  shows "\<exists>f. incseq f \<and> (\<forall>i. (\<forall>x. f i x < top) \<and> simple_function M (f i)) \<and> u = (SUP i. f i)"
-proof -
-  define f where [abs_def]:
-    "f i x = real_of_int (floor (enn2real (min i (u x)) * 2^i)) / 2^i" for i x
-
-  have [simp]: "0 \<le> f i x" for i x
-    by (auto simp: f_def intro!: divide_nonneg_nonneg mult_nonneg_nonneg enn2real_nonneg)
-
-  have *: "2^n * real_of_int x = real_of_int (2^n * x)" for n x
-    by simp
-
-  have "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = real_of_int \<lfloor>i * 2 ^ i\<rfloor>" for i
-    by (intro arg_cong[where f=real_of_int]) simp
-  then have [simp]: "real_of_int \<lfloor>real i * 2 ^ i\<rfloor> = i * 2 ^ i" for i
-    unfolding floor_of_nat by simp
-
-  have "incseq f"
-  proof (intro monoI le_funI)
-    fix m n :: nat and x assume "m \<le> n"
-    moreover
-    { fix d :: nat
-      have "\<lfloor>2^d::real\<rfloor> * \<lfloor>2^m * enn2real (min (of_nat m) (u x))\<rfloor> \<le>
-        \<lfloor>2^d * (2^m * enn2real (min (of_nat m) (u x)))\<rfloor>"
-        by (rule le_mult_floor) (auto simp: enn2real_nonneg)
-      also have "\<dots> \<le> \<lfloor>2^d * (2^m * enn2real (min (of_nat d + of_nat m) (u x)))\<rfloor>"
-        by (intro floor_mono mult_mono enn2real_mono min.mono)
-           (auto simp: enn2real_nonneg min_less_iff_disj of_nat_less_top)
-      finally have "f m x \<le> f (m + d) x"
-        unfolding f_def
-        by (auto simp: field_simps power_add * simp del: of_int_mult) }
-    ultimately show "f m x \<le> f n x"
-      by (auto simp add: le_iff_add)
-  qed
-  then have inc_f: "incseq (\<lambda>i. ennreal (f i x))" for x
-    by (auto simp: incseq_def le_fun_def)
-  then have "incseq (\<lambda>i x. ennreal (f i x))"
-    by (auto simp: incseq_def le_fun_def)
-  moreover
-  have "simple_function M (f i)" for i
-  proof (rule simple_function_borel_measurable)
-    have "\<lfloor>enn2real (min (of_nat i) (u x)) * 2 ^ i\<rfloor> \<le> \<lfloor>int i * 2 ^ i\<rfloor>" for x
-      by (cases "u x" rule: ennreal_cases)
-         (auto split: split_min intro!: floor_mono)
-    then have "f i ` space M \<subseteq> (\<lambda>n. real_of_int n / 2^i) ` {0 .. of_nat i * 2^i}"
-      unfolding floor_of_int by (auto simp: f_def enn2real_nonneg intro!: imageI)
-    then show "finite (f i ` space M)"
-      by (rule finite_subset) auto
-    show "f i \<in> borel_measurable M"
-      unfolding f_def enn2real_def by measurable
-  qed
-  moreover
-  { fix x
-    have "(SUP i. ennreal (f i x)) = u x"
-    proof (cases "u x" rule: ennreal_cases)
-      case top then show ?thesis
-        by (simp add: f_def inf_min[symmetric] ennreal_of_nat_eq_real_of_nat[symmetric]
-                      ennreal_SUP_of_nat_eq_top)
-    next
-      case (real r)
-      obtain n where "r \<le> of_nat n" using real_arch_simple by auto
-      then have min_eq_r: "\<forall>\<^sub>F x in sequentially. min (real x) r = r"
-        by (auto simp: eventually_sequentially intro!: exI[of _ n] split: split_min)
-
-      have "(\<lambda>i. real_of_int \<lfloor>min (real i) r * 2^i\<rfloor> / 2^i) \<longlonglongrightarrow> r"
-      proof (rule tendsto_sandwich)
-        show "(\<lambda>n. r - (1/2)^n) \<longlonglongrightarrow> r"
-          by (auto intro!: tendsto_eq_intros LIMSEQ_power_zero)
-        show "\<forall>\<^sub>F n in sequentially. real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n \<le> r"
-          using min_eq_r by eventually_elim (auto simp: field_simps)
-        have *: "r * (2 ^ n * 2 ^ n) \<le> 2^n + 2^n * real_of_int \<lfloor>r * 2 ^ n\<rfloor>" for n
-          using real_of_int_floor_ge_diff_one[of "r * 2^n", THEN mult_left_mono, of "2^n"]
-          by (auto simp: field_simps)
-        show "\<forall>\<^sub>F n in sequentially. r - (1/2)^n \<le> real_of_int \<lfloor>min (real n) r * 2 ^ n\<rfloor> / 2 ^ n"
-          using min_eq_r by eventually_elim (insert *, auto simp: field_simps)
-      qed auto
-      then have "(\<lambda>i. ennreal (f i x)) \<longlonglongrightarrow> ennreal r"
-        by (simp add: real f_def ennreal_of_nat_eq_real_of_nat min_ennreal)
-      from LIMSEQ_unique[OF LIMSEQ_SUP[OF inc_f] this]
-      show ?thesis
-        by (simp add: real)
-    qed }
-  ultimately show ?thesis
-    by (intro exI[of _ "\<lambda>i x. ennreal (f i x)"]) auto
-qed
-
-lemma borel_measurable_implies_simple_function_sequence':
-  fixes u :: "'a \<Rightarrow> ennreal"
-  assumes u: "u \<in> borel_measurable M"
-  obtains f where
-    "\<And>i. simple_function M (f i)" "incseq f" "\<And>i x. f i x < top" "\<And>x. (SUP i. f i x) = u x"
-  using borel_measurable_implies_simple_function_sequence[OF u] by (auto simp: fun_eq_iff) blast
-
-lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]:
-  fixes u :: "'a \<Rightarrow> ennreal"
-  assumes u: "simple_function M u"
-  assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
-  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
-  assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)"
-  assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
-  shows "P u"
-proof (rule cong)
-  from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
-  proof eventually_elim
-    fix x assume x: "x \<in> space M"
-    from simple_function_indicator_representation[OF u x]
-    show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
-  qed
-next
-  from u have "finite (u ` space M)"
-    unfolding simple_function_def by auto
-  then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
-  proof induct
-    case empty show ?case
-      using set[of "{}"] by (simp add: indicator_def[abs_def])
-  qed (auto intro!: add mult set simple_functionD u)
-next
-  show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
-    apply (subst simple_function_cong)
-    apply (rule simple_function_indicator_representation[symmetric])
-    apply (auto intro: u)
-    done
-qed fact
-
-lemma simple_function_induct_nn[consumes 1, case_names cong set mult add]:
-  fixes u :: "'a \<Rightarrow> ennreal"
-  assumes u: "simple_function M u"
-  assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
-  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
-  assumes mult: "\<And>u c. simple_function M u \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
-  assumes add: "\<And>u v. simple_function M u \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<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)"
-  shows "P u"
-proof -
-  show ?thesis
-  proof (rule cong)
-    fix x assume x: "x \<in> space M"
-    from simple_function_indicator_representation[OF u x]
-    show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
-  next
-    show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
-      apply (subst simple_function_cong)
-      apply (rule simple_function_indicator_representation[symmetric])
-      apply (auto intro: u)
-      done
-  next
-    from u have "finite (u ` space M)"
-      unfolding simple_function_def by auto
-    then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
-    proof induct
-      case empty show ?case
-        using set[of "{}"] by (simp add: indicator_def[abs_def])
-    next
-      case (insert x S)
-      { fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
-          x * indicator (u -` {x} \<inter> space M) z = 0"
-          using insert by (subst setsum_eq_0_iff) (auto simp: indicator_def) }
-      note disj = this
-      from insert show ?case
-        by (auto intro!: add mult set simple_functionD u simple_function_setsum disj)
-    qed
-  qed fact
-qed
-
-lemma borel_measurable_induct[consumes 1, case_names cong set mult add seq, induct set: borel_measurable]:
-  fixes u :: "'a \<Rightarrow> ennreal"
-  assumes u: "u \<in> borel_measurable M"
-  assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f"
-  assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)"
-  assumes mult': "\<And>u c. c < top \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
-  assumes add: "\<And>u v. u \<in> borel_measurable M\<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < top) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < top) \<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. x \<in> space M \<Longrightarrow> U i x < top) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)"
-  shows "P u"
-  using u
-proof (induct rule: borel_measurable_implies_simple_function_sequence')
-  fix U assume 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) = u x"
-  have u_eq: "u = (SUP i. U i)"
-    using u sup by auto
-
-  have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < top"
-    using U by (auto simp: image_iff eq_commute)
-
-  from U have "\<And>i. U i \<in> borel_measurable M"
-    by (simp add: borel_measurable_simple_function)
-
-  show "P u"
-    unfolding u_eq
-  proof (rule seq)
-    fix i show "P (U i)"
-      using \<open>simple_function M (U i)\<close> not_inf[of _ i]
-    proof (induct rule: simple_function_induct_nn)
-      case (mult u c)
-      show ?case
-      proof cases
-        assume "c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0)"
-        with mult(1) show ?thesis
-          by (intro cong[of "\<lambda>x. c * u x" "indicator {}"] set)
-             (auto dest!: borel_measurable_simple_function)
-      next
-        assume "\<not> (c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0))"
-        then obtain x where "space M \<noteq> {}" and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0"
-          by auto
-        with mult(3)[of x] have "c < top"
-          by (auto simp: ennreal_mult_less_top)
-        then have u_fin: "x' \<in> space M \<Longrightarrow> u x' < top" for x'
-          using mult(3)[of x'] \<open>c \<noteq> 0\<close> by (auto simp: ennreal_mult_less_top)
-        then have "P u"
-          by (rule mult)
-        with u_fin \<open>c < top\<close> mult(1) show ?thesis
-          by (intro mult') (auto dest!: borel_measurable_simple_function)
-      qed
-    qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function)
-  qed fact+
-qed
-
-lemma simple_function_If_set:
-  assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
-  shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
-proof -
-  define F where "F x = f -` {x} \<inter> space M" for x
-  define G where "G x = g -` {x} \<inter> space M" for x
-  show ?thesis unfolding simple_function_def
-  proof safe
-    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
-    from finite_subset[OF this] assms
-    show "finite (?IF ` space M)" unfolding simple_function_def by auto
-  next
-    fix x assume "x \<in> space M"
-    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
-      then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M))))
-      else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))"
-      using sets.sets_into_space[OF A] by (auto split: if_split_asm simp: G_def F_def)
-    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
-      unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto
-    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
-  qed
-qed
-
-lemma simple_function_If:
-  assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
-  shows "simple_function M (\<lambda>x. if P x then f x else g x)"
-proof -
-  have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
-  with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
-qed
-
-lemma simple_function_subalgebra:
-  assumes "simple_function N f"
-  and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M"
-  shows "simple_function M f"
-  using assms unfolding simple_function_def by auto
-
-lemma simple_function_comp:
-  assumes T: "T \<in> measurable M M'"
-    and f: "simple_function M' f"
-  shows "simple_function M (\<lambda>x. f (T x))"
-proof (intro simple_function_def[THEN iffD2] conjI ballI)
-  have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
-    using T unfolding measurable_def by auto
-  then show "finite ((\<lambda>x. f (T x)) ` space M)"
-    using f unfolding simple_function_def by (auto intro: finite_subset)
-  fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M"
-  then have "i \<in> f ` space M'"
-    using T unfolding measurable_def by auto
-  then have "f -` {i} \<inter> space M' \<in> sets M'"
-    using f unfolding simple_function_def by auto
-  then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
-    using T unfolding measurable_def by auto
-  also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
-    using T unfolding measurable_def by auto
-  finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
-qed
-
-subsection "Simple integral"
-
-definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>S") where
-  "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
-
-syntax
-  "_simple_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
-
-translations
-  "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
-
-lemma simple_integral_cong:
-  assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
-  shows "integral\<^sup>S M f = integral\<^sup>S M g"
-proof -
-  have "f ` space M = g ` space M"
-    "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
-    using assms by (auto intro!: image_eqI)
-  thus ?thesis unfolding simple_integral_def by simp
-qed
-
-lemma simple_integral_const[simp]:
-  "(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)"
-proof (cases "space M = {}")
-  case True thus ?thesis unfolding simple_integral_def by simp
-next
-  case False hence "(\<lambda>x. c) ` space M = {c}" by auto
-  thus ?thesis unfolding simple_integral_def by simp
-qed
-
-lemma simple_function_partition:
-  assumes f: "simple_function M f" and g: "simple_function M g"
-  assumes sub: "\<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 "integral\<^sup>S M f = (\<Sum>y\<in>g ` space M. v y * emeasure M {x\<in>space M. g x = y})"
-    (is "_ = ?r")
-proof -
-  from f g have [simp]: "finite (f`space M)" "finite (g`space M)"
-    by (auto simp: simple_function_def)
-  from f g have [measurable]: "f \<in> measurable M (count_space UNIV)" "g \<in> measurable M (count_space UNIV)"
-    by (auto intro: measurable_simple_function)
-
-  { 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 "integral\<^sup>S M f =
-    (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M.
-      if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
-    unfolding simple_integral_def
-  proof (safe intro!: setsum.cong ennreal_mult_left_cong)
-    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:"(\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i}) =
-        f -` {f y} \<inter> space M"
-      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
-    then show "emeasure M (f -` {f y} \<inter> space M) =
-      (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then emeasure M {x \<in> space M. g x = z} else 0)"
-      apply (simp add: setsum.If_cases)
-      apply (subst setsum_emeasure)
-      apply (auto simp: disjoint_family_on_def eq)
-      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 y * emeasure M {x\<in>space M. g x = z} else 0))"
-    by (auto intro!: setsum.cong simp: setsum_right_distrib)
-  also have "\<dots> = ?r"
-    by (subst setsum.commute)
-       (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq)
-  finally show "integral\<^sup>S M f = ?r" .
-qed
-
-lemma simple_integral_add[simp]:
-  assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
-  shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g"
-proof -
-  have "(\<integral>\<^sup>Sx. f x + g x \<partial>M) =
-    (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\<in>space M. (f x, g x) = y})"
-    by (intro simple_function_partition) (auto intro: f g)
-  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
-    (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
-    using assms(2,4) by (auto intro!: setsum.cong distrib_right simp: setsum.distrib[symmetric])
-  also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
-    by (intro simple_function_partition[symmetric]) (auto intro: f g)
-  also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
-    by (intro simple_function_partition[symmetric]) (auto intro: f g)
-  finally show ?thesis .
-qed
-
-lemma simple_integral_setsum[simp]:
-  assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
-  assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
-  shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))"
-proof cases
-  assume "finite P"
-  from this assms show ?thesis
-    by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
-qed auto
-
-lemma simple_integral_mult[simp]:
-  assumes f: "simple_function M f"
-  shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f"
-proof -
-  have "(\<integral>\<^sup>Sx. c * f x \<partial>M) = (\<Sum>y\<in>f ` space M. (c * y) * emeasure M {x\<in>space M. f x = y})"
-    using f by (intro simple_function_partition) auto
-  also have "\<dots> = c * integral\<^sup>S M f"
-    using f unfolding simple_integral_def
-    by (subst setsum_right_distrib) (auto simp: mult.assoc Int_def conj_commute)
-  finally show ?thesis .
-qed
-
-lemma simple_integral_mono_AE:
-  assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g"
-  and mono: "AE x in M. f x \<le> g x"
-  shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
-proof -
-  let ?\<mu> = "\<lambda>P. emeasure M {x\<in>space M. P x}"
-  have "integral\<^sup>S M f = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
-    using f g by (intro simple_function_partition) auto
-  also have "\<dots> \<le> (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * ?\<mu> (\<lambda>x. (f x, g x) = y))"
-  proof (clarsimp intro!: setsum_mono)
-    fix x assume "x \<in> space M"
-    let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)"
-    show "f x * ?M \<le> g x * ?M"
-    proof cases
-      assume "?M \<noteq> 0"
-      then have "0 < ?M"
-        by (simp add: less_le)
-      also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)"
-        using mono by (intro emeasure_mono_AE) auto
-      finally have "\<not> \<not> f x \<le> g x"
-        by (intro notI) auto
-      then show ?thesis
-        by (intro mult_right_mono) auto
-    qed simp
-  qed
-  also have "\<dots> = integral\<^sup>S M g"
-    using f g by (intro simple_function_partition[symmetric]) auto
-  finally show ?thesis .
-qed
-
-lemma simple_integral_mono:
-  assumes "simple_function M f" and "simple_function M g"
-  and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
-  shows "integral\<^sup>S M f \<le> integral\<^sup>S M g"
-  using assms by (intro simple_integral_mono_AE) auto
-
-lemma simple_integral_cong_AE:
-  assumes "simple_function M f" and "simple_function M g"
-  and "AE x in M. f x = g x"
-  shows "integral\<^sup>S M f = integral\<^sup>S M g"
-  using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
-
-lemma simple_integral_cong':
-  assumes sf: "simple_function M f" "simple_function M g"
-  and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
-  shows "integral\<^sup>S M f = integral\<^sup>S M g"
-proof (intro simple_integral_cong_AE sf AE_I)
-  show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
-  show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
-    using sf[THEN borel_measurable_simple_function] by auto
-qed simp
-
-lemma simple_integral_indicator:
-  assumes A: "A \<in> sets M"
-  assumes f: "simple_function M f"
-  shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
-    (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
-proof -
-  have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ennreal))`A"
-    using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: if_split_asm)
-  have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
-    by (auto simp: image_iff)
-
-  have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) =
-    (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\<in>space M. (f x, indicator A x) = y})"
-    using assms by (intro simple_function_partition) auto
-  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ennreal))`space M.
-    if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
-    by (auto simp: indicator_def split: if_split_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong)
-  also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ennreal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
-    using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq)
-  also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ennreal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
-    by (subst setsum.reindex [of fst]) (auto simp: inj_on_def)
-  also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
-    using A[THEN sets.sets_into_space]
-    by (intro setsum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2)
-  finally show ?thesis .
-qed
-
-lemma simple_integral_indicator_only[simp]:
-  assumes "A \<in> sets M"
-  shows "integral\<^sup>S M (indicator A) = emeasure M A"
-  using simple_integral_indicator[OF assms, of "\<lambda>x. 1"] sets.sets_into_space[OF assms]
-  by (simp_all add: image_constant_conv Int_absorb1 split: if_split_asm)
-
-lemma simple_integral_null_set:
-  assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M"
-  shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0"
-proof -
-  have "AE x in M. indicator N x = (0 :: ennreal)"
-    using \<open>N \<in> null_sets M\<close> by (auto simp: indicator_def intro!: AE_I[of _ _ N])
-  then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)"
-    using assms apply (intro simple_integral_cong_AE) by auto
-  then show ?thesis by simp
-qed
-
-lemma simple_integral_cong_AE_mult_indicator:
-  assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M"
-  shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)"
-  using assms by (intro simple_integral_cong_AE) auto
-
-lemma simple_integral_cmult_indicator:
-  assumes A: "A \<in> sets M"
-  shows "(\<integral>\<^sup>Sx. c * indicator A x \<partial>M) = c * emeasure M A"
-  using simple_integral_mult[OF simple_function_indicator[OF A]]
-  unfolding simple_integral_indicator_only[OF A] by simp
-
-lemma simple_integral_nonneg:
-  assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x"
-  shows "0 \<le> integral\<^sup>S M f"
-proof -
-  have "integral\<^sup>S M (\<lambda>x. 0) \<le> integral\<^sup>S M f"
-    using simple_integral_mono_AE[OF _ f ae] by auto
-  then show ?thesis by simp
-qed
-
-subsection \<open>Integral on nonnegative functions\<close>
-
-definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> ennreal" ("integral\<^sup>N") where
-  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> f}. integral\<^sup>S M g)"
-
-syntax
-  "_nn_integral" :: "pttrn \<Rightarrow> ennreal \<Rightarrow> 'a measure \<Rightarrow> ennreal" ("\<integral>\<^sup>+((2 _./ _)/ \<partial>_)" [60,61] 110)
-
-translations
-  "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)"
-
-lemma nn_integral_def_finite:
-  "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> f \<and> (\<forall>x. g x < top)}. integral\<^sup>S M g)"
-    (is "_ = SUPREMUM ?A ?f")
-  unfolding nn_integral_def
-proof (safe intro!: antisym SUP_least)
-  fix g assume g[measurable]: "simple_function M g" "g \<le> f"
-
-  show "integral\<^sup>S M g \<le> SUPREMUM ?A ?f"
-  proof cases
-    assume ae: "AE x in M. g x \<noteq> top"
-    let ?G = "{x \<in> space M. g x \<noteq> top}"
-    have "integral\<^sup>S M g = integral\<^sup>S M (\<lambda>x. g x * indicator ?G x)"
-    proof (rule simple_integral_cong_AE)
-      show "AE x in M. g x = g x * indicator ?G x"
-        using ae AE_space by eventually_elim auto
-    qed (insert g, auto)
-    also have "\<dots> \<le> SUPREMUM ?A ?f"
-      using g by (intro SUP_upper) (auto simp: le_fun_def less_top split: split_indicator)
-    finally show ?thesis .
-  next
-    assume nAE: "\<not> (AE x in M. g x \<noteq> top)"
-    then have "emeasure M {x\<in>space M. g x = top} \<noteq> 0" (is "emeasure M ?G \<noteq> 0")
-      by (subst (asm) AE_iff_measurable[OF _ refl]) auto
-    then have "top = (SUP n. (\<integral>\<^sup>Sx. of_nat n * indicator ?G x \<partial>M))"
-      by (simp add: ennreal_SUP_of_nat_eq_top ennreal_top_eq_mult_iff SUP_mult_right_ennreal[symmetric])
-    also have "\<dots> \<le> SUPREMUM ?A ?f"
-      using g
-      by (safe intro!: SUP_least SUP_upper)
-         (auto simp: le_fun_def of_nat_less_top top_unique[symmetric] split: split_indicator
-               intro: order_trans[of _ "g x" "f x" for x, OF order_trans[of _ top]])
-    finally show ?thesis
-      by (simp add: top_unique del: SUP_eq_top_iff Sup_eq_top_iff)
-  qed
-qed (auto intro: SUP_upper)
-
-lemma nn_integral_mono_AE:
-  assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v"
-  unfolding nn_integral_def
-proof (safe intro!: SUP_mono)
-  fix n assume n: "simple_function M n" "n \<le> u"
-  from ae[THEN AE_E] guess N . note N = this
-  then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in)
-  let ?n = "\<lambda>x. n x * indicator (space M - N) x"
-  have "AE x in M. n x \<le> ?n x" "simple_function M ?n"
-    using n N ae_N by auto
-  moreover
-  { fix x have "?n x \<le> v x"
-    proof cases
-      assume x: "x \<in> space M - N"
-      with N have "u x \<le> v x" by auto
-      with n(2)[THEN le_funD, of x] x show ?thesis
-        by (auto simp: max_def split: if_split_asm)
-    qed simp }
-  then have "?n \<le> v" by (auto simp: le_funI)
-  moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n"
-    using ae_N N n by (auto intro!: simple_integral_mono_AE)
-  ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
-    by force
-qed
-
-lemma nn_integral_mono:
-  "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>N M u \<le> integral\<^sup>N M v"
-  by (auto intro: nn_integral_mono_AE)
-
-lemma mono_nn_integral: "mono F \<Longrightarrow> mono (\<lambda>x. integral\<^sup>N M (F x))"
-  by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono)
-
-lemma nn_integral_cong_AE:
-  "AE x in M. u x = v x \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
-  by (auto simp: eq_iff intro!: nn_integral_mono_AE)
-
-lemma nn_integral_cong:
-  "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
-  by (auto intro: nn_integral_cong_AE)
-
-lemma nn_integral_cong_simp:
-  "(\<And>x. x \<in> space M =simp=> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v"
-  by (auto intro: nn_integral_cong simp: simp_implies_def)
-
-lemma nn_integral_cong_strong:
-  "M = N \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N N v"
-  by (auto intro: nn_integral_cong)
-
-lemma incseq_nn_integral:
-  assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))"
-proof -
-  have "\<And>i x. f i x \<le> f (Suc i) x"
-    using assms by (auto dest!: incseq_SucD simp: le_fun_def)
-  then show ?thesis
-    by (auto intro!: incseq_SucI nn_integral_mono)
-qed
-
-lemma nn_integral_eq_simple_integral:
-  assumes f: "simple_function M f" shows "integral\<^sup>N M f = integral\<^sup>S M f"
-proof -
-  let ?f = "\<lambda>x. f x * indicator (space M) x"
-  have f': "simple_function M ?f" using f by auto
-  have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f'
-    by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def)
-  moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f"
-    unfolding nn_integral_def
-    using f' by (auto intro!: SUP_upper)
-  ultimately show ?thesis
-    by (simp cong: nn_integral_cong simple_integral_cong)
-qed
-
-text \<open>Beppo-Levi monotone convergence theorem\<close>
-lemma nn_integral_monotone_convergence_SUP:
-  assumes f: "incseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
-proof (rule antisym)
-  show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
-    unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"]
-  proof (safe intro!: SUP_least)
-    fix u assume sf_u[simp]: "simple_function M u" and
-      u: "u \<le> (\<lambda>x. SUP i. f i x)" and u_range: "\<forall>x. u x < top"
-    note sf_u[THEN borel_measurable_simple_function, measurable]
-    show "integral\<^sup>S M u \<le> (SUP j. \<integral>\<^sup>+x. f j x \<partial>M)"
-    proof (rule ennreal_approx_unit)
-      fix a :: ennreal assume "a < 1"
-      let ?au = "\<lambda>x. a * u x"
-
-      let ?B = "\<lambda>c i. {x\<in>space M. ?au x = c \<and> c \<le> f i x}"
-      have "integral\<^sup>S M ?au = (\<Sum>c\<in>?au`space M. c * (SUP i. emeasure M (?B c i)))"
-        unfolding simple_integral_def
-      proof (intro setsum.cong ennreal_mult_left_cong refl)
-        fix c assume "c \<in> ?au ` space M" "c \<noteq> 0"
-        { fix x' assume x': "x' \<in> space M" "?au x' = c"
-          with \<open>c \<noteq> 0\<close> u_range have "?au x' < 1 * u x'"
-            by (intro ennreal_mult_strict_right_mono \<open>a < 1\<close>) (auto simp: less_le)
-          also have "\<dots> \<le> (SUP i. f i x')"
-            using u by (auto simp: le_fun_def)
-          finally have "\<exists>i. ?au x' \<le> f i x'"
-            by (auto simp: less_SUP_iff intro: less_imp_le) }
-        then have *: "?au -` {c} \<inter> space M = (\<Union>i. ?B c i)"
-          by auto
-        show "emeasure M (?au -` {c} \<inter> space M) = (SUP i. emeasure M (?B c i))"
-          unfolding * using f
-          by (intro SUP_emeasure_incseq[symmetric])
-             (auto simp: incseq_def le_fun_def intro: order_trans)
-      qed
-      also have "\<dots> = (SUP i. \<Sum>c\<in>?au`space M. c * emeasure M (?B c i))"
-        unfolding SUP_mult_left_ennreal using f
-        by (intro ennreal_SUP_setsum[symmetric])
-           (auto intro!: mult_mono emeasure_mono simp: incseq_def le_fun_def intro: order_trans)
-      also have "\<dots> \<le> (SUP i. integral\<^sup>N M (f i))"
-      proof (intro SUP_subset_mono order_refl)
-        fix i
-        have "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) =
-          (\<integral>\<^sup>Sx. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
-          by (subst simple_integral_indicator)
-             (auto intro!: setsum.cong ennreal_mult_left_cong arg_cong2[where f=emeasure])
-        also have "\<dots> = (\<integral>\<^sup>+x. (a * u x) * indicator {x\<in>space M. a * u x \<le> f i x} x \<partial>M)"
-          by (rule nn_integral_eq_simple_integral[symmetric]) simp
-        also have "\<dots> \<le> (\<integral>\<^sup>+x. f i x \<partial>M)"
-          by (intro nn_integral_mono) (auto split: split_indicator)
-        finally show "(\<Sum>c\<in>?au`space M. c * emeasure M (?B c i)) \<le> (\<integral>\<^sup>+x. f i x \<partial>M)" .
-      qed
-      finally show "a * integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))"
-        by simp
-    qed
-  qed
-qed (auto intro!: SUP_least SUP_upper nn_integral_mono)
-
-lemma sup_continuous_nn_integral[order_continuous_intros]:
-  assumes f: "\<And>y. sup_continuous (f y)"
-  assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
-  shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
-  unfolding sup_continuous_def
-proof safe
-  fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C"
-  with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (SUPREMUM UNIV C) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
-    unfolding sup_continuousD[OF f C]
-    by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def)
-qed
-
-lemma nn_integral_monotone_convergence_SUP_AE:
-  assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x" "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))"
-proof -
-  from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x"
-    by (simp add: AE_all_countable)
-  from this[THEN AE_E] guess N . note N = this
-  let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0"
-  have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N])
-  then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)"
-    by (auto intro!: nn_integral_cong_AE)
-  also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))"
-  proof (rule nn_integral_monotone_convergence_SUP)
-    show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI)
-    { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
-        using f N(3) by (intro measurable_If_set) auto }
-  qed
-  also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))"
-    using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] nn_integral_cong_AE ext)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_monotone_convergence_simple:
-  "incseq f \<Longrightarrow> (\<And>i. simple_function M (f i)) \<Longrightarrow> (SUP i. \<integral>\<^sup>S x. f i x \<partial>M) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
-  using nn_integral_monotone_convergence_SUP[of f M]
-  by (simp add: nn_integral_eq_simple_integral[symmetric] borel_measurable_simple_function)
-
-lemma SUP_simple_integral_sequences:
-  assumes f: "incseq f" "\<And>i. simple_function M (f i)"
-  and g: "incseq g" "\<And>i. simple_function M (g i)"
-  and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)"
-  shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))"
-    (is "SUPREMUM _ ?F = SUPREMUM _ ?G")
-proof -
-  have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)"
-    using f by (rule nn_integral_monotone_convergence_simple)
-  also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)"
-    unfolding eq[THEN nn_integral_cong_AE] ..
-  also have "\<dots> = (SUP i. ?G i)"
-    using g by (rule nn_integral_monotone_convergence_simple[symmetric])
-  finally show ?thesis by simp
-qed
-
-lemma nn_integral_const[simp]: "(\<integral>\<^sup>+ x. c \<partial>M) = c * emeasure M (space M)"
-  by (subst nn_integral_eq_simple_integral) auto
-
-lemma nn_integral_linear:
-  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g"
-    (is "integral\<^sup>N M ?L = _")
-proof -
-  from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u .
-  note u = nn_integral_monotone_convergence_simple[OF this(2,1)] this
-  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v .
-  note v = nn_integral_monotone_convergence_simple[OF this(2,1)] this
-  let ?L' = "\<lambda>i x. a * u i x + v i x"
-
-  have "?L \<in> borel_measurable M" using assms by auto
-  from borel_measurable_implies_simple_function_sequence'[OF this] guess l .
-  note l = nn_integral_monotone_convergence_simple[OF this(2,1)] this
-
-  have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))"
-    using u v by (auto simp: incseq_Suc_iff le_fun_def intro!: add_mono mult_left_mono simple_integral_mono)
-
-  have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))"
-  proof (rule SUP_simple_integral_sequences[OF l(3,2)])
-    show "incseq ?L'" "\<And>i. simple_function M (?L' i)"
-      using u v unfolding incseq_Suc_iff le_fun_def
-      by (auto intro!: add_mono mult_left_mono)
-    { fix x
-      have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)"
-        using u(3) v(3) u(4)[of _ x] v(4)[of _ x] unfolding SUP_mult_left_ennreal
-        by (auto intro!: ennreal_SUP_add simp: incseq_Suc_iff le_fun_def add_mono mult_left_mono) }
-    then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)"
-      unfolding l(5) using u(5) v(5) by (intro AE_I2) auto
-  qed
-  also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))"
-    using u(2) v(2) by auto
-  finally show ?thesis
-    unfolding l(5)[symmetric] l(1)[symmetric]
-    by (simp add: ennreal_SUP_add[OF inc] v u SUP_mult_left_ennreal[symmetric])
-qed
-
-lemma nn_integral_cmult: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f"
-  using nn_integral_linear[of f M "\<lambda>x. 0" c] by simp
-
-lemma nn_integral_multc: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c"
-  unfolding mult.commute[of _ c] nn_integral_cmult by simp
-
-lemma nn_integral_divide: "f \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c"
-   unfolding divide_ennreal_def by (rule nn_integral_multc)
-
-lemma nn_integral_indicator[simp]: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A"
-  by (subst nn_integral_eq_simple_integral) (auto simp: simple_integral_indicator)
-
-lemma nn_integral_cmult_indicator: "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * emeasure M A"
-  by (subst nn_integral_eq_simple_integral)
-     (auto simp: simple_function_indicator simple_integral_indicator)
-
-lemma nn_integral_indicator':
-  assumes [measurable]: "A \<inter> space M \<in> sets M"
-  shows "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)"
-proof -
-  have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)"
-    by (intro nn_integral_cong) (simp split: split_indicator)
-  also have "\<dots> = emeasure M (A \<inter> space M)"
-    by simp
-  finally show ?thesis .
-qed
-
-lemma nn_integral_indicator_singleton[simp]:
-  assumes [measurable]: "{y} \<in> sets M" shows "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = f y * emeasure M {y}"
-proof -
-  have "(\<integral>\<^sup>+x. f x * indicator {y} x \<partial>M) = (\<integral>\<^sup>+x. f y * indicator {y} x \<partial>M)"
-    by (auto intro!: nn_integral_cong split: split_indicator)
-  then show ?thesis
-    by (simp add: nn_integral_cmult)
-qed
-
-lemma nn_integral_set_ennreal:
-  "(\<integral>\<^sup>+x. ennreal (f x) * indicator A x \<partial>M) = (\<integral>\<^sup>+x. ennreal (f x * indicator A x) \<partial>M)"
-  by (rule nn_integral_cong) (simp split: split_indicator)
-
-lemma nn_integral_indicator_singleton'[simp]:
-  assumes [measurable]: "{y} \<in> sets M"
-  shows "(\<integral>\<^sup>+x. ennreal (f x * indicator {y} x) \<partial>M) = f y * emeasure M {y}"
-  by (subst nn_integral_set_ennreal[symmetric]) (simp add: nn_integral_indicator_singleton)
-
-lemma nn_integral_add:
-  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g"
-  using nn_integral_linear[of f M g 1] by simp
-
-lemma nn_integral_setsum:
-  "(\<And>i. i \<in> P \<Longrightarrow> f i \<in> borel_measurable M) \<Longrightarrow> (\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))"
-  by (induction P rule: infinite_finite_induct) (auto simp: nn_integral_add)
-
-lemma nn_integral_suminf:
-  assumes f: "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))"
-proof -
-  have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
-    using assms by (auto simp: AE_all_countable)
-  have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))"
-    by (rule suminf_eq_SUP)
-  also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)"
-    unfolding nn_integral_setsum[OF f] ..
-  also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos
-    by (intro nn_integral_monotone_convergence_SUP_AE[symmetric])
-       (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3)
-  also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos
-    by (intro nn_integral_cong_AE) (auto simp: suminf_eq_SUP)
-  finally show ?thesis by simp
-qed
-
-lemma nn_integral_bound_simple_function:
-  assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>"
-  assumes f[measurable]: "simple_function M f"
-  assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
-  shows "nn_integral M f < \<infinity>"
-proof cases
-  assume "space M = {}"
-  then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)"
-    by (intro nn_integral_cong) auto
-  then show ?thesis by simp
-next
-  assume "space M \<noteq> {}"
-  with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>"
-    by (subst Max_less_iff) (auto simp: Max_ge_iff)
-
-  have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)"
-  proof (rule nn_integral_mono)
-    fix x assume "x \<in> space M"
-    with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x"
-      by (auto split: split_indicator intro!: Max_ge simple_functionD)
-  qed
-  also have "\<dots> < \<infinity>"
-    using bnd supp by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_Markov_inequality:
-  assumes u: "u \<in> borel_measurable M" and "A \<in> sets M"
-  shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
-    (is "(emeasure M) ?A \<le> _ * ?PI")
-proof -
-  have "?A \<in> sets M"
-    using \<open>A \<in> sets M\<close> u by auto
-  hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)"
-    using nn_integral_indicator by simp
-  also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)"
-    using u by (auto intro!: nn_integral_mono_AE simp: indicator_def)
-  also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
-    using assms by (auto intro!: nn_integral_cmult)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_noteq_infinite:
-  assumes g: "g \<in> borel_measurable M" and "integral\<^sup>N M g \<noteq> \<infinity>"
-  shows "AE x in M. g x \<noteq> \<infinity>"
-proof (rule ccontr)
-  assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)"
-  have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
-    using c g by (auto simp add: AE_iff_null)
-  then have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}"
-    by (auto simp: zero_less_iff_neq_zero)
-  then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}"
-    by (auto simp: ennreal_top_eq_mult_iff)
-  also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
-    using g by (subst nn_integral_cmult_indicator) auto
-  also have "\<dots> \<le> integral\<^sup>N M g"
-    using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def)
-  finally show False
-    using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by (auto simp: top_unique)
-qed
-
-lemma nn_integral_PInf:
-  assumes f: "f \<in> borel_measurable M" and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>"
-  shows "emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
-proof -
-  have "\<infinity> * emeasure M (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
-    using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets)
-  also have "\<dots> \<le> integral\<^sup>N M f"
-    by (auto intro!: nn_integral_mono simp: indicator_def)
-  finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
-    by simp
-  then show ?thesis
-    using assms by (auto simp: ennreal_top_mult top_unique split: if_split_asm)
-qed
-
-lemma simple_integral_PInf:
-  "simple_function M f \<Longrightarrow> integral\<^sup>S M f \<noteq> \<infinity> \<Longrightarrow> emeasure M (f -` {\<infinity>} \<inter> space M) = 0"
-  by (rule nn_integral_PInf) (auto simp: nn_integral_eq_simple_integral borel_measurable_simple_function)
-
-lemma nn_integral_PInf_AE:
-  assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
-proof (rule AE_I)
-  show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
-    by (rule nn_integral_PInf[OF assms])
-  show "f -` {\<infinity>} \<inter> space M \<in> sets M"
-    using assms by (auto intro: borel_measurable_vimage)
-qed auto
-
-lemma nn_integral_diff:
-  assumes f: "f \<in> borel_measurable M"
-  and g: "g \<in> borel_measurable M"
-  and fin: "integral\<^sup>N M g \<noteq> \<infinity>"
-  and mono: "AE x in M. g x \<le> f x"
-  shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g"
-proof -
-  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-    using assms by auto
-  have "AE x in M. f x = f x - g x + g x"
-    using diff_add_cancel_ennreal mono nn_integral_noteq_infinite[OF g fin] assms by auto
-  then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g"
-    unfolding nn_integral_add[OF diff g, symmetric]
-    by (rule nn_integral_cong_AE)
-  show ?thesis unfolding **
-    using fin
-    by (cases rule: ennreal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto
-qed
-
-lemma nn_integral_mult_bounded_inf:
-  assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" and c: "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x"
-  shows "(\<integral>\<^sup>+x. g x \<partial>M) < \<infinity>"
-proof -
-  have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)"
-    by (intro nn_integral_mono_AE ae)
-  also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>"
-    using c f by (subst nn_integral_cmult) (auto simp: ennreal_mult_less_top top_unique not_less)
-  finally show ?thesis .
-qed
-
-text \<open>Fatou's lemma: convergence theorem on limes inferior\<close>
-
-lemma nn_integral_monotone_convergence_INF_AE':
-  assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
-    and *: "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>"
-  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
-proof (rule ennreal_minus_cancel)
-  have "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+x. f 0 x - (INF i. f i x) \<partial>M)"
-  proof (rule nn_integral_diff[symmetric])
-    have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
-      by (intro nn_integral_mono INF_lower) simp
-    with * show "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) \<noteq> \<infinity>"
-      by simp
-  qed (auto intro: INF_lower)
-  also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. f 0 x - f i x) \<partial>M)"
-    by (simp add: ennreal_INF_const_minus)
-  also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f 0 x - f i x \<partial>M))"
-  proof (intro nn_integral_monotone_convergence_SUP_AE)
-    show "AE x in M. f 0 x - f i x \<le> f 0 x - f (Suc i) x" for i
-      using f[of i] by eventually_elim (auto simp: ennreal_mono_minus)
-  qed simp
-  also have "\<dots> = (SUP i. nn_integral M (f 0) - (\<integral>\<^sup>+x. f i x \<partial>M))"
-  proof (subst nn_integral_diff[symmetric])
-    fix i
-    have dec: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
-      unfolding AE_all_countable using f by auto
-    then show "AE x in M. f i x \<le> f 0 x"
-      using dec by eventually_elim (auto intro: lift_Suc_antimono_le[of "\<lambda>i. f i x" 0 i for x])
-    then have "(\<integral>\<^sup>+ x. f i x \<partial>M) \<le> (\<integral>\<^sup>+ x. f 0 x \<partial>M)"
-      by (rule nn_integral_mono_AE)
-    with * show "(\<integral>\<^sup>+ x. f i x \<partial>M) \<noteq> \<infinity>"
-      by simp
-  qed (insert f, auto simp: decseq_def le_fun_def)
-  finally show "integral\<^sup>N M (f 0) - (\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) =
-    integral\<^sup>N M (f 0) - (INF i. \<integral>\<^sup>+ x. f i x \<partial>M)"
-    by (simp add: ennreal_INF_const_minus)
-qed (insert *, auto intro!: nn_integral_mono intro: INF_lower)
-
-lemma nn_integral_monotone_convergence_INF_AE:
-  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
-  assumes f: "\<And>i. AE x in M. f (Suc i) x \<le> f i x"
-    and [measurable]: "\<And>i. f i \<in> borel_measurable M"
-    and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
-  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
-proof -
-  { fix f :: "nat \<Rightarrow> ennreal" and j assume "decseq f"
-    then have "(INF i. f i) = (INF i. f (i + j))"
-      apply (intro INF_eq)
-      apply (rule_tac x="i" in bexI)
-      apply (auto simp: decseq_def le_fun_def)
-      done }
-  note INF_shift = this
-  have mono: "AE x in M. \<forall>i. f (Suc i) x \<le> f i x"
-    using f by (auto simp: AE_all_countable)
-  then have "AE x in M. (INF i. f i x) = (INF n. f (n + i) x)"
-    by eventually_elim (auto intro!: decseq_SucI INF_shift)
-  then have "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF n. f (n + i) x) \<partial>M)"
-    by (rule nn_integral_cong_AE)
-  also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f (n + i) x \<partial>M))"
-    by (rule nn_integral_monotone_convergence_INF_AE') (insert assms, auto)
-  also have "\<dots> = (INF n. (\<integral>\<^sup>+ x. f n x \<partial>M))"
-    by (intro INF_shift[symmetric] decseq_SucI nn_integral_mono_AE f)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_monotone_convergence_INF_decseq:
-  assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
-  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
-  using nn_integral_monotone_convergence_INF_AE[of f M i, OF _ *] f by (auto simp: decseq_Suc_iff le_fun_def)
-
-lemma nn_integral_liminf:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
-  assumes u: "\<And>i. u i \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
-proof -
-  have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (SUP n. \<integral>\<^sup>+ x. (INF i:{n..}. u i x) \<partial>M)"
-    unfolding liminf_SUP_INF using u
-    by (intro nn_integral_monotone_convergence_SUP_AE)
-       (auto intro!: AE_I2 intro: INF_greatest INF_superset_mono)
-  also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
-    by (auto simp: liminf_SUP_INF intro!: SUP_mono INF_greatest nn_integral_mono INF_lower)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_limsup:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ennreal"
-  assumes [measurable]: "\<And>i. u i \<in> borel_measurable M" "w \<in> borel_measurable M"
-  assumes bounds: "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>"
-  shows "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
-proof -
-  have bnd: "AE x in M. \<forall>i. u i x \<le> w x"
-    using bounds by (auto simp: AE_all_countable)
-  then have "(\<integral>\<^sup>+ x. (SUP n. u n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. w x \<partial>M)"
-    by (auto intro!: nn_integral_mono_AE elim: eventually_mono intro: SUP_least)
-  then have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (INF n. \<integral>\<^sup>+ x. (SUP i:{n..}. u i x) \<partial>M)"
-    unfolding limsup_INF_SUP using bnd w
-    by (intro nn_integral_monotone_convergence_INF_AE')
-       (auto intro!: AE_I2 intro: SUP_least SUP_subset_mono)
-  also have "\<dots> \<ge> limsup (\<lambda>n. integral\<^sup>N M (u n))"
-    by (auto simp: limsup_INF_SUP intro!: INF_mono SUP_least exI nn_integral_mono SUP_upper)
-  finally (xtrans) show ?thesis .
-qed
-
-lemma nn_integral_LIMSEQ:
-  assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M"
-    and u: "\<And>x. (\<lambda>i. f i x) \<longlonglongrightarrow> u x"
-  shows "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> integral\<^sup>N M u"
-proof -
-  have "(\<lambda>n. integral\<^sup>N M (f n)) \<longlonglongrightarrow> (SUP n. integral\<^sup>N M (f n))"
-    using f by (intro LIMSEQ_SUP[of "\<lambda>n. integral\<^sup>N M (f n)"] incseq_nn_integral)
-  also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)"
-    using f by (intro nn_integral_monotone_convergence_SUP[symmetric])
-  also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)"
-    using f by (subst LIMSEQ_SUP[THEN LIMSEQ_unique, OF _ u]) (auto simp: incseq_def le_fun_def)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_dominated_convergence:
-  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. 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. u i x \<partial>M)) \<longlonglongrightarrow> (\<integral>\<^sup>+x. u' x \<partial>M)"
-proof -
-  have "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)"
-    by (intro nn_integral_limsup[OF _ _ bound w]) auto
-  moreover have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
-    using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
-  moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)"
-    using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot)
-  moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))"
-    by (intro nn_integral_liminf) auto
-  moreover have "liminf (\<lambda>n. integral\<^sup>N M (u n)) \<le> limsup (\<lambda>n. integral\<^sup>N M (u n))"
-    by (intro Liminf_le_Limsup sequentially_bot)
-  ultimately show ?thesis
-    by (intro Liminf_eq_Limsup) auto
-qed
-
-lemma inf_continuous_nn_integral[order_continuous_intros]:
-  assumes f: "\<And>y. inf_continuous (f y)"
-  assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
-  assumes bnd: "\<And>x. (\<integral>\<^sup>+ y. f y x \<partial>M) \<noteq> \<infinity>"
-  shows "inf_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))"
-  unfolding inf_continuous_def
-proof safe
-  fix C :: "nat \<Rightarrow> 'b" assume C: "decseq C"
-  then show "(\<integral>\<^sup>+ y. f y (INFIMUM UNIV C) \<partial>M) = (INF i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)"
-    using inf_continuous_mono[OF f] bnd
-    by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def less_top
-             intro!: nn_integral_monotone_convergence_INF_decseq)
-qed
-
-lemma nn_integral_null_set:
-  assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0"
-proof -
-  have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
-  proof (intro nn_integral_cong_AE AE_I)
-    show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
-      by (auto simp: indicator_def)
-    show "(emeasure M) N = 0" "N \<in> sets M"
-      using assms by auto
-  qed
-  then show ?thesis by simp
-qed
-
-lemma nn_integral_0_iff:
-  assumes u: "u \<in> borel_measurable M"
-  shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
-    (is "_ \<longleftrightarrow> (emeasure M) ?A = 0")
-proof -
-  have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>N M u"
-    by (auto intro!: nn_integral_cong simp: indicator_def)
-  show ?thesis
-  proof
-    assume "(emeasure M) ?A = 0"
-    with nn_integral_null_set[of ?A M u] u
-    show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def)
-  next
-    assume *: "integral\<^sup>N M u = 0"
-    let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
-    have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))"
-    proof -
-      { fix n :: nat
-        from nn_integral_Markov_inequality[OF u, of ?A "of_nat n"] u
-        have "(emeasure M) (?M n \<inter> ?A) \<le> 0"
-          by (simp add: ennreal_of_nat_eq_real_of_nat u_eq *)
-        moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto
-        ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto }
-      thus ?thesis by simp
-    qed
-    also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)"
-    proof (safe intro!: SUP_emeasure_incseq)
-      fix n show "?M n \<inter> ?A \<in> sets M"
-        using u by (auto intro!: sets.Int)
-    next
-      show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
-      proof (safe intro!: incseq_SucI)
-        fix n :: nat and x
-        assume *: "1 \<le> real n * u x"
-        also have "real n * u x \<le> real (Suc n) * u x"
-          by (auto intro!: mult_right_mono)
-        finally show "1 \<le> real (Suc n) * u x" by auto
-      qed
-    qed
-    also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
-    proof (safe intro!: arg_cong[where f="(emeasure M)"])
-      fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
-      show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
-      proof (cases "u x" rule: ennreal_cases)
-        case (real r) with \<open>0 < u x\<close> have "0 < r" by auto
-        obtain j :: nat where "1 / r \<le> real j" using real_arch_simple ..
-        hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto
-        hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto
-        thus ?thesis using \<open>0 < r\<close> real
-          by (auto simp: ennreal_of_nat_eq_real_of_nat ennreal_1[symmetric] ennreal_mult[symmetric]
-                   simp del: ennreal_1)
-      qed (insert \<open>0 < u x\<close>, auto simp: ennreal_mult_top)
-    qed (auto simp: zero_less_iff_neq_zero)
-    finally show "emeasure M ?A = 0"
-      by (simp add: zero_less_iff_neq_zero)
-  qed
-qed
-
-lemma nn_integral_0_iff_AE:
-  assumes u: "u \<in> borel_measurable M"
-  shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
-proof -
-  have sets: "{x\<in>space M. u x \<noteq> 0} \<in> sets M"
-    using u by auto
-  show "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x = 0)"
-    using nn_integral_0_iff[of u] AE_iff_null[OF sets] u by auto
-qed
-
-lemma AE_iff_nn_integral:
-  "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>N M (indicator {x. \<not> P x}) = 0"
-  by (subst nn_integral_0_iff_AE) (auto simp: indicator_def[abs_def])
-
-lemma nn_integral_less:
-  assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes f: "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>"
-  assumes ord: "AE x in M. f x \<le> g x" "\<not> (AE x in M. g x \<le> f x)"
-  shows "(\<integral>\<^sup>+x. f x \<partial>M) < (\<integral>\<^sup>+x. g x \<partial>M)"
-proof -
-  have "0 < (\<integral>\<^sup>+x. g x - f x \<partial>M)"
-  proof (intro order_le_neq_trans notI)
-    assume "0 = (\<integral>\<^sup>+x. g x - f x \<partial>M)"
-    then have "AE x in M. g x - f x = 0"
-      using nn_integral_0_iff_AE[of "\<lambda>x. g x - f x" M] by simp
-    with ord(1) have "AE x in M. g x \<le> f x"
-      by eventually_elim (auto simp: ennreal_minus_eq_0)
-    with ord show False
-      by simp
-  qed simp
-  also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M) - (\<integral>\<^sup>+x. f x \<partial>M)"
-    using f by (subst nn_integral_diff) (auto simp: ord)
-  finally show ?thesis
-    using f by (auto dest!: ennreal_minus_pos_iff[rotated] simp: less_top)
-qed
-
-lemma nn_integral_subalgebra:
-  assumes f: "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>N N f = integral\<^sup>N M f"
-proof -
-  have [simp]: "\<And>f :: 'a \<Rightarrow> ennreal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M"
-    using N by (auto simp: measurable_def)
-  have [simp]: "\<And>P. (AE x in N. P x) \<Longrightarrow> (AE x in M. P x)"
-    using N by (auto simp add: eventually_ae_filter null_sets_def subset_eq)
-  have [simp]: "\<And>A. A \<in> sets N \<Longrightarrow> A \<in> sets M"
-    using N by auto
-  from f show ?thesis
-    apply induct
-    apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N)
-    apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric])
-    done
-qed
-
-lemma nn_integral_nat_function:
-  fixes f :: "'a \<Rightarrow> nat"
-  assumes "f \<in> measurable M (count_space UNIV)"
-  shows "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
-proof -
-  define F where "F i = {x\<in>space M. i < f x}" for i
-  with assms have [measurable]: "\<And>i. F i \<in> sets M"
-    by auto
-
-  { fix x assume "x \<in> space M"
-    have "(\<lambda>i. if i < f x then 1 else 0) sums (of_nat (f x)::real)"
-      using sums_If_finite[of "\<lambda>i. i < f x" "\<lambda>_. 1::real"] by simp
-    then have "(\<lambda>i. ennreal (if i < f x then 1 else 0)) sums of_nat(f x)"
-      unfolding ennreal_of_nat_eq_real_of_nat
-      by (subst sums_ennreal) auto
-    moreover have "\<And>i. ennreal (if i < f x then 1 else 0) = indicator (F i) x"
-      using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def)
-    ultimately have "of_nat (f x) = (\<Sum>i. indicator (F i) x :: ennreal)"
-      by (simp add: sums_iff) }
-  then have "(\<integral>\<^sup>+x. of_nat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
-    by (simp cong: nn_integral_cong)
-  also have "\<dots> = (\<Sum>i. emeasure M (F i))"
-    by (simp add: nn_integral_suminf)
-  finally show ?thesis
-    by (simp add: F_def)
-qed
-
-lemma nn_integral_lfp:
-  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
-  assumes f: "sup_continuous f"
-  assumes g: "sup_continuous g"
-  assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
-  assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
-  shows "(\<integral>\<^sup>+\<omega>. lfp f \<omega> \<partial>M s) = lfp g s"
-proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f and P="\<lambda>f. f \<in> borel_measurable N", symmetric])
-  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N"
-  then show "(\<lambda>s. \<integral>\<^sup>+x. (SUP i. C i) x \<partial>M s) = (SUP i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
-    unfolding SUP_apply[abs_def]
-    by (subst nn_integral_monotone_convergence_SUP)
-       (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
-qed (auto simp add: step le_fun_def SUP_apply[abs_def] bot_fun_def bot_ennreal intro!: meas f g)
-
-lemma nn_integral_gfp:
-  assumes sets[simp]: "\<And>s. sets (M s) = sets N"
-  assumes f: "inf_continuous f" and g: "inf_continuous g"
-  assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N"
-  assumes bound: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f F x \<partial>M s) < \<infinity>"
-  assumes non_zero: "\<And>s. emeasure (M s) (space (M s)) \<noteq> 0"
-  assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s"
-  shows "(\<integral>\<^sup>+\<omega>. gfp f \<omega> \<partial>M s) = gfp g s"
-proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f
-    and P="\<lambda>F. F \<in> borel_measurable N \<and> (\<forall>s. (\<integral>\<^sup>+x. F x \<partial>M s) < \<infinity>)", symmetric])
-  fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ennreal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)"
-  then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
-    unfolding INF_apply[abs_def]
-    by (subst nn_integral_monotone_convergence_INF_decseq)
-       (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
-next
-  show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))"
-    by (subst step)
-       (auto simp add: top_fun_def less_le non_zero le_fun_def ennreal_top_mult
-             cong del: if_weak_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD])
-next
-  fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C"
-  with bound show "INFIMUM UNIV C \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (INFIMUM UNIV C) < \<infinity>)"
-    unfolding INF_apply[abs_def]
-    by (subst nn_integral_monotone_convergence_INF_decseq)
-       (auto simp: INF_less_iff cong: measurable_cong_sets intro!: borel_measurable_INF)
-next
-  show "\<And>x. x \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) x < \<infinity>) \<Longrightarrow>
-         (\<lambda>s. integral\<^sup>N (M s) (f x)) = g (\<lambda>s. integral\<^sup>N (M s) x)"
-    by (subst step) auto
-qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g)
-
-subsection \<open>Integral under concrete measures\<close>
-
-lemma nn_integral_mono_measure:
-  assumes "sets M = sets N" "M \<le> N" shows "nn_integral M f \<le> nn_integral N f"
-  unfolding nn_integral_def
-proof (intro SUP_subset_mono)
-  note \<open>sets M = sets N\<close>[simp]  \<open>sets M = sets N\<close>[THEN sets_eq_imp_space_eq, simp]
-  show "{g. simple_function M g \<and> g \<le> f} \<subseteq> {g. simple_function N g \<and> g \<le> f}"
-    by (simp add: simple_function_def)
-  show "integral\<^sup>S M x \<le> integral\<^sup>S N x" for x
-    using le_measureD3[OF \<open>M \<le> N\<close>]
-    by (auto simp add: simple_integral_def intro!: setsum_mono mult_mono)
-qed
-
-lemma nn_integral_empty:
-  assumes "space M = {}"
-  shows "nn_integral M f = 0"
-proof -
-  have "(\<integral>\<^sup>+ x. f x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)"
-    by(rule nn_integral_cong)(simp add: assms)
-  thus ?thesis by simp
-qed
-
-lemma nn_integral_bot[simp]: "nn_integral bot f = 0"
-  by (simp add: nn_integral_empty)
-
-subsubsection \<open>Distributions\<close>
-
-lemma nn_integral_distr:
-  assumes T: "T \<in> measurable M M'" and f: "f \<in> borel_measurable (distr M M' T)"
-  shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)"
-  using f
-proof induct
-  case (cong f g)
-  with T show ?case
-    apply (subst nn_integral_cong[of _ f g])
-    apply simp
-    apply (subst nn_integral_cong[of _ "\<lambda>x. f (T x)" "\<lambda>x. g (T x)"])
-    apply (simp add: measurable_def Pi_iff)
-    apply simp
-    done
-next
-  case (set A)
-  then have eq: "\<And>x. x \<in> space M \<Longrightarrow> indicator A (T x) = indicator (T -` A \<inter> space M) x"
-    by (auto simp: indicator_def)
-  from set T show ?case
-    by (subst nn_integral_cong[OF eq])
-       (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets)
-qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add
-                   nn_integral_monotone_convergence_SUP le_fun_def incseq_def)
-
-subsubsection \<open>Counting space\<close>
-
-lemma simple_function_count_space[simp]:
-  "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
-  unfolding simple_function_def by simp
-
-lemma nn_integral_count_space:
-  assumes A: "finite {a\<in>A. 0 < f a}"
-  shows "integral\<^sup>N (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
-proof -
-  have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) =
-    (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
-    by (auto intro!: nn_integral_cong
-             simp add: indicator_def if_distrib setsum.If_cases[OF A] max_def le_less)
-  also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
-    by (subst nn_integral_setsum) (simp_all add: AE_count_space  less_imp_le)
-  also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
-    by (auto intro!: setsum.cong simp: one_ennreal_def[symmetric] max_def)
-  finally show ?thesis by (simp add: max.absorb2)
-qed
-
-lemma nn_integral_count_space_finite:
-    "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
-  by (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le)
-
-lemma nn_integral_count_space':
-  assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "A \<subseteq> B"
-  shows "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>x\<in>A. f x)"
-proof -
-  have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 0 < f a. f a)"
-    using assms(2,3)
-    by (intro nn_integral_count_space finite_subset[OF _ \<open>finite A\<close>]) (auto simp: less_le)
-  also have "\<dots> = (\<Sum>a\<in>A. f a)"
-    using assms by (intro setsum.mono_neutral_cong_left) (auto simp: less_le)
-  finally show ?thesis .
-qed
-
-lemma nn_integral_bij_count_space:
-  assumes g: "bij_betw g A B"
-  shows "(\<integral>\<^sup>+x. f (g x) \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
-  using g[THEN bij_betw_imp_funcset]
-  by (subst distr_bij_count_space[OF g, symmetric])
-     (auto intro!: nn_integral_distr[symmetric])
-
-lemma nn_integral_indicator_finite:
-  fixes f :: "'a \<Rightarrow> ennreal"
-  assumes f: "finite A" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
-  shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
-proof -
-  from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
-    by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] setsum.If_cases)
-  also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
-    by (subst nn_integral_setsum) auto
-  finally show ?thesis .
-qed
-
-lemma nn_integral_count_space_nat:
-  fixes f :: "nat \<Rightarrow> ennreal"
-  shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)"
-proof -
-  have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) =
-    (\<integral>\<^sup>+i. (\<Sum>j. f j * indicator {j} i) \<partial>count_space UNIV)"
-  proof (intro nn_integral_cong)
-    fix i
-    have "f i = (\<Sum>j\<in>{i}. f j * indicator {j} i)"
-      by simp
-    also have "\<dots> = (\<Sum>j. f j * indicator {j} i)"
-      by (rule suminf_finite[symmetric]) auto
-    finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
-  qed
-  also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
-    by (rule nn_integral_suminf) auto
-  finally show ?thesis
-    by simp
-qed
-
-lemma nn_integral_enat_function:
-  assumes f: "f \<in> measurable M (count_space UNIV)"
-  shows "(\<integral>\<^sup>+ x. ennreal_of_enat (f x) \<partial>M) = (\<Sum>t. emeasure M {x \<in> space M. t < f x})"
-proof -
-  define F where "F i = {x\<in>space M. i < f x}" for i :: nat
-  with assms have [measurable]: "\<And>i. F i \<in> sets M"
-    by auto
-
-  { fix x assume "x \<in> space M"
-    have "(\<lambda>i::nat. if i < f x then 1 else 0) sums ennreal_of_enat (f x)"
-      using sums_If_finite[of "\<lambda>r. r < f x" "\<lambda>_. 1 :: ennreal"]
-      by (cases "f x") (simp_all add: sums_def of_nat_tendsto_top_ennreal)
-    also have "(\<lambda>i. (if i < f x then 1 else 0)) = (\<lambda>i. indicator (F i) x)"
-      using \<open>x \<in> space M\<close> by (simp add: one_ennreal_def F_def fun_eq_iff)
-    finally have "ennreal_of_enat (f x) = (\<Sum>i. indicator (F i) x)"
-      by (simp add: sums_iff) }
-  then have "(\<integral>\<^sup>+x. ennreal_of_enat (f x) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)"
-    by (simp cong: nn_integral_cong)
-  also have "\<dots> = (\<Sum>i. emeasure M (F i))"
-    by (simp add: nn_integral_suminf)
-  finally show ?thesis
-    by (simp add: F_def)
-qed
-
-lemma nn_integral_count_space_nn_integral:
-  fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ennreal"
-  assumes "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M"
-  shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)"
-proof cases
-  assume "finite I" then show ?thesis
-    by (simp add: nn_integral_count_space_finite nn_integral_setsum)
-next
-  assume "infinite I"
-  then have [simp]: "I \<noteq> {}"
-    by auto
-  note * = bij_betw_from_nat_into[OF \<open>countable I\<close> \<open>infinite I\<close>]
-  have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>n. f (from_nat_into I n))"
-    by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat)
-  show ?thesis
-    by (simp add: ** nn_integral_suminf from_nat_into)
-qed
-
-lemma emeasure_UN_countable:
-  assumes sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I[simp]: "countable I"
-  assumes disj: "disjoint_family_on X I"
-  shows "emeasure M (UNION I X) = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)"
-proof -
-  have eq: "\<And>x. indicator (UNION I X) x = \<integral>\<^sup>+ i. indicator (X i) x \<partial>count_space I"
-  proof cases
-    fix x assume x: "x \<in> UNION I X"
-    then obtain j where j: "x \<in> X j" "j \<in> I"
-      by auto
-    with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ennreal)"
-      by (auto simp: disjoint_family_on_def split: split_indicator)
-    with x j show "?thesis x"
-      by (simp cong: nn_integral_cong_simp)
-  qed (auto simp: nn_integral_0_iff_AE)
-
-  note sets.countable_UN'[unfolded subset_eq, measurable]
-  have "emeasure M (UNION I X) = (\<integral>\<^sup>+x. indicator (UNION I X) x \<partial>M)"
-    by simp
-  also have "\<dots> = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. indicator (X i) x \<partial>M \<partial>count_space I)"
-    by (simp add: eq nn_integral_count_space_nn_integral)
-  finally show ?thesis
-    by (simp cong: nn_integral_cong_simp)
-qed
-
-lemma emeasure_countable_singleton:
-  assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" and X: "countable X"
-  shows "emeasure M X = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
-proof -
-  have "emeasure M (\<Union>i\<in>X. {i}) = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
-    using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def)
-  also have "(\<Union>i\<in>X. {i}) = X" by auto
-  finally show ?thesis .
-qed
-
-lemma measure_eqI_countable:
-  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A"
-  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
-  shows "M = N"
-proof (rule measure_eqI)
-  fix X assume "X \<in> sets M"
-  then have X: "X \<subseteq> A" by auto
-  moreover from A X have "countable X" by (auto dest: countable_subset)
-  ultimately have
-    "emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
-    "emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
-    by (auto intro!: emeasure_countable_singleton)
-  moreover have "(\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X) = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
-    using X by (intro nn_integral_cong eq) auto
-  ultimately show "emeasure M X = emeasure N X"
-    by simp
-qed simp
-
-lemma measure_eqI_countable_AE:
-  assumes [simp]: "sets M = UNIV" "sets N = UNIV"
-  assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>" and [simp]: "countable \<Omega>"
-  assumes eq: "\<And>x. x \<in> \<Omega> \<Longrightarrow> emeasure M {x} = emeasure N {x}"
-  shows "M = N"
-proof (rule measure_eqI)
-  fix A
-  have "emeasure N A = emeasure N {x\<in>\<Omega>. x \<in> A}"
-    using ae by (intro emeasure_eq_AE) auto
-  also have "\<dots> = (\<integral>\<^sup>+x. emeasure N {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
-    by (intro emeasure_countable_singleton) auto
-  also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
-    by (intro nn_integral_cong eq[symmetric]) auto
-  also have "\<dots> = emeasure M {x\<in>\<Omega>. x \<in> A}"
-    by (intro emeasure_countable_singleton[symmetric]) auto
-  also have "\<dots> = emeasure M A"
-    using ae by (intro emeasure_eq_AE) auto
-  finally show "emeasure M A = emeasure N A" ..
-qed simp
-
-lemma nn_integral_monotone_convergence_SUP_nat:
-  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ennreal"
-  assumes chain: "Complete_Partial_Order.chain op \<le> (f ` Y)"
-  and nonempty: "Y \<noteq> {}"
-  shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space UNIV) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))"
-  (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _")
-proof (rule order_class.order.antisym)
-  show "?rhs \<le> ?lhs"
-    by (auto intro!: SUP_least SUP_upper nn_integral_mono)
-next
-  have "\<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i:Y. f i x) = (SUP i. g i)" for x
-    by (rule ennreal_Sup_countable_SUP) (simp add: nonempty)
-  then obtain g where incseq: "\<And>x. incseq (g x)"
-    and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y"
-    and sup: "\<And>x. (SUP i:Y. f i x) = (SUP i. g x i)" by moura
-  from incseq have incseq': "incseq (\<lambda>i x. g x i)"
-    by(blast intro: incseq_SucI le_funI dest: incseq_SucD)
-
-  have "?lhs = \<integral>\<^sup>+ x. (SUP i. g x i) \<partial>?M" by(simp add: sup)
-  also have "\<dots> = (SUP i. \<integral>\<^sup>+ x. g x i \<partial>?M)" using incseq'
-    by(rule nn_integral_monotone_convergence_SUP) simp
-  also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
-  proof(rule SUP_least)
-    fix n
-    have "\<And>x. \<exists>i. g x n = f i x \<and> i \<in> Y" using range by blast
-    then obtain I where I: "\<And>x. g x n = f (I x) x" "\<And>x. I x \<in> Y" by moura
-
-    have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)"
-      by(rule nn_integral_count_space_nat)
-    also have "\<dots> = (SUP m. \<Sum>x<m. g x n)"
-      by(rule suminf_eq_SUP)
-    also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)"
-    proof(rule SUP_mono)
-      fix m
-      show "\<exists>m'\<in>Y. (\<Sum>x<m. g x n) \<le> (\<integral>\<^sup>+ x. f m' x \<partial>?M)"
-      proof(cases "m > 0")
-        case False
-        thus ?thesis using nonempty by auto
-      next
-        case True
-        let ?Y = "I ` {..<m}"
-        have "f ` ?Y \<subseteq> f ` Y" using I by auto
-        with chain have chain': "Complete_Partial_Order.chain op \<le> (f ` ?Y)" by(rule chain_subset)
-        hence "Sup (f ` ?Y) \<in> f ` ?Y"
-          by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff)
-        then obtain m' where "m' < m" and m': "(SUP i:?Y. f i) = f (I m')" by auto
-        have "I m' \<in> Y" using I by blast
-        have "(\<Sum>x<m. g x n) \<le> (\<Sum>x<m. f (I m') x)"
-        proof(rule setsum_mono)
-          fix x
-          assume "x \<in> {..<m}"
-          hence "x < m" by simp
-          have "g x n = f (I x) x" by(simp add: I)
-          also have "\<dots> \<le> (SUP i:?Y. f i) x" unfolding Sup_fun_def image_image
-            using \<open>x \<in> {..<m}\<close> by (rule Sup_upper [OF imageI])
-          also have "\<dots> = f (I m') x" unfolding m' by simp
-          finally show "g x n \<le> f (I m') x" .
-        qed
-        also have "\<dots> \<le> (SUP m. (\<Sum>x<m. f (I m') x))"
-          by(rule SUP_upper) simp
-        also have "\<dots> = (\<Sum>x. f (I m') x)"
-          by(rule suminf_eq_SUP[symmetric])
-        also have "\<dots> = (\<integral>\<^sup>+ x. f (I m') x \<partial>?M)"
-          by(rule nn_integral_count_space_nat[symmetric])
-        finally show ?thesis using \<open>I m' \<in> Y\<close> by blast
-      qed
-    qed
-    finally show "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) \<le> \<dots>" .
-  qed
-  finally show "?lhs \<le> ?rhs" .
-qed
-
-lemma power_series_tendsto_at_left:
-  assumes nonneg: "\<And>i. 0 \<le> f i" and summable: "\<And>z. 0 \<le> z \<Longrightarrow> z < 1 \<Longrightarrow> summable (\<lambda>n. f n * z^n)"
-  shows "((\<lambda>z. ennreal (\<Sum>n. f n * z^n)) \<longlongrightarrow> (\<Sum>n. ennreal (f n))) (at_left (1::real))"
-proof (intro tendsto_at_left_sequentially)
-  show "0 < (1::real)" by simp
-  fix S :: "nat \<Rightarrow> real" assume S: "\<And>n. S n < 1" "\<And>n. 0 < S n" "S \<longlonglongrightarrow> 1" "incseq S"
-  then have S_nonneg: "\<And>i. 0 \<le> S i" by (auto intro: less_imp_le)
-
-  have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) \<longlonglongrightarrow> (\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV)"
-  proof (rule nn_integral_LIMSEQ)
-    show "incseq (\<lambda>i n. ennreal (f n * S i^n))"
-      using S by (auto intro!: mult_mono power_mono nonneg ennreal_leI
-                       simp: incseq_def le_fun_def less_imp_le)
-    fix n have "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n * 1^n)"
-      by (intro tendsto_intros tendsto_ennrealI S)
-    then show "(\<lambda>i. ennreal (f n * S i^n)) \<longlonglongrightarrow> ennreal (f n)"
-      by simp
-  qed (auto simp: S_nonneg intro!: mult_nonneg_nonneg nonneg)
-  also have "(\<lambda>i. (\<integral>\<^sup>+n. f n * S i^n \<partial>count_space UNIV)) = (\<lambda>i. \<Sum>n. f n * S i^n)"
-    by (subst nn_integral_count_space_nat)
-       (intro ext suminf_ennreal2 mult_nonneg_nonneg nonneg S_nonneg
-              zero_le_power summable S)+
-  also have "(\<integral>\<^sup>+n. ennreal (f n) \<partial>count_space UNIV) = (\<Sum>n. ennreal (f n))"
-    by (simp add: nn_integral_count_space_nat nonneg)
-  finally show "(\<lambda>n. ennreal (\<Sum>na. f na * S n ^ na)) \<longlonglongrightarrow> (\<Sum>n. ennreal (f n))" .
-qed
-
-subsubsection \<open>Measures with Restricted Space\<close>
-
-lemma simple_function_restrict_space_ennreal:
-  fixes f :: "'a \<Rightarrow> ennreal"
-  assumes "\<Omega> \<inter> space M \<in> sets M"
-  shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)"
-proof -
-  { assume "finite (f ` space (restrict_space M \<Omega>))"
-    then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
-    then have "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
-      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
-  moreover
-  { assume "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
-    then have "finite (f ` space (restrict_space M \<Omega>))"
-      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
-  ultimately show ?thesis
-    unfolding
-      simple_function_iff_borel_measurable borel_measurable_restrict_space_iff_ennreal[OF assms]
-    by auto
-qed
-
-lemma simple_function_restrict_space:
-  fixes f :: "'a \<Rightarrow> 'b::real_normed_vector"
-  assumes "\<Omega> \<inter> space M \<in> sets M"
-  shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)"
-proof -
-  { assume "finite (f ` space (restrict_space M \<Omega>))"
-    then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
-    then have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
-      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
-  moreover
-  { assume "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
-    then have "finite (f ` space (restrict_space M \<Omega>))"
-      by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) }
-  ultimately show ?thesis
-    unfolding simple_function_iff_borel_measurable
-      borel_measurable_restrict_space_iff[OF assms]
-    by auto
-qed
-
-lemma simple_integral_restrict_space:
-  assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f"
-  shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)"
-  using simple_function_restrict_space_ennreal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)]
-  by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def
-           split: split_indicator split_indicator_asm
-           intro!: setsum.mono_neutral_cong_left ennreal_mult_left_cong arg_cong2[where f=emeasure])
-
-lemma nn_integral_restrict_space:
-  assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M"
-  shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)"
-proof -
-  let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> f \<and> (\<forall>x. s x < top)}"
-  have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)"
-  proof (safe intro!: image_eqI)
-    fix s assume s: "simple_function ?R s" "s \<le> f" "\<forall>x. s x < top"
-    from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)"
-      by (intro simple_integral_restrict_space) auto
-    from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
-      by (simp add: simple_function_restrict_space_ennreal)
-    from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> (\<lambda>x. f x * indicator \<Omega> x)"
-      "\<And>x. s x * indicator \<Omega> x < top"
-      by (auto split: split_indicator simp: le_fun_def image_subset_iff)
-  next
-    fix s assume s: "simple_function M s" "s \<le> (\<lambda>x. f x * indicator \<Omega> x)" "\<forall>x. s x < top"
-    then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s')
-      by (intro simple_function_mult simple_function_indicator) auto
-    also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)"
-      by (rule simple_function_cong) (auto split: split_indicator)
-    finally show sf: "simple_function (restrict_space M \<Omega>) s"
-      by (simp add: simple_function_restrict_space_ennreal)
-
-    from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)"
-      by (auto simp add: fun_eq_iff le_fun_def image_subset_iff
-                  split: split_indicator split_indicator_asm
-                  intro: antisym)
-
-    show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s"
-      by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf])
-    show "\<And>x. s x < top"
-      using s by (auto simp: image_subset_iff)
-    from s show "s \<le> f"
-      by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm)
-  qed
-  then show ?thesis
-    unfolding nn_integral_def_finite by (simp cong del: strong_SUP_cong)
-qed
-
-lemma nn_integral_count_space_indicator:
-  assumes "NO_MATCH (UNIV::'a set) (X::'a set)"
-  shows "(\<integral>\<^sup>+x. f x \<partial>count_space X) = (\<integral>\<^sup>+x. f x * indicator X x \<partial>count_space UNIV)"
-  by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space)
-
-lemma nn_integral_count_space_eq:
-  "(\<And>x. x \<in> A - B \<Longrightarrow> f x = 0) \<Longrightarrow> (\<And>x. x \<in> B - A \<Longrightarrow> f x = 0) \<Longrightarrow>
-    (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)"
-  by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)
-
-lemma nn_integral_ge_point:
-  assumes "x \<in> A"
-  shows "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
-proof -
-  from assms have "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space {x}"
-    by(auto simp add: nn_integral_count_space_finite max_def)
-  also have "\<dots> = \<integral>\<^sup>+ x'. p x' * indicator {x} x' \<partial>count_space A"
-    using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong)
-  also have "\<dots> \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A"
-    by(rule nn_integral_mono)(simp add: indicator_def)
-  finally show ?thesis .
-qed
-
-subsubsection \<open>Measure spaces with an associated density\<close>
-
-definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
-  "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
-
-lemma
-  shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M"
-    and space_density[simp]: "space (density M f) = space M"
-  by (auto simp: density_def)
-
-(* FIXME: add conversion to simplify space, sets and measurable *)
-lemma space_density_imp[measurable_dest]:
-  "\<And>x M f. x \<in> space (density M f) \<Longrightarrow> x \<in> space M" by auto
-
-lemma
-  shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'"
-    and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'"
-    and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u"
-  unfolding measurable_def simple_function_def by simp_all
-
-lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow>
-  (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'"
-  unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed)
-
-lemma emeasure_density:
-  assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M"
-  shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)"
-    (is "_ = ?\<mu> A")
-  unfolding density_def
-proof (rule emeasure_measure_of_sigma)
-  show "sigma_algebra (space M) (sets M)" ..
-  show "positive (sets M) ?\<mu>"
-    using f by (auto simp: positive_def)
-  show "countably_additive (sets M) ?\<mu>"
-  proof (intro countably_additiveI)
-    fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M"
-    then have "\<And>i. A i \<in> sets M" by auto
-    then have *: "\<And>i. (\<lambda>x. f x * indicator (A i) x) \<in> borel_measurable M"
-      by auto
-    assume disj: "disjoint_family A"
-    then have "(\<Sum>n. ?\<mu> (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M)"
-       using f * by (subst nn_integral_suminf) auto
-    also have "(\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (A n) x) \<partial>M) = (\<integral>\<^sup>+ x. f x * (\<Sum>n. indicator (A n) x) \<partial>M)"
-      using f by (auto intro!: ennreal_suminf_cmult nn_integral_cong_AE)
-    also have "\<dots> = (\<integral>\<^sup>+ x. f x * indicator (\<Union>n. A n) x \<partial>M)"
-      unfolding suminf_indicator[OF disj] ..
-    finally show "(\<Sum>i. \<integral>\<^sup>+ x. f x * indicator (A i) x \<partial>M) = \<integral>\<^sup>+ x. f x * indicator (\<Union>i. A i) x \<partial>M" .
-  qed
-qed fact
-
-lemma null_sets_density_iff:
-  assumes f: "f \<in> borel_measurable M"
-  shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x = 0)"
-proof -
-  { assume "A \<in> sets M"
-    have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> emeasure M {x \<in> space M. f x * indicator A x \<noteq> 0} = 0"
-      using f \<open>A \<in> sets M\<close> by (intro nn_integral_0_iff) auto
-    also have "\<dots> \<longleftrightarrow> (AE x in M. f x * indicator A x = 0)"
-      using f \<open>A \<in> sets M\<close> by (intro AE_iff_measurable[OF _ refl, symmetric]) auto
-    also have "(AE x in M. f x * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
-      by (auto simp add: indicator_def max_def split: if_split_asm)
-    finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
-  with f show ?thesis
-    by (simp add: null_sets_def emeasure_density cong: conj_cong)
-qed
-
-lemma AE_density:
-  assumes f: "f \<in> borel_measurable M"
-  shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)"
-proof
-  assume "AE x in density M f. P x"
-  with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x = 0"
-    by (auto simp: eventually_ae_filter null_sets_density_iff)
-  then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto
-  with ae show "AE x in M. 0 < f x \<longrightarrow> P x"
-    by (rule eventually_elim2) auto
-next
-  fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x"
-  then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
-    by (auto simp: eventually_ae_filter)
-  then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. f x = 0}"
-    "N \<union> {x\<in>space M. f x = 0} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
-    using f by (auto simp: subset_eq zero_less_iff_neq_zero intro!: AE_not_in)
-  show "AE x in density M f. P x"
-    using ae2
-    unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f]
-    by (intro exI[of _ "N \<union> {x\<in>space M. f x = 0}"] conjI *) (auto elim: eventually_elim2)
-qed
-
-lemma nn_integral_density:
-  assumes f: "f \<in> borel_measurable M"
-  assumes g: "g \<in> borel_measurable M"
-  shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)"
-using g proof induct
-  case (cong u v)
-  then show ?case
-    apply (subst nn_integral_cong[OF cong(3)])
-    apply (simp_all cong: nn_integral_cong)
-    done
-next
-  case (set A) then show ?case
-    by (simp add: emeasure_density f)
-next
-  case (mult u c)
-  moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps)
-  ultimately show ?case
-    using f by (simp add: nn_integral_cmult)
-next
-  case (add u v)
-  then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x"
-    by (simp add: distrib_left)
-  with add f show ?case
-    by (auto simp add: nn_integral_add intro!: nn_integral_add[symmetric])
-next
-  case (seq U)
-  have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)"
-    by eventually_elim (simp add: SUP_mult_left_ennreal seq)
-  from seq f show ?case
-    apply (simp add: nn_integral_monotone_convergence_SUP)
-    apply (subst nn_integral_cong_AE[OF eq])
-    apply (subst nn_integral_monotone_convergence_SUP_AE)
-    apply (auto simp: incseq_def le_fun_def intro!: mult_left_mono)
-    done
-qed
-
-lemma density_distr:
-  assumes [measurable]: "f \<in> borel_measurable N" "X \<in> measurable M N"
-  shows "density (distr M N X) f = distr (density M (\<lambda>x. f (X x))) N X"
-  by (intro measure_eqI)
-     (auto simp add: emeasure_density nn_integral_distr emeasure_distr
-           split: split_indicator intro!: nn_integral_cong)
-
-lemma emeasure_restricted:
-  assumes S: "S \<in> sets M" and X: "X \<in> sets M"
-  shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)"
-proof -
-  have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)"
-    using S X by (simp add: emeasure_density)
-  also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)"
-    by (auto intro!: nn_integral_cong simp: indicator_def)
-  also have "\<dots> = emeasure M (S \<inter> X)"
-    using S X by (simp add: sets.Int)
-  finally show ?thesis .
-qed
-
-lemma measure_restricted:
-  "S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)"
-  by (simp add: emeasure_restricted measure_def)
-
-lemma (in finite_measure) finite_measure_restricted:
-  "S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))"
-  by standard (simp add: emeasure_restricted)
-
-lemma emeasure_density_const:
-  "A \<in> sets M \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A"
-  by (auto simp: nn_integral_cmult_indicator emeasure_density)
-
-lemma measure_density_const:
-  "A \<in> sets M \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = enn2real c * measure M A"
-  by (auto simp: emeasure_density_const measure_def enn2real_mult)
-
-lemma density_density_eq:
-   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
-   density (density M f) g = density M (\<lambda>x. f x * g x)"
-  by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps)
-
-lemma distr_density_distr:
-  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
-    and inv: "\<forall>x\<in>space M. T' (T x) = x"
-  assumes f: "f \<in> borel_measurable M'"
-  shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L")
-proof (rule measure_eqI)
-  fix A assume A: "A \<in> sets ?R"
-  { fix x assume "x \<in> space M"
-    with sets.sets_into_space[OF A]
-    have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ennreal)"
-      using T inv by (auto simp: indicator_def measurable_space) }
-  with A T T' f show "emeasure ?R A = emeasure ?L A"
-    by (simp add: measurable_comp emeasure_density emeasure_distr
-                  nn_integral_distr measurable_sets cong: nn_integral_cong)
-qed simp
-
-lemma density_density_divide:
-  fixes f g :: "'a \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
-  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
-  assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
-  shows "density (density M f) (\<lambda>x. g x / f x) = density M g"
-proof -
-  have "density M g = density M (\<lambda>x. ennreal (f x) * ennreal (g x / f x))"
-    using f g ac by (auto intro!: density_cong measurable_If simp: ennreal_mult[symmetric])
-  then show ?thesis
-    using f g by (subst density_density_eq) auto
-qed
-
-lemma density_1: "density M (\<lambda>_. 1) = M"
-  by (intro measure_eqI) (auto simp: emeasure_density)
-
-lemma emeasure_density_add:
-  assumes X: "X \<in> sets M"
-  assumes Mf[measurable]: "f \<in> borel_measurable M"
-  assumes Mg[measurable]: "g \<in> borel_measurable M"
-  shows "emeasure (density M f) X + emeasure (density M g) X =
-           emeasure (density M (\<lambda>x. f x + g x)) X"
-  using assms
-  apply (subst (1 2 3) emeasure_density, simp_all) []
-  apply (subst nn_integral_add[symmetric], simp_all) []
-  apply (intro nn_integral_cong, simp split: split_indicator)
-  done
-
-subsubsection \<open>Point measure\<close>
-
-definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
-  "point_measure A f = density (count_space A) f"
-
-lemma
-  shows space_point_measure: "space (point_measure A f) = A"
-    and sets_point_measure: "sets (point_measure A f) = Pow A"
-  by (auto simp: point_measure_def)
-
-lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)"
-  by (simp add: sets_point_measure)
-
-lemma measurable_point_measure_eq1[simp]:
-  "g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M"
-  unfolding point_measure_def by simp
-
-lemma measurable_point_measure_eq2_finite[simp]:
-  "finite A \<Longrightarrow>
-   g \<in> measurable M (point_measure A f) \<longleftrightarrow>
-    (g \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. g -` {a} \<inter> space M \<in> sets M))"
-  unfolding point_measure_def by (simp add: measurable_count_space_eq2)
-
-lemma simple_function_point_measure[simp]:
-  "simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)"
-  by (simp add: point_measure_def)
-
-lemma emeasure_point_measure:
-  assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
-  shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)"
-proof -
-  have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
-    using \<open>X \<subseteq> A\<close> by auto
-  with A show ?thesis
-    by (simp add: emeasure_density nn_integral_count_space point_measure_def indicator_def)
-qed
-
-lemma emeasure_point_measure_finite:
-  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
-  by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
-
-lemma emeasure_point_measure_finite2:
-  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)"
-  by (subst emeasure_point_measure)
-     (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le)
-
-lemma null_sets_point_measure_iff:
-  "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x = 0)"
- by (auto simp: AE_count_space null_sets_density_iff point_measure_def)
-
-lemma AE_point_measure:
-  "(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)"
-  unfolding point_measure_def
-  by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def)
-
-lemma nn_integral_point_measure:
-  "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
-    integral\<^sup>N (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
-  unfolding point_measure_def
-  by (subst nn_integral_density)
-     (simp_all add: nn_integral_density nn_integral_count_space ennreal_zero_less_mult_iff)
-
-lemma nn_integral_point_measure_finite:
-  "finite A \<Longrightarrow> integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
-  by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le)
-
-subsubsection \<open>Uniform measure\<close>
-
-definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
-
-lemma
-  shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M"
-    and space_uniform_measure[simp]: "space (uniform_measure M A) = space M"
-  by (auto simp: uniform_measure_def)
-
-lemma emeasure_uniform_measure[simp]:
-  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
-  shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A"
-proof -
-  from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
-    by (auto simp add: uniform_measure_def emeasure_density divide_ennreal_def split: split_indicator
-             intro!: nn_integral_cong)
-  also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A"
-    using A B
-    by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int divide_ennreal_def mult.commute)
-  finally show ?thesis .
-qed
-
-lemma measure_uniform_measure[simp]:
-  assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M"
-  shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A"
-  using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A
-  by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ennreal2_cases)
-     (simp_all add: measure_def divide_ennreal top_ennreal.rep_eq top_ereal_def ennreal_top_divide)
-
-lemma AE_uniform_measureI:
-  "A \<in> sets M \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x) \<Longrightarrow> (AE x in uniform_measure M A. P x)"
-  unfolding uniform_measure_def by (auto simp: AE_density divide_ennreal_def)
-
-lemma emeasure_uniform_measure_1:
-  "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> emeasure (uniform_measure M S) S = 1"
-  by (subst emeasure_uniform_measure)
-     (simp_all add: emeasure_neq_0_sets emeasure_eq_ennreal_measure divide_ennreal
-                    zero_less_iff_neq_zero[symmetric])
-
-lemma nn_integral_uniform_measure:
-  assumes f[measurable]: "f \<in> borel_measurable M" and S[measurable]: "S \<in> sets M"
-  shows "(\<integral>\<^sup>+x. f x \<partial>uniform_measure M S) = (\<integral>\<^sup>+x. f x * indicator S x \<partial>M) / emeasure M S"
-proof -
-  { assume "emeasure M S = \<infinity>"
-    then have ?thesis
-      by (simp add: uniform_measure_def nn_integral_density f) }
-  moreover
-  { assume [simp]: "emeasure M S = 0"
-    then have ae: "AE x in M. x \<notin> S"
-      using sets.sets_into_space[OF S]
-      by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong)
-    from ae have "(\<integral>\<^sup>+ x. indicator S x / 0 * f x \<partial>M) = 0"
-      by (subst nn_integral_0_iff_AE) auto
-    moreover from ae have "(\<integral>\<^sup>+ x. f x * indicator S x \<partial>M) = 0"
-      by (subst nn_integral_0_iff_AE) auto
-    ultimately have ?thesis
-      by (simp add: uniform_measure_def nn_integral_density f) }
-  moreover have "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> ?thesis"
-    unfolding uniform_measure_def
-    by (subst nn_integral_density)
-       (auto simp: ennreal_times_divide f nn_integral_divide[symmetric] mult.commute)
-  ultimately show ?thesis by blast
-qed
-
-lemma AE_uniform_measure:
-  assumes "emeasure M A \<noteq> 0" "emeasure M A < \<infinity>"
-  shows "(AE x in uniform_measure M A. P x) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x)"
-proof -
-  have "A \<in> sets M"
-    using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets)
-  moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> A"
-    using assms
-    by (cases "emeasure M A") (auto split: split_indicator simp: ennreal_zero_less_divide)
-  ultimately show ?thesis
-    unfolding uniform_measure_def by (simp add: AE_density)
-qed
-
-subsubsection \<open>Null measure\<close>
-
-lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 0)"
-  by (intro measure_eqI) (simp_all add: emeasure_density)
-
-lemma nn_integral_null_measure[simp]: "(\<integral>\<^sup>+x. f x \<partial>null_measure M) = 0"
-  by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ennreal_def le_fun_def
-           intro!: exI[of _ "\<lambda>x. 0"])
-
-lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M"
-proof (intro measure_eqI)
-  fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A"
-    by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure)
-qed simp
-
-subsubsection \<open>Uniform count measure\<close>
-
-definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
-
-lemma
-  shows space_uniform_count_measure: "space (uniform_count_measure A) = A"
-    and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A"
-    unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure)
-
-lemma sets_uniform_count_measure_count_space[measurable_cong]:
-  "sets (uniform_count_measure A) = sets (count_space A)"
-  by (simp add: sets_uniform_count_measure)
-
-lemma emeasure_uniform_count_measure:
-  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A"
-  by (simp add: emeasure_point_measure_finite uniform_count_measure_def divide_inverse ennreal_mult
-                ennreal_of_nat_eq_real_of_nat)
-
-lemma measure_uniform_count_measure:
-  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A"
-  by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def enn2real_mult)
-
-lemma space_uniform_count_measure_empty_iff [simp]:
-  "space (uniform_count_measure X) = {} \<longleftrightarrow> X = {}"
-by(simp add: space_uniform_count_measure)
-
-lemma sets_uniform_count_measure_eq_UNIV [simp]:
-  "sets (uniform_count_measure UNIV) = UNIV \<longleftrightarrow> True"
-  "UNIV = sets (uniform_count_measure UNIV) \<longleftrightarrow> True"
-by(simp_all add: sets_uniform_count_measure)
-
-subsubsection \<open>Scaled measure\<close>
-
-lemma nn_integral_scale_measure:
-  assumes f: "f \<in> borel_measurable M"
-  shows "nn_integral (scale_measure r M) f = r * nn_integral M f"
-  using f
-proof induction
-  case (cong f g)
-  thus ?case
-    by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp)
-next
-  case (mult f c)
-  thus ?case
-    by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute)
-next
-  case (add f g)
-  thus ?case
-    by(simp add: nn_integral_add distrib_left)
-next
-  case (seq U)
-  thus ?case
-    by(simp add: nn_integral_monotone_convergence_SUP SUP_mult_left_ennreal)
-qed simp
-
-end
--- a/src/HOL/Probability/Probability.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Probability.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -5,14 +5,10 @@
 theory Probability
 imports
   Discrete_Topology
-  Complete_Measure
-  Projective_Limit
-  Probability_Mass_Function
   SPMF
   PMF_Impl
   Stream_Space
   Random_Permutations
-  Embed_Measure
   Central_Limit_Theorem
 begin
 
--- a/src/HOL/Probability/Probability_Measure.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -6,51 +6,9 @@
 section \<open>Probability measure\<close>
 
 theory Probability_Measure
-  imports Lebesgue_Measure Radon_Nikodym
+  imports "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
 begin
 
-lemma (in finite_measure) countable_support:
-  "countable {x. measure M {x} \<noteq> 0}"
-proof cases
-  assume "measure M (space M) = 0"
-  with bounded_measure measure_le_0_iff have "{x. measure M {x} \<noteq> 0} = {}"
-    by auto
-  then show ?thesis
-    by simp
-next
-  let ?M = "measure M (space M)" and ?m = "\<lambda>x. measure M {x}"
-  assume "?M \<noteq> 0"
-  then have *: "{x. ?m x \<noteq> 0} = (\<Union>n. {x. ?M / Suc n < ?m x})"
-    using reals_Archimedean[of "?m x / ?M" for x]
-    by (auto simp: field_simps not_le[symmetric] measure_nonneg divide_le_0_iff measure_le_0_iff)
-  have **: "\<And>n. finite {x. ?M / Suc n < ?m x}"
-  proof (rule ccontr)
-    fix n assume "infinite {x. ?M / Suc n < ?m x}" (is "infinite ?X")
-    then obtain X where "finite X" "card X = Suc (Suc n)" "X \<subseteq> ?X"
-      by (metis infinite_arbitrarily_large)
-    from this(3) have *: "\<And>x. x \<in> X \<Longrightarrow> ?M / Suc n \<le> ?m x"
-      by auto
-    { fix x assume "x \<in> X"
-      from \<open>?M \<noteq> 0\<close> *[OF this] have "?m x \<noteq> 0" by (auto simp: field_simps measure_le_0_iff)
-      then have "{x} \<in> sets M" by (auto dest: measure_notin_sets) }
-    note singleton_sets = this
-    have "?M < (\<Sum>x\<in>X. ?M / Suc n)"
-      using \<open>?M \<noteq> 0\<close>
-      by (simp add: \<open>card X = Suc (Suc n)\<close> of_nat_Suc field_simps less_le measure_nonneg)
-    also have "\<dots> \<le> (\<Sum>x\<in>X. ?m x)"
-      by (rule setsum_mono) fact
-    also have "\<dots> = measure M (\<Union>x\<in>X. {x})"
-      using singleton_sets \<open>finite X\<close>
-      by (intro finite_measure_finite_Union[symmetric]) (auto simp: disjoint_family_on_def)
-    finally have "?M < measure M (\<Union>x\<in>X. {x})" .
-    moreover have "measure M (\<Union>x\<in>X. {x}) \<le> ?M"
-      using singleton_sets[THEN sets.sets_into_space] by (intro finite_measure_mono) auto
-    ultimately show False by simp
-  qed
-  show ?thesis
-    unfolding * by (intro countable_UN countableI_type countable_finite[OF **])
-qed
-
 locale prob_space = finite_measure +
   assumes emeasure_space_1: "emeasure M (space M) = 1"
 
@@ -129,7 +87,7 @@
   using emeasure_space[of M X] by (simp add: emeasure_space_1)
 
 lemma (in prob_space) measure_ge_1_iff: "measure M A \<ge> 1 \<longleftrightarrow> measure M A = 1"
-  by (auto intro!: antisym)  
+  by (auto intro!: antisym)
 
 lemma (in prob_space) AE_I_eq_1:
   assumes "emeasure M {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
@@ -1267,4 +1225,7 @@
   by (subst emeasure_distr)
      (auto simp: measurable_cong_sets[OF Q] prod_emb_def space_PiM[symmetric] intro!: measurable_restrict)
 
+lemma (in prob_space) prob_space_completion: "prob_space (completion M)"
+  by (rule prob_spaceI) (simp add: emeasure_space_1)
+
 end
--- a/src/HOL/Probability/Projective_Family.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Projective_Family.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -6,7 +6,7 @@
 section \<open>Projective Family\<close>
 
 theory Projective_Family
-imports Finite_Product_Measure Giry_Monad
+imports Giry_Monad
 begin
 
 lemma vimage_restrict_preseve_mono:
--- a/src/HOL/Probability/Projective_Limit.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -5,13 +5,10 @@
 section \<open>Projective Limit\<close>
 
 theory Projective_Limit
-  imports
-    Caratheodory
-    Fin_Map
-    Regularity
-    Projective_Family
-    Infinite_Product_Measure
-    "~~/src/HOL/Library/Diagonal_Subsequence"
+imports
+  Fin_Map
+  Infinite_Product_Measure
+  "~~/src/HOL/Library/Diagonal_Subsequence"
 begin
 
 subsection \<open>Sequences of Finite Maps in Compact Sets\<close>
--- a/src/HOL/Probability/Radon_Nikodym.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1064 +0,0 @@
-(*  Title:      HOL/Probability/Radon_Nikodym.thy
-    Author:     Johannes Hölzl, TU München
-*)
-
-section \<open>Radon-Nikod{\'y}m derivative\<close>
-
-theory Radon_Nikodym
-imports Bochner_Integration
-begin
-
-definition diff_measure :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a measure"
-where
-  "diff_measure M N = measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)"
-
-lemma
-  shows space_diff_measure[simp]: "space (diff_measure M N) = space M"
-    and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M"
-  by (auto simp: diff_measure_def)
-
-lemma emeasure_diff_measure:
-  assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N"
-  assumes pos: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure N A \<le> emeasure M A" and A: "A \<in> sets M"
-  shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\<mu> A")
-  unfolding diff_measure_def
-proof (rule emeasure_measure_of_sigma)
-  show "sigma_algebra (space M) (sets M)" ..
-  show "positive (sets M) ?\<mu>"
-    using pos by (simp add: positive_def)
-  show "countably_additive (sets M) ?\<mu>"
-  proof (rule countably_additiveI)
-    fix A :: "nat \<Rightarrow> _"  assume A: "range A \<subseteq> sets M" and "disjoint_family A"
-    then have suminf:
-      "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
-      "(\<Sum>i. emeasure N (A i)) = emeasure N (\<Union>i. A i)"
-      by (simp_all add: suminf_emeasure sets_eq)
-    with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
-      (\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))"
-      using fin pos[of "A _"]
-      by (intro ennreal_suminf_minus)
-         (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure)
-    then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
-      emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) "
-      by (simp add: suminf)
-  qed
-qed fact
-
-lemma (in sigma_finite_measure) Ex_finite_integrable_function:
-  "\<exists>h\<in>borel_measurable M. integral\<^sup>N M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>)"
-proof -
-  obtain A :: "nat \<Rightarrow> 'a set" where
-    range[measurable]: "range A \<subseteq> sets M" and
-    space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and
-    disjoint: "disjoint_family A"
-    using sigma_finite_disjoint by blast
-  let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)"
-  have [measurable]: "\<And>i. A i \<in> sets M"
-    using range by fastforce+
-  have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
-  proof
-    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
-      using measure[of i]
-      by (auto intro!: dense simp: ennreal_inverse_positive ennreal_mult_eq_top_iff power_eq_top_ennreal)
-  qed
-  from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
-    "\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto
-  { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
-  let ?h = "\<lambda>x. \<Sum>i. n i * indicator (A i) x"
-  show ?thesis
-  proof (safe intro!: bexI[of _ ?h] del: notI)
-    have "integral\<^sup>N M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
-      by (simp add: nn_integral_suminf nn_integral_cmult_indicator)
-    also have "\<dots> \<le> (\<Sum>i. ennreal ((1/2)^Suc i))"
-    proof (intro suminf_le allI)
-      fix N
-      have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)"
-        using n[of N] by (intro mult_right_mono) auto
-      also have "\<dots> = (1/2)^Suc N * (inverse (emeasure M (A N)) * emeasure M (A N))"
-        using measure[of N]
-        by (simp add: ennreal_inverse_power divide_ennreal_def ennreal_inverse_mult
-                      power_eq_top_ennreal less_top[symmetric] mult_ac
-                 del: power_Suc)
-      also have "\<dots> \<le> inverse (ennreal 2) ^ Suc N"
-        using measure[of N]
-        by (cases "emeasure M (A N)"; cases "emeasure M (A N) = 0")
-           (auto simp: inverse_ennreal ennreal_mult[symmetric] divide_ennreal_def simp del: power_Suc)
-      also have "\<dots> = ennreal (inverse 2 ^ Suc N)"
-        by (subst ennreal_power[symmetric], simp) (simp add: inverse_ennreal)
-      finally show "n N * emeasure M (A N) \<le> ennreal ((1/2)^Suc N)"
-        by simp
-    qed auto
-    also have "\<dots> < top"
-      unfolding less_top[symmetric]
-      by (rule ennreal_suminf_neq_top)
-         (auto simp: summable_geometric summable_Suc_iff simp del: power_Suc)
-    finally show "integral\<^sup>N M ?h \<noteq> \<infinity>"
-      by (auto simp: top_unique)
-  next
-    { fix x assume "x \<in> space M"
-      then obtain i where "x \<in> A i" using space[symmetric] by auto
-      with disjoint n have "?h x = n i"
-        by (auto intro!: suminf_cmult_indicator intro: less_imp_le)
-      then show "0 < ?h x" and "?h x < \<infinity>" using n[of i] by (auto simp: less_top[symmetric]) }
-    note pos = this
-  qed measurable
-qed
-
-subsection "Absolutely continuous"
-
-definition absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
-  "absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N"
-
-lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M"
-  unfolding absolutely_continuous_def by (auto simp: null_sets_count_space)
-
-lemma absolutely_continuousI_density:
-  "f \<in> borel_measurable M \<Longrightarrow> absolutely_continuous M (density M f)"
-  by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in)
-
-lemma absolutely_continuousI_point_measure_finite:
-  "(\<And>x. \<lbrakk> x \<in> A ; f x \<le> 0 \<rbrakk> \<Longrightarrow> g x \<le> 0) \<Longrightarrow> absolutely_continuous (point_measure A f) (point_measure A g)"
-  unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff)
-
-lemma absolutely_continuousD:
-  "absolutely_continuous M N \<Longrightarrow> A \<in> sets M \<Longrightarrow> emeasure M A = 0 \<Longrightarrow> emeasure N A = 0"
-  by (auto simp: absolutely_continuous_def null_sets_def)
-
-lemma absolutely_continuous_AE:
-  assumes sets_eq: "sets M' = sets M"
-    and "absolutely_continuous M M'" "AE x in M. P x"
-   shows "AE x in M'. P x"
-proof -
-  from \<open>AE x in M. P x\<close> obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
-    unfolding eventually_ae_filter by auto
-  show "AE x in M'. P x"
-  proof (rule AE_I')
-    show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp
-    from \<open>absolutely_continuous M M'\<close> show "N \<in> null_sets M'"
-      using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto
-  qed
-qed
-
-subsection "Existence of the Radon-Nikodym derivative"
-
-lemma (in finite_measure) Radon_Nikodym_finite_measure:
-  assumes "finite_measure N" and sets_eq[simp]: "sets N = sets M"
-  assumes "absolutely_continuous M N"
-  shows "\<exists>f \<in> borel_measurable M. density M f = N"
-proof -
-  interpret N: finite_measure N by fact
-  define G where "G = {g \<in> borel_measurable M. \<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A}"
-  have [measurable_dest]: "f \<in> G \<Longrightarrow> f \<in> borel_measurable M"
-    and G_D: "\<And>A. f \<in> G \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) \<le> N A" for f
-    by (auto simp: G_def)
-  note this[measurable_dest]
-  have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
-  hence "G \<noteq> {}" by auto
-  { fix f g assume f[measurable]: "f \<in> G" and g[measurable]: "g \<in> G"
-    have "(\<lambda>x. max (g x) (f x)) \<in> G" (is "?max \<in> G") unfolding G_def
-    proof safe
-      let ?A = "{x \<in> space M. f x \<le> g x}"
-      have "?A \<in> sets M" using f g unfolding G_def by auto
-      fix A assume [measurable]: "A \<in> sets M"
-      have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
-        using sets.sets_into_space[OF \<open>A \<in> sets M\<close>] by auto
-      have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
-        g x * indicator (?A \<inter> A) x + f x * indicator ((space M - ?A) \<inter> A) x"
-        by (auto simp: indicator_def max_def)
-      hence "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) =
-        (\<integral>\<^sup>+x. g x * indicator (?A \<inter> A) x \<partial>M) +
-        (\<integral>\<^sup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
-        by (auto cong: nn_integral_cong intro!: nn_integral_add)
-      also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)"
-        using f g unfolding G_def by (auto intro!: add_mono)
-      also have "\<dots> = N A"
-        using union by (subst plus_emeasure) auto
-      finally show "(\<integral>\<^sup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .
-    qed auto }
-  note max_in_G = this
-  { fix f assume  "incseq f" and f: "\<And>i. f i \<in> G"
-    then have [measurable]: "\<And>i. f i \<in> borel_measurable M" by (auto simp: G_def)
-    have "(\<lambda>x. SUP i. f i x) \<in> G" unfolding G_def
-    proof safe
-      show "(\<lambda>x. SUP i. f i x) \<in> borel_measurable M" by measurable
-    next
-      fix A assume "A \<in> sets M"
-      have "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) =
-        (\<integral>\<^sup>+x. (SUP i. f i x * indicator A x) \<partial>M)"
-        by (intro nn_integral_cong) (simp split: split_indicator)
-      also have "\<dots> = (SUP i. (\<integral>\<^sup>+x. f i x * indicator A x \<partial>M))"
-        using \<open>incseq f\<close> f \<open>A \<in> sets M\<close>
-        by (intro nn_integral_monotone_convergence_SUP)
-           (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
-      finally show "(\<integral>\<^sup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
-        using f \<open>A \<in> sets M\<close> by (auto intro!: SUP_least simp: G_D)
-    qed }
-  note SUP_in_G = this
-  let ?y = "SUP g : G. integral\<^sup>N M g"
-  have y_le: "?y \<le> N (space M)" unfolding G_def
-  proof (safe intro!: SUP_least)
-    fix g assume "\<forall>A\<in>sets M. (\<integral>\<^sup>+x. g x * indicator A x \<partial>M) \<le> N A"
-    from this[THEN bspec, OF sets.top] show "integral\<^sup>N M g \<le> N (space M)"
-      by (simp cong: nn_integral_cong)
-  qed
-  from ennreal_SUP_countable_SUP [OF \<open>G \<noteq> {}\<close>, of "integral\<^sup>N M"] guess ys .. note ys = this
-  then have "\<forall>n. \<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n"
-  proof safe
-    fix n assume "range ys \<subseteq> integral\<^sup>N M ` G"
-    hence "ys n \<in> integral\<^sup>N M ` G" by auto
-    thus "\<exists>g. g\<in>G \<and> integral\<^sup>N M g = ys n" by auto
-  qed
-  from choice[OF this] obtain gs where "\<And>i. gs i \<in> G" "\<And>n. integral\<^sup>N M (gs n) = ys n" by auto
-  hence y_eq: "?y = (SUP i. integral\<^sup>N M (gs i))" using ys by auto
-  let ?g = "\<lambda>i x. Max ((\<lambda>n. gs n x) ` {..i})"
-  define f where [abs_def]: "f x = (SUP i. ?g i x)" for x
-  let ?F = "\<lambda>A x. f x * indicator A x"
-  have gs_not_empty: "\<And>i x. (\<lambda>n. gs n x) ` {..i} \<noteq> {}" by auto
-  { fix i have "?g i \<in> G"
-    proof (induct i)
-      case 0 thus ?case by simp fact
-    next
-      case (Suc i)
-      with Suc gs_not_empty \<open>gs (Suc i) \<in> G\<close> show ?case
-        by (auto simp add: atMost_Suc intro!: max_in_G)
-    qed }
-  note g_in_G = this
-  have "incseq ?g" using gs_not_empty
-    by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc)
-
-  from SUP_in_G[OF this g_in_G] have [measurable]: "f \<in> G" unfolding f_def .
-  then have [measurable]: "f \<in> borel_measurable M" unfolding G_def by auto
-
-  have "integral\<^sup>N M f = (SUP i. integral\<^sup>N M (?g i))" unfolding f_def
-    using g_in_G \<open>incseq ?g\<close> by (auto intro!: nn_integral_monotone_convergence_SUP simp: G_def)
-  also have "\<dots> = ?y"
-  proof (rule antisym)
-    show "(SUP i. integral\<^sup>N M (?g i)) \<le> ?y"
-      using g_in_G by (auto intro: SUP_mono)
-    show "?y \<le> (SUP i. integral\<^sup>N M (?g i))" unfolding y_eq
-      by (auto intro!: SUP_mono nn_integral_mono Max_ge)
-  qed
-  finally have int_f_eq_y: "integral\<^sup>N M f = ?y" .
-
-  have upper_bound: "\<forall>A\<in>sets M. N A \<le> density M f A"
-  proof (rule ccontr)
-    assume "\<not> ?thesis"
-    then obtain A where A[measurable]: "A \<in> sets M" and f_less_N: "density M f A < N A"
-      by (auto simp: not_le)
-    then have pos_A: "0 < M A"
-      using \<open>absolutely_continuous M N\<close>[THEN absolutely_continuousD, OF A]
-      by (auto simp: zero_less_iff_neq_zero)
-
-    define b where "b = (N A - density M f A) / M A / 2"
-    with f_less_N pos_A have "0 < b" "b \<noteq> top"
-      by (auto intro!: diff_gr0_ennreal simp: zero_less_iff_neq_zero diff_eq_0_iff_ennreal ennreal_divide_eq_top_iff)
-
-    let ?f = "\<lambda>x. f x + b"
-    have "nn_integral M f \<noteq> top"
-      using `f \<in> G`[THEN G_D, of "space M"] by (auto simp: top_unique cong: nn_integral_cong)
-    with \<open>b \<noteq> top\<close> interpret Mf: finite_measure "density M ?f"
-      by (intro finite_measureI)
-         (auto simp: field_simps mult_indicator_subset ennreal_mult_eq_top_iff
-                     emeasure_density nn_integral_cmult_indicator nn_integral_add
-               cong: nn_integral_cong)
-
-    from unsigned_Hahn_decomposition[of "density M ?f" N A]
-    obtain Y where [measurable]: "Y \<in> sets M" and [simp]: "Y \<subseteq> A"
-       and Y1: "\<And>C. C \<in> sets M \<Longrightarrow> C \<subseteq> Y \<Longrightarrow> density M ?f C \<le> N C"
-       and Y2: "\<And>C. C \<in> sets M \<Longrightarrow> C \<subseteq> A \<Longrightarrow> C \<inter> Y = {} \<Longrightarrow> N C \<le> density M ?f C"
-       by auto
-
-    let ?f' = "\<lambda>x. f x + b * indicator Y x"
-    have "M Y \<noteq> 0"
-    proof
-      assume "M Y = 0"
-      then have "N Y = 0"
-        using \<open>absolutely_continuous M N\<close>[THEN absolutely_continuousD, of Y] by auto
-      then have "N A = N (A - Y)"
-        by (subst emeasure_Diff) auto
-      also have "\<dots> \<le> density M ?f (A - Y)"
-        by (rule Y2) auto
-      also have "\<dots> \<le> density M ?f A - density M ?f Y"
-        by (subst emeasure_Diff) auto
-      also have "\<dots> \<le> density M ?f A - 0"
-        by (intro ennreal_minus_mono) auto
-      also have "density M ?f A = b * M A + density M f A"
-        by (simp add: emeasure_density field_simps mult_indicator_subset nn_integral_add nn_integral_cmult_indicator)
-      also have "\<dots> < N A"
-        using f_less_N pos_A
-        by (cases "density M f A"; cases "M A"; cases "N A")
-           (auto simp: b_def ennreal_less_iff ennreal_minus divide_ennreal ennreal_numeral[symmetric]
-                       ennreal_plus[symmetric] ennreal_mult[symmetric] field_simps
-                    simp del: ennreal_numeral ennreal_plus)
-      finally show False
-        by simp
-    qed
-    then have "nn_integral M f < nn_integral M ?f'"
-      using \<open>0 < b\<close> \<open>nn_integral M f \<noteq> top\<close>
-      by (simp add: nn_integral_add nn_integral_cmult_indicator ennreal_zero_less_mult_iff zero_less_iff_neq_zero)
-    moreover
-    have "?f' \<in> G"
-      unfolding G_def
-    proof safe
-      fix X assume [measurable]: "X \<in> sets M"
-      have "(\<integral>\<^sup>+ x. ?f' x * indicator X x \<partial>M) = density M f (X - Y) + density M ?f (X \<inter> Y)"
-        by (auto simp add: emeasure_density nn_integral_add[symmetric] split: split_indicator intro!: nn_integral_cong)
-      also have "\<dots> \<le> N (X - Y) + N (X \<inter> Y)"
-        using G_D[OF \<open>f \<in> G\<close>] by (intro add_mono Y1) (auto simp: emeasure_density)
-      also have "\<dots> = N X"
-        by (subst plus_emeasure) (auto intro!: arg_cong2[where f=emeasure])
-      finally show "(\<integral>\<^sup>+ x. ?f' x * indicator X x \<partial>M) \<le> N X" .
-    qed simp
-    then have "nn_integral M ?f' \<le> ?y"
-      by (rule SUP_upper)
-    ultimately show False
-      by (simp add: int_f_eq_y)
-  qed
-  show ?thesis
-  proof (intro bexI[of _ f] measure_eqI conjI antisym)
-    fix A assume "A \<in> sets (density M f)" then show "emeasure (density M f) A \<le> emeasure N A"
-      by (auto simp: emeasure_density intro!: G_D[OF \<open>f \<in> G\<close>])
-  next
-    fix A assume A: "A \<in> sets (density M f)" then show "emeasure N A \<le> emeasure (density M f) A"
-      using upper_bound by auto
-  qed auto
-qed
-
-lemma (in finite_measure) split_space_into_finite_sets_and_rest:
-  assumes ac: "absolutely_continuous M N" and sets_eq[simp]: "sets N = sets M"
-  shows "\<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> (\<forall>i. N (B i) \<noteq> \<infinity>) \<and>
-    (\<forall>A\<in>sets M. A \<inter> (\<Union>i. B i) = {} \<longrightarrow> (emeasure M A = 0 \<and> N A = 0) \<or> (emeasure M A > 0 \<and> N A = \<infinity>))"
-proof -
-  let ?Q = "{Q\<in>sets M. N Q \<noteq> \<infinity>}"
-  let ?a = "SUP Q:?Q. emeasure M Q"
-  have "{} \<in> ?Q" by auto
-  then have Q_not_empty: "?Q \<noteq> {}" by blast
-  have "?a \<le> emeasure M (space M)" using sets.sets_into_space
-    by (auto intro!: SUP_least emeasure_mono)
-  then have "?a \<noteq> \<infinity>"
-    using finite_emeasure_space
-    by (auto simp: less_top[symmetric] top_unique simp del: SUP_eq_top_iff Sup_eq_top_iff)
-  from ennreal_SUP_countable_SUP [OF Q_not_empty, of "emeasure M"]
-  obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
-    by auto
-  then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto
-  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = emeasure M (Q' i)" "\<And>i. Q' i \<in> ?Q"
-    by auto
-  then have a_Lim: "?a = (SUP i::nat. emeasure M (Q' i))" using a by simp
-  let ?O = "\<lambda>n. \<Union>i\<le>n. Q' i"
-  have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\<Union>i. ?O i)"
-  proof (rule SUP_emeasure_incseq[of ?O])
-    show "range ?O \<subseteq> sets M" using Q' by auto
-    show "incseq ?O" by (fastforce intro!: incseq_SucI)
-  qed
-  have Q'_sets[measurable]: "\<And>i. Q' i \<in> sets M" using Q' by auto
-  have O_sets: "\<And>i. ?O i \<in> sets M" using Q' by auto
-  then have O_in_G: "\<And>i. ?O i \<in> ?Q"
-  proof (safe del: notI)
-    fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto
-    then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))"
-      by (simp add: emeasure_subadditive_finite)
-    also have "\<dots> < \<infinity>" using Q' by (simp add: less_top)
-    finally show "N (?O i) \<noteq> \<infinity>" by simp
-  qed auto
-  have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce
-  have a_eq: "?a = emeasure M (\<Union>i. ?O i)" unfolding Union[symmetric]
-  proof (rule antisym)
-    show "?a \<le> (SUP i. emeasure M (?O i))" unfolding a_Lim
-      using Q' by (auto intro!: SUP_mono emeasure_mono)
-    show "(SUP i. emeasure M (?O i)) \<le> ?a"
-    proof (safe intro!: Sup_mono, unfold bex_simps)
-      fix i
-      have *: "(\<Union>(Q' ` {..i})) = ?O i" by auto
-      then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and>
-        emeasure M (\<Union>(Q' ` {..i})) \<le> emeasure M x"
-        using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
-    qed
-  qed
-  let ?O_0 = "(\<Union>i. ?O i)"
-  have "?O_0 \<in> sets M" using Q' by auto
-  have "disjointed Q' i \<in> sets M" for i
-    using sets.range_disjointed_sets[of Q' M] using Q'_sets by (auto simp: subset_eq)
-  note Q_sets = this
-  show ?thesis
-  proof (intro bexI exI conjI ballI impI allI)
-    show "disjoint_family (disjointed Q')"
-      by (rule disjoint_family_disjointed)
-    show "range (disjointed Q') \<subseteq> sets M"
-      using Q'_sets by (intro sets.range_disjointed_sets) auto
-    { fix A assume A: "A \<in> sets M" "A \<inter> (\<Union>i. disjointed Q' i) = {}"
-      then have A1: "A \<inter> (\<Union>i. Q' i) = {}"
-        unfolding UN_disjointed_eq by auto
-      show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
-      proof (rule disjCI, simp)
-        assume *: "emeasure M A = 0 \<or> N A \<noteq> top"
-        show "emeasure M A = 0 \<and> N A = 0"
-        proof (cases "emeasure M A = 0")
-          case True
-          with ac A have "N A = 0"
-            unfolding absolutely_continuous_def by auto
-          with True show ?thesis by simp
-        next
-          case False
-          with * have "N A \<noteq> \<infinity>" by auto
-          with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)"
-            using Q' A1 by (auto intro!: plus_emeasure simp: set_eq_iff)
-          also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))"
-          proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
-            show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
-              using \<open>N A \<noteq> \<infinity>\<close> O_sets A by auto
-          qed (fastforce intro!: incseq_SucI)
-          also have "\<dots> \<le> ?a"
-          proof (safe intro!: SUP_least)
-            fix i have "?O i \<union> A \<in> ?Q"
-            proof (safe del: notI)
-              show "?O i \<union> A \<in> sets M" using O_sets A by auto
-              from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A"
-                using emeasure_subadditive[of "?O i" N A] A O_sets by auto
-              with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>"
-                using \<open>N A \<noteq> \<infinity>\<close> by (auto simp: top_unique)
-            qed
-            then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper)
-          qed
-          finally have "emeasure M A = 0"
-            unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure)
-          with \<open>emeasure M A \<noteq> 0\<close> show ?thesis by auto
-        qed
-      qed }
-    { fix i
-      have "N (disjointed Q' i) \<le> N (Q' i)"
-        by (auto intro!: emeasure_mono simp: disjointed_def)
-      then show "N (disjointed Q' i) \<noteq> \<infinity>"
-        using Q'(2)[of i] by (auto simp: top_unique) }
-  qed
-qed
-
-lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
-  assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M"
-  shows "\<exists>f\<in>borel_measurable M. density M f = N"
-proof -
-  from split_space_into_finite_sets_and_rest[OF assms]
-  obtain Q :: "nat \<Rightarrow> 'a set"
-    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
-    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
-    and Q_fin: "\<And>i. N (Q i) \<noteq> \<infinity>" by force
-  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
-  let ?N = "\<lambda>i. density N (indicator (Q i))" and ?M = "\<lambda>i. density M (indicator (Q i))"
-  have "\<forall>i. \<exists>f\<in>borel_measurable (?M i). density (?M i) f = ?N i"
-  proof (intro allI finite_measure.Radon_Nikodym_finite_measure)
-    fix i
-    from Q show "finite_measure (?M i)"
-      by (auto intro!: finite_measureI cong: nn_integral_cong
-               simp add: emeasure_density subset_eq sets_eq)
-    from Q have "emeasure (?N i) (space N) = emeasure N (Q i)"
-      by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: nn_integral_cong)
-    with Q_fin show "finite_measure (?N i)"
-      by (auto intro!: finite_measureI)
-    show "sets (?N i) = sets (?M i)" by (simp add: sets_eq)
-    have [measurable]: "\<And>A. A \<in> sets M \<Longrightarrow> A \<in> sets N" by (simp add: sets_eq)
-    show "absolutely_continuous (?M i) (?N i)"
-      using \<open>absolutely_continuous M N\<close> \<open>Q i \<in> sets M\<close>
-      by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq
-               intro!: absolutely_continuous_AE[OF sets_eq])
-  qed
-  from choice[OF this[unfolded Bex_def]]
-  obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
-    and f_density: "\<And>i. density (?M i) (f i) = ?N i"
-    by force
-  { fix A i assume A: "A \<in> sets M"
-    with Q borel have "(\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"
-      by (auto simp add: emeasure_density nn_integral_density subset_eq
-               intro!: nn_integral_cong split: split_indicator)
-    also have "\<dots> = emeasure N (Q i \<inter> A)"
-      using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq)
-    finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^sup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }
-  note integral_eq = this
-  let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator (space M - (\<Union>i. Q i)) x"
-  show ?thesis
-  proof (safe intro!: bexI[of _ ?f])
-    show "?f \<in> borel_measurable M" using borel Q_sets
-      by (auto intro!: measurable_If)
-    show "density M ?f = N"
-    proof (rule measure_eqI)
-      fix A assume "A \<in> sets (density M ?f)"
-      then have "A \<in> sets M" by simp
-      have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
-      have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
-        "\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
-        using borel Qi \<open>A \<in> sets M\<close> by auto
-      have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator ((space M - (\<Union>i. Q i)) \<inter> A) x \<partial>M)"
-        using borel by (intro nn_integral_cong) (auto simp: indicator_def)
-      also have "\<dots> = (\<integral>\<^sup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)"
-        using borel Qi \<open>A \<in> sets M\<close>
-        by (subst nn_integral_add)
-           (auto simp add: nn_integral_cmult_indicator sets.Int intro!: suminf_0_le)
-      also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)"
-        by (subst integral_eq[OF \<open>A \<in> sets M\<close>], subst nn_integral_suminf) auto
-      finally have "(\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A)" .
-      moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"
-        using Q Q_sets \<open>A \<in> sets M\<close>
-        by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)
-      moreover
-      have "(space M - (\<Union>x. Q x)) \<inter> A \<inter> (\<Union>x. Q x) = {}"
-        by auto
-      then have "\<infinity> * emeasure M ((space M - (\<Union>i. Q i)) \<inter> A) = N ((space M - (\<Union>i. Q i)) \<inter> A)"
-        using in_Q0[of "(space M - (\<Union>i. Q i)) \<inter> A"] \<open>A \<in> sets M\<close> Q by (auto simp: ennreal_top_mult)
-      moreover have "(space M - (\<Union>i. Q i)) \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
-        using Q_sets \<open>A \<in> sets M\<close> by auto
-      moreover have "((\<Union>i. Q i) \<inter> A) \<union> ((space M - (\<Union>i. Q i)) \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> ((space M - (\<Union>i. Q i)) \<inter> A) = {}"
-        using \<open>A \<in> sets M\<close> sets.sets_into_space by auto
-      ultimately have "N A = (\<integral>\<^sup>+x. ?f x * indicator A x \<partial>M)"
-        using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "(space M - (\<Union>i. Q i)) \<inter> A"] by (simp add: sets_eq)
-      with \<open>A \<in> sets M\<close> borel Q show "emeasure (density M ?f) A = N A"
-        by (auto simp: subset_eq emeasure_density)
-    qed (simp add: sets_eq)
-  qed
-qed
-
-lemma (in sigma_finite_measure) Radon_Nikodym:
-  assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M"
-  shows "\<exists>f \<in> borel_measurable M. density M f = N"
-proof -
-  from Ex_finite_integrable_function
-  obtain h where finite: "integral\<^sup>N M h \<noteq> \<infinity>" and
-    borel: "h \<in> borel_measurable M" and
-    nn: "\<And>x. 0 \<le> h x" and
-    pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
-    "\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto
-  let ?T = "\<lambda>A. (\<integral>\<^sup>+x. h x * indicator A x \<partial>M)"
-  let ?MT = "density M h"
-  from borel finite nn interpret T: finite_measure ?MT
-    by (auto intro!: finite_measureI cong: nn_integral_cong simp: emeasure_density)
-  have "absolutely_continuous ?MT N" "sets N = sets ?MT"
-  proof (unfold absolutely_continuous_def, safe)
-    fix A assume "A \<in> null_sets ?MT"
-    with borel have "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> h x \<le> 0"
-      by (auto simp add: null_sets_density_iff)
-    with pos sets.sets_into_space have "AE x in M. x \<notin> A"
-      by (elim eventually_mono) (auto simp: not_le[symmetric])
-    then have "A \<in> null_sets M"
-      using \<open>A \<in> sets M\<close> by (simp add: AE_iff_null_sets)
-    with ac show "A \<in> null_sets N"
-      by (auto simp: absolutely_continuous_def)
-  qed (auto simp add: sets_eq)
-  from T.Radon_Nikodym_finite_measure_infinite[OF this]
-  obtain f where f_borel: "f \<in> borel_measurable M" "density ?MT f = N" by auto
-  with nn borel show ?thesis
-    by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)
-qed
-
-subsection \<open>Uniqueness of densities\<close>
-
-lemma finite_density_unique:
-  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
-  and fin: "integral\<^sup>N M f \<noteq> \<infinity>"
-  shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"
-proof (intro iffI ballI)
-  fix A assume eq: "AE x in M. f x = g x"
-  with borel show "density M f = density M g"
-    by (auto intro: density_cong)
-next
-  let ?P = "\<lambda>f A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M"
-  assume "density M f = density M g"
-  with borel have eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
-    by (simp add: emeasure_density[symmetric])
-  from this[THEN bspec, OF sets.top] fin
-  have g_fin: "integral\<^sup>N M g \<noteq> \<infinity>" by (simp cong: nn_integral_cong)
-  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-      and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
-      and g_fin: "integral\<^sup>N M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
-    let ?N = "{x\<in>space M. g x < f x}"
-    have N: "?N \<in> sets M" using borel by simp
-    have "?P g ?N \<le> integral\<^sup>N M g" using pos
-      by (intro nn_integral_mono_AE) (auto split: split_indicator)
-    then have Pg_fin: "?P g ?N \<noteq> \<infinity>" using g_fin by (auto simp: top_unique)
-    have "?P (\<lambda>x. (f x - g x)) ?N = (\<integral>\<^sup>+x. f x * indicator ?N x - g x * indicator ?N x \<partial>M)"
-      by (auto intro!: nn_integral_cong simp: indicator_def)
-    also have "\<dots> = ?P f ?N - ?P g ?N"
-    proof (rule nn_integral_diff)
-      show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
-        using borel N by auto
-      show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x"
-        using pos by (auto split: split_indicator)
-    qed fact
-    also have "\<dots> = 0"
-      unfolding eq[THEN bspec, OF N] using Pg_fin by auto
-    finally have "AE x in M. f x \<le> g x"
-      using pos borel nn_integral_PInf_AE[OF borel(2) g_fin]
-      by (subst (asm) nn_integral_0_iff_AE)
-         (auto split: split_indicator simp: not_less ennreal_minus_eq_0) }
-  from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
-  show "AE x in M. f x = g x" by auto
-qed
-
-lemma (in finite_measure) density_unique_finite_measure:
-  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
-  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x"
-  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. f' x * indicator A x \<partial>M)"
-    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
-  shows "AE x in M. f x = f' x"
-proof -
-  let ?D = "\<lambda>f. density M f"
-  let ?N = "\<lambda>A. ?P f A" and ?N' = "\<lambda>A. ?P f' A"
-  let ?f = "\<lambda>A x. f x * indicator A x" and ?f' = "\<lambda>A x. f' x * indicator A x"
-
-  have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M"
-    using borel by (auto intro!: absolutely_continuousI_density)
-  from split_space_into_finite_sets_and_rest[OF this]
-  obtain Q :: "nat \<Rightarrow> 'a set"
-    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
-    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> ?D f A = 0 \<or> 0 < emeasure M A \<and> ?D f A = \<infinity>"
-    and Q_fin: "\<And>i. ?D f (Q i) \<noteq> \<infinity>" by force
-  with borel pos have in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<inter> (\<Union>i. Q i) = {} \<Longrightarrow> emeasure M A = 0 \<and> ?N A = 0 \<or> 0 < emeasure M A \<and> ?N A = \<infinity>"
-    and Q_fin: "\<And>i. ?N (Q i) \<noteq> \<infinity>" by (auto simp: emeasure_density subset_eq)
-
-  from Q have Q_sets[measurable]: "\<And>i. Q i \<in> sets M" by auto
-  let ?D = "{x\<in>space M. f x \<noteq> f' x}"
-  have "?D \<in> sets M" using borel by auto
-  have *: "\<And>i x A. \<And>y::ennreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
-    unfolding indicator_def by auto
-  have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
-    by (intro finite_density_unique[THEN iffD1] allI)
-       (auto intro!: f measure_eqI simp: emeasure_density * subset_eq)
-  moreover have "AE x in M. ?f (space M - (\<Union>i. Q i)) x = ?f' (space M - (\<Union>i. Q i)) x"
-  proof (rule AE_I')
-    { fix f :: "'a \<Rightarrow> ennreal" assume borel: "f \<in> borel_measurable M"
-        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M)"
-      let ?A = "\<lambda>i. (space M - (\<Union>i. Q i)) \<inter> {x \<in> space M. f x < (i::nat)}"
-      have "(\<Union>i. ?A i) \<in> null_sets M"
-      proof (rule null_sets_UN)
-        fix i ::nat have "?A i \<in> sets M"
-          using borel by auto
-        have "?N (?A i) \<le> (\<integral>\<^sup>+x. (i::ennreal) * indicator (?A i) x \<partial>M)"
-          unfolding eq[OF \<open>?A i \<in> sets M\<close>]
-          by (auto intro!: nn_integral_mono simp: indicator_def)
-        also have "\<dots> = i * emeasure M (?A i)"
-          using \<open>?A i \<in> sets M\<close> by (auto intro!: nn_integral_cmult_indicator)
-        also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by (auto simp: ennreal_mult_less_top of_nat_less_top)
-        finally have "?N (?A i) \<noteq> \<infinity>" by simp
-        then show "?A i \<in> null_sets M" using in_Q0[OF \<open>?A i \<in> sets M\<close>] \<open>?A i \<in> sets M\<close> by auto
-      qed
-      also have "(\<Union>i. ?A i) = (space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}"
-        by (auto simp: ennreal_Ex_less_of_nat less_top[symmetric])
-      finally have "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp }
-    from this[OF borel(1) refl] this[OF borel(2) f]
-    have "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "(space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all
-    then show "((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets M" by (rule null_sets.Un)
-    show "{x \<in> space M. ?f (space M - (\<Union>i. Q i)) x \<noteq> ?f' (space M - (\<Union>i. Q i)) x} \<subseteq>
-      ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> ((space M - (\<Union>i. Q i)) \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def)
-  qed
-  moreover have "AE x in M. (?f (space M - (\<Union>i. Q i)) x = ?f' (space M - (\<Union>i. Q i)) x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
-    ?f (space M) x = ?f' (space M) x"
-    by (auto simp: indicator_def)
-  ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x"
-    unfolding AE_all_countable[symmetric]
-    by eventually_elim (auto split: if_split_asm simp: indicator_def)
-  then show "AE x in M. f x = f' x" by auto
-qed
-
-lemma (in sigma_finite_measure) density_unique:
-  assumes f: "f \<in> borel_measurable M"
-  assumes f': "f' \<in> borel_measurable M"
-  assumes density_eq: "density M f = density M f'"
-  shows "AE x in M. f x = f' x"
-proof -
-  obtain h where h_borel: "h \<in> borel_measurable M"
-    and fin: "integral\<^sup>N M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
-    using Ex_finite_integrable_function by auto
-  then have h_nn: "AE x in M. 0 \<le> h x" by auto
-  let ?H = "density M h"
-  interpret h: finite_measure ?H
-    using fin h_borel pos
-    by (intro finite_measureI) (simp cong: nn_integral_cong emeasure_density add: fin)
-  let ?fM = "density M f"
-  let ?f'M = "density M f'"
-  { fix A assume "A \<in> sets M"
-    then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
-      using pos(1) sets.sets_into_space by (force simp: indicator_def)
-    then have "(\<integral>\<^sup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
-      using h_borel \<open>A \<in> sets M\<close> h_nn by (subst nn_integral_0_iff) auto }
-  note h_null_sets = this
-  { fix A assume "A \<in> sets M"
-    have "(\<integral>\<^sup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?fM)"
-      using \<open>A \<in> sets M\<close> h_borel h_nn f f'
-      by (intro nn_integral_density[symmetric]) auto
-    also have "\<dots> = (\<integral>\<^sup>+x. h x * indicator A x \<partial>?f'M)"
-      by (simp_all add: density_eq)
-    also have "\<dots> = (\<integral>\<^sup>+x. f' x * (h x * indicator A x) \<partial>M)"
-      using \<open>A \<in> sets M\<close> h_borel h_nn f f'
-      by (intro nn_integral_density) auto
-    finally have "(\<integral>\<^sup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^sup>+x. h x * (f' x * indicator A x) \<partial>M)"
-      by (simp add: ac_simps)
-    then have "(\<integral>\<^sup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^sup>+x. (f' x * indicator A x) \<partial>?H)"
-      using \<open>A \<in> sets M\<close> h_borel h_nn f f'
-      by (subst (asm) (1 2) nn_integral_density[symmetric]) auto }
-  then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
-    by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M]) auto
-  with AE_space[of M] pos show "AE x in M. f x = f' x"
-    unfolding AE_density[OF h_borel] by auto
-qed
-
-lemma (in sigma_finite_measure) density_unique_iff:
-  assumes f: "f \<in> borel_measurable M" and f': "f' \<in> borel_measurable M"
-  shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)"
-  using density_unique[OF assms] density_cong[OF f f'] by auto
-
-lemma sigma_finite_density_unique:
-  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  and fin: "sigma_finite_measure (density M f)"
-  shows "density M f = density M g \<longleftrightarrow> (AE x in M. f x = g x)"
-proof
-  assume "AE x in M. f x = g x" with borel show "density M f = density M g"
-    by (auto intro: density_cong)
-next
-  assume eq: "density M f = density M g"
-  interpret f: sigma_finite_measure "density M f" by fact
-  from f.sigma_finite_incseq guess A . note cover = this
-
-  have "AE x in M. \<forall>i. x \<in> A i \<longrightarrow> f x = g x"
-    unfolding AE_all_countable
-  proof
-    fix i
-    have "density (density M f) (indicator (A i)) = density (density M g) (indicator (A i))"
-      unfolding eq ..
-    moreover have "(\<integral>\<^sup>+x. f x * indicator (A i) x \<partial>M) \<noteq> \<infinity>"
-      using cover(1) cover(3)[of i] borel by (auto simp: emeasure_density subset_eq)
-    ultimately have "AE x in M. f x * indicator (A i) x = g x * indicator (A i) x"
-      using borel cover(1)
-      by (intro finite_density_unique[THEN iffD1]) (auto simp: density_density_eq subset_eq)
-    then show "AE x in M. x \<in> A i \<longrightarrow> f x = g x"
-      by auto
-  qed
-  with AE_space show "AE x in M. f x = g x"
-    apply eventually_elim
-    using cover(2)[symmetric]
-    apply auto
-    done
-qed
-
-lemma (in sigma_finite_measure) sigma_finite_iff_density_finite':
-  assumes f: "f \<in> borel_measurable M"
-  shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
-    (is "sigma_finite_measure ?N \<longleftrightarrow> _")
-proof
-  assume "sigma_finite_measure ?N"
-  then interpret N: sigma_finite_measure ?N .
-  from N.Ex_finite_integrable_function obtain h where
-    h: "h \<in> borel_measurable M" "integral\<^sup>N ?N h \<noteq> \<infinity>" and
-    fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>"
-    by auto
-  have "AE x in M. f x * h x \<noteq> \<infinity>"
-  proof (rule AE_I')
-    have "integral\<^sup>N ?N h = (\<integral>\<^sup>+x. f x * h x \<partial>M)"
-      using f h by (auto intro!: nn_integral_density)
-    then have "(\<integral>\<^sup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
-      using h(2) by simp
-    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
-      using f h(1) by (auto intro!: nn_integral_PInf[unfolded infinity_ennreal_def] borel_measurable_vimage)
-  qed auto
-  then show "AE x in M. f x \<noteq> \<infinity>"
-    using fin by (auto elim!: AE_Ball_mp simp: less_top ennreal_mult_less_top)
-next
-  assume AE: "AE x in M. f x \<noteq> \<infinity>"
-  from sigma_finite guess Q . note Q = this
-  define A where "A i =
-    f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ennreal(of_nat (Suc n))}) \<inter> space M" for i
-  { fix i j have "A i \<inter> Q j \<in> sets M"
-    unfolding A_def using f Q
-    apply (rule_tac sets.Int)
-    by (cases i) (auto intro: measurable_sets[OF f(1)]) }
-  note A_in_sets = this
-
-  show "sigma_finite_measure ?N"
-  proof (standard, intro exI conjI ballI)
-    show "countable (range (\<lambda>(i, j). A i \<inter> Q j))"
-      by auto
-    show "range (\<lambda>(i, j). A i \<inter> Q j) \<subseteq> sets (density M f)"
-      using A_in_sets by auto
-  next
-    have "\<Union>range (\<lambda>(i, j). A i \<inter> Q j) = (\<Union>i j. A i \<inter> Q j)"
-      by auto
-    also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto
-    also have "(\<Union>i. A i) = space M"
-    proof safe
-      fix x assume x: "x \<in> space M"
-      show "x \<in> (\<Union>i. A i)"
-      proof (cases "f x" rule: ennreal_cases)
-        case top with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0])
-      next
-        case (real r)
-        with ennreal_Ex_less_of_nat[of "f x"] obtain n :: nat where "f x < n"
-          by auto
-        also have "n < (Suc n :: ennreal)"
-          by simp
-        finally show ?thesis
-          using x real by (auto simp: A_def ennreal_of_nat_eq_real_of_nat intro!: exI[of _ "Suc n"])
-      qed
-    qed (auto simp: A_def)
-    finally show "\<Union>range (\<lambda>(i, j). A i \<inter> Q j) = space ?N" by simp
-  next
-    fix X assume "X \<in> range (\<lambda>(i, j). A i \<inter> Q j)"
-    then obtain i j where [simp]:"X = A i \<inter> Q j" by auto
-    have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
-    proof (cases i)
-      case 0
-      have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0"
-        using AE by (auto simp: A_def \<open>i = 0\<close>)
-      from nn_integral_cong_AE[OF this] show ?thesis by simp
-    next
-      case (Suc n)
-      then have "(\<integral>\<^sup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
-        (\<integral>\<^sup>+x. (Suc n :: ennreal) * indicator (Q j) x \<partial>M)"
-        by (auto intro!: nn_integral_mono simp: indicator_def A_def ennreal_of_nat_eq_real_of_nat)
-      also have "\<dots> = Suc n * emeasure M (Q j)"
-        using Q by (auto intro!: nn_integral_cmult_indicator)
-      also have "\<dots> < \<infinity>"
-        using Q by (auto simp: ennreal_mult_less_top less_top of_nat_less_top)
-      finally show ?thesis by simp
-    qed
-    then show "emeasure ?N X \<noteq> \<infinity>"
-      using A_in_sets Q f by (auto simp: emeasure_density)
-  qed
-qed
-
-lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
-  "f \<in> borel_measurable M \<Longrightarrow> sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
-  by (subst sigma_finite_iff_density_finite')
-     (auto simp: max_def intro!: measurable_If)
-
-subsection \<open>Radon-Nikodym derivative\<close>
-
-definition RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ennreal" where
-  "RN_deriv M N =
-    (if \<exists>f. f \<in> borel_measurable M \<and> density M f = N
-       then SOME f. f \<in> borel_measurable M \<and> density M f = N
-       else (\<lambda>_. 0))"
-
-lemma RN_derivI:
-  assumes "f \<in> borel_measurable M" "density M f = N"
-  shows "density M (RN_deriv M N) = N"
-proof -
-  have *: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
-    using assms by auto
-  then have "density M (SOME f. f \<in> borel_measurable M \<and> density M f = N) = N"
-    by (rule someI2_ex) auto
-  with * show ?thesis
-    by (auto simp: RN_deriv_def)
-qed
-
-lemma borel_measurable_RN_deriv[measurable]: "RN_deriv M N \<in> borel_measurable M"
-proof -
-  { assume ex: "\<exists>f. f \<in> borel_measurable M \<and> density M f = N"
-    have 1: "(SOME f. f \<in> borel_measurable M \<and> density M f = N) \<in> borel_measurable M"
-      using ex by (rule someI2_ex) auto }
-  from this show ?thesis
-    by (auto simp: RN_deriv_def)
-qed
-
-lemma density_RN_deriv_density:
-  assumes f: "f \<in> borel_measurable M"
-  shows "density M (RN_deriv M (density M f)) = density M f"
-  by (rule RN_derivI[OF f]) simp
-
-lemma (in sigma_finite_measure) density_RN_deriv:
-  "absolutely_continuous M N \<Longrightarrow> sets N = sets M \<Longrightarrow> density M (RN_deriv M N) = N"
-  by (metis RN_derivI Radon_Nikodym)
-
-lemma (in sigma_finite_measure) RN_deriv_nn_integral:
-  assumes N: "absolutely_continuous M N" "sets N = sets M"
-    and f: "f \<in> borel_measurable M"
-  shows "integral\<^sup>N N f = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
-proof -
-  have "integral\<^sup>N N f = integral\<^sup>N (density M (RN_deriv M N)) f"
-    using N by (simp add: density_RN_deriv)
-  also have "\<dots> = (\<integral>\<^sup>+x. RN_deriv M N x * f x \<partial>M)"
-    using f by (simp add: nn_integral_density)
-  finally show ?thesis by simp
-qed
-
-lemma null_setsD_AE: "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
-  using AE_iff_null_sets[of N M] by auto
-
-lemma (in sigma_finite_measure) RN_deriv_unique:
-  assumes f: "f \<in> borel_measurable M"
-  and eq: "density M f = N"
-  shows "AE x in M. f x = RN_deriv M N x"
-  unfolding eq[symmetric]
-  by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv
-            density_RN_deriv_density[symmetric])
-
-lemma RN_deriv_unique_sigma_finite:
-  assumes f: "f \<in> borel_measurable M"
-  and eq: "density M f = N" and fin: "sigma_finite_measure N"
-  shows "AE x in M. f x = RN_deriv M N x"
-  using fin unfolding eq[symmetric]
-  by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv
-            density_RN_deriv_density[symmetric])
-
-lemma (in sigma_finite_measure) RN_deriv_distr:
-  fixes T :: "'a \<Rightarrow> 'b"
-  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
-    and inv: "\<forall>x\<in>space M. T' (T x) = x"
-  and ac[simp]: "absolutely_continuous (distr M M' T) (distr N M' T)"
-  and N: "sets N = sets M"
-  shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x"
-proof (rule RN_deriv_unique)
-  have [simp]: "sets N = sets M" by fact
-  note sets_eq_imp_space_eq[OF N, simp]
-  have measurable_N[simp]: "\<And>M'. measurable N M' = measurable M M'" by (auto simp: measurable_def)
-  { fix A assume "A \<in> sets M"
-    with inv T T' sets.sets_into_space[OF this]
-    have "T -` T' -` A \<inter> T -` space M' \<inter> space M = A"
-      by (auto simp: measurable_def) }
-  note eq = this[simp]
-  { fix A assume "A \<in> sets M"
-    with inv T T' sets.sets_into_space[OF this]
-    have "(T' \<circ> T) -` A \<inter> space M = A"
-      by (auto simp: measurable_def) }
-  note eq2 = this[simp]
-  let ?M' = "distr M M' T" and ?N' = "distr N M' T"
-  interpret M': sigma_finite_measure ?M'
-  proof
-    from sigma_finite_countable guess F .. note F = this
-    show "\<exists>A. countable A \<and> A \<subseteq> sets (distr M M' T) \<and> \<Union>A = space (distr M M' T) \<and> (\<forall>a\<in>A. emeasure (distr M M' T) a \<noteq> \<infinity>)"
-    proof (intro exI conjI ballI)
-      show *: "(\<lambda>A. T' -` A \<inter> space ?M') ` F \<subseteq> sets ?M'"
-        using F T' by (auto simp: measurable_def)
-      show "\<Union>((\<lambda>A. T' -` A \<inter> space ?M')`F) = space ?M'"
-        using F T'[THEN measurable_space] by (auto simp: set_eq_iff)
-    next
-      fix X assume "X \<in> (\<lambda>A. T' -` A \<inter> space ?M')`F"
-      then obtain A where [simp]: "X = T' -` A \<inter> space ?M'" and "A \<in> F" by auto
-      have "X \<in> sets M'" using F T' \<open>A\<in>F\<close> by auto
-      moreover
-      have Fi: "A \<in> sets M" using F \<open>A\<in>F\<close> by auto
-      ultimately show "emeasure ?M' X \<noteq> \<infinity>"
-        using F T T' \<open>A\<in>F\<close> by (simp add: emeasure_distr)
-    qed (insert F, auto)
-  qed
-  have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M"
-    using T ac by measurable
-  then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M"
-    by (simp add: comp_def)
-
-  have "N = distr N M (T' \<circ> T)"
-    by (subst measure_of_of_measure[of N, symmetric])
-       (auto simp add: distr_def sets.sigma_sets_eq intro!: measure_of_eq sets.space_closed)
-  also have "\<dots> = distr (distr N M' T) M T'"
-    using T T' by (simp add: distr_distr)
-  also have "\<dots> = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'"
-    using ac by (simp add: M'.density_RN_deriv)
-  also have "\<dots> = density M (RN_deriv (distr M M' T) (distr N M' T) \<circ> T)"
-    by (simp add: distr_density_distr[OF T T', OF inv])
-  finally show "density M (\<lambda>x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N"
-    by (simp add: comp_def)
-qed
-
-lemma (in sigma_finite_measure) RN_deriv_finite:
-  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
-  shows "AE x in M. RN_deriv M N x \<noteq> \<infinity>"
-proof -
-  interpret N: sigma_finite_measure N by fact
-  from N show ?thesis
-    using sigma_finite_iff_density_finite[OF borel_measurable_RN_deriv, of N] density_RN_deriv[OF ac]
-    by simp
-qed
-
-lemma (in sigma_finite_measure)
-  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
-    and f: "f \<in> borel_measurable M"
-  shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>
-      integrable M (\<lambda>x. enn2real (RN_deriv M N x) * f x)" (is ?integrable)
-    and RN_deriv_integral: "integral\<^sup>L N f = (\<integral>x. enn2real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
-proof -
-  note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]
-  interpret N: sigma_finite_measure N by fact
-
-  have eq: "density M (RN_deriv M N) = density M (\<lambda>x. enn2real (RN_deriv M N x))"
-  proof (rule density_cong)
-    from RN_deriv_finite[OF assms(1,2,3)]
-    show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
-      by eventually_elim (auto simp: less_top)
-  qed (insert ac, auto)
-
-  show ?integrable
-    apply (subst density_RN_deriv[OF ac, symmetric])
-    unfolding eq
-    apply (intro integrable_real_density f AE_I2 enn2real_nonneg)
-    apply (insert ac, auto)
-    done
-
-  show ?integral
-    apply (subst density_RN_deriv[OF ac, symmetric])
-    unfolding eq
-    apply (intro integral_real_density f AE_I2 enn2real_nonneg)
-    apply (insert ac, auto)
-    done
-qed
-
-lemma (in sigma_finite_measure) real_RN_deriv:
-  assumes "finite_measure N"
-  assumes ac: "absolutely_continuous M N" "sets N = sets M"
-  obtains D where "D \<in> borel_measurable M"
-    and "AE x in M. RN_deriv M N x = ennreal (D x)"
-    and "AE x in N. 0 < D x"
-    and "\<And>x. 0 \<le> D x"
-proof
-  interpret N: finite_measure N by fact
-
-  note RN = borel_measurable_RN_deriv density_RN_deriv[OF ac]
-
-  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}"
-
-  show "(\<lambda>x. enn2real (RN_deriv M N x)) \<in> borel_measurable M"
-    using RN by auto
-
-  have "N (?RN \<infinity>) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
-    using RN(1) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
-  also have "\<dots> = (\<integral>\<^sup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
-    by (intro nn_integral_cong) (auto simp: indicator_def)
-  also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"
-    using RN by (intro nn_integral_cmult_indicator) auto
-  finally have eq: "N (?RN \<infinity>) = \<infinity> * emeasure M (?RN \<infinity>)" .
-  moreover
-  have "emeasure M (?RN \<infinity>) = 0"
-  proof (rule ccontr)
-    assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0"
-    then have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}"
-      by (auto simp: zero_less_iff_neq_zero)
-    with eq have "N (?RN \<infinity>) = \<infinity>" by (simp add: ennreal_mult_eq_top_iff)
-    with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto
-  qed
-  ultimately have "AE x in M. RN_deriv M N x < \<infinity>"
-    using RN by (intro AE_iff_measurable[THEN iffD2]) (auto simp: less_top[symmetric])
-  then show "AE x in M. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
-    by auto
-  then have eq: "AE x in N. RN_deriv M N x = ennreal (enn2real (RN_deriv M N x))"
-    using ac absolutely_continuous_AE by auto
-
-
-  have "N (?RN 0) = (\<integral>\<^sup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
-    by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
-  also have "\<dots> = (\<integral>\<^sup>+ x. 0 \<partial>M)"
-    by (intro nn_integral_cong) (auto simp: indicator_def)
-  finally have "AE x in N. RN_deriv M N x \<noteq> 0"
-    using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)
-  with eq show "AE x in N. 0 < enn2real (RN_deriv M N x)"
-    by (auto simp: enn2real_positive_iff less_top[symmetric] zero_less_iff_neq_zero)
-qed (rule enn2real_nonneg)
-
-lemma (in sigma_finite_measure) RN_deriv_singleton:
-  assumes ac: "absolutely_continuous M N" "sets N = sets M"
-  and x: "{x} \<in> sets M"
-  shows "N {x} = RN_deriv M N x * emeasure M {x}"
-proof -
-  from \<open>{x} \<in> sets M\<close>
-  have "density M (RN_deriv M N) {x} = (\<integral>\<^sup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
-    by (auto simp: indicator_def emeasure_density intro!: nn_integral_cong)
-  with x density_RN_deriv[OF ac] show ?thesis
-    by (auto simp: max_def)
-qed
-
-end
--- a/src/HOL/Probability/Regularity.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,383 +0,0 @@
-(*  Title:      HOL/Probability/Regularity.thy
-    Author:     Fabian Immler, TU München
-*)
-
-section \<open>Regularity of Measures\<close>
-
-theory Regularity
-imports Measure_Space Borel_Space
-begin
-
-lemma
-  fixes M::"'a::{second_countable_topology, complete_space} measure"
-  assumes sb: "sets M = sets borel"
-  assumes "emeasure M (space M) \<noteq> \<infinity>"
-  assumes "B \<in> sets borel"
-  shows inner_regular: "emeasure M B =
-    (SUP K : {K. K \<subseteq> B \<and> compact K}. emeasure M K)" (is "?inner B")
-  and outer_regular: "emeasure M B =
-    (INF U : {U. B \<subseteq> U \<and> open U}. emeasure M U)" (is "?outer B")
-proof -
-  have Us: "UNIV = space M" by (metis assms(1) sets_eq_imp_space_eq space_borel)
-  hence sU: "space M = UNIV" by simp
-  interpret finite_measure M by rule fact
-  have approx_inner: "\<And>A. A \<in> sets M \<Longrightarrow>
-    (\<And>e. e > 0 \<Longrightarrow> \<exists>K. K \<subseteq> A \<and> compact K \<and> emeasure M A \<le> emeasure M K + ennreal e) \<Longrightarrow> ?inner A"
-    by (rule ennreal_approx_SUP)
-      (force intro!: emeasure_mono simp: compact_imp_closed emeasure_eq_measure)+
-  have approx_outer: "\<And>A. A \<in> sets M \<Longrightarrow>
-    (\<And>e. e > 0 \<Longrightarrow> \<exists>B. A \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M A + ennreal e) \<Longrightarrow> ?outer A"
-    by (rule ennreal_approx_INF)
-       (force intro!: emeasure_mono simp: emeasure_eq_measure sb)+
-  from countable_dense_setE guess X::"'a set"  . note X = this
-  {
-    fix r::real assume "r > 0" hence "\<And>y. open (ball y r)" "\<And>y. ball y r \<noteq> {}" by auto
-    with X(2)[OF this]
-    have x: "space M = (\<Union>x\<in>X. cball x r)"
-      by (auto simp add: sU) (metis dist_commute order_less_imp_le)
-    let ?U = "\<Union>k. (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)"
-    have "(\<lambda>k. emeasure M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M ?U"
-      by (rule Lim_emeasure_incseq) (auto intro!: borel_closed bexI simp: incseq_def Us sb)
-    also have "?U = space M"
-    proof safe
-      fix x from X(2)[OF open_ball[of x r]] \<open>r > 0\<close> obtain d where d: "d\<in>X" "d \<in> ball x r" by auto
-      show "x \<in> ?U"
-        using X(1) d
-        by simp (auto intro!: exI [where x = "to_nat_on X d"] simp: dist_commute Bex_def)
-    qed (simp add: sU)
-    finally have "(\<lambda>k. M (\<Union>n\<in>{0..k}. cball (from_nat_into X n) r)) \<longlonglongrightarrow> M (space M)" .
-  } note M_space = this
-  {
-    fix e ::real and n :: nat assume "e > 0" "n > 0"
-    hence "1/n > 0" "e * 2 powr - n > 0" by (auto)
-    from M_space[OF \<open>1/n>0\<close>]
-    have "(\<lambda>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) \<longlonglongrightarrow> measure M (space M)"
-      unfolding emeasure_eq_measure by (auto simp: measure_nonneg)
-    from metric_LIMSEQ_D[OF this \<open>0 < e * 2 powr -n\<close>]
-    obtain k where "dist (measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n))) (measure M (space M)) <
-      e * 2 powr -n"
-      by auto
-    hence "measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge>
-      measure M (space M) - e * 2 powr -real n"
-      by (auto simp: dist_real_def)
-    hence "\<exists>k. measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge>
-      measure M (space M) - e * 2 powr - real n" ..
-  } note k=this
-  hence "\<forall>e\<in>{0<..}. \<forall>(n::nat)\<in>{0<..}. \<exists>k.
-    measure M (\<Union>i\<in>{0..k}. cball (from_nat_into X i) (1/real n)) \<ge> measure M (space M) - e * 2 powr - real n"
-    by blast
-  then obtain k where k: "\<forall>e\<in>{0<..}. \<forall>n\<in>{0<..}. measure M (space M) - e * 2 powr - real (n::nat)
-    \<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
-    by metis
-  hence k: "\<And>e n. e > 0 \<Longrightarrow> n > 0 \<Longrightarrow> measure M (space M) - e * 2 powr - n
-    \<le> measure M (\<Union>i\<in>{0..k e n}. cball (from_nat_into X i) (1 / n))"
-    unfolding Ball_def by blast
-  have approx_space:
-    "\<exists>K \<in> {K. K \<subseteq> space M \<and> compact K}. emeasure M (space M) \<le> emeasure M K + ennreal e"
-    (is "?thesis e") if "0 < e" for e :: real
-  proof -
-    define B where [abs_def]:
-      "B n = (\<Union>i\<in>{0..k e (Suc n)}. cball (from_nat_into X i) (1 / Suc n))" for n
-    have "\<And>n. closed (B n)" by (auto simp: B_def)
-    hence [simp]: "\<And>n. B n \<in> sets M" by (simp add: sb)
-    from k[OF \<open>e > 0\<close> zero_less_Suc]
-    have "\<And>n. measure M (space M) - measure M (B n) \<le> e * 2 powr - real (Suc n)"
-      by (simp add: algebra_simps B_def finite_measure_compl)
-    hence B_compl_le: "\<And>n::nat. measure M (space M - B n) \<le> e * 2 powr - real (Suc n)"
-      by (simp add: finite_measure_compl)
-    define K where "K = (\<Inter>n. B n)"
-    from \<open>closed (B _)\<close> have "closed K" by (auto simp: K_def)
-    hence [simp]: "K \<in> sets M" by (simp add: sb)
-    have "measure M (space M) - measure M K = measure M (space M - K)"
-      by (simp add: finite_measure_compl)
-    also have "\<dots> = emeasure M (\<Union>n. space M - B n)" by (auto simp: K_def emeasure_eq_measure)
-    also have "\<dots> \<le> (\<Sum>n. emeasure M (space M - B n))"
-      by (rule emeasure_subadditive_countably) (auto simp: summable_def)
-    also have "\<dots> \<le> (\<Sum>n. ennreal (e*2 powr - real (Suc n)))"
-      using B_compl_le by (intro suminf_le) (simp_all add: measure_nonneg emeasure_eq_measure ennreal_leI)
-    also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
-      by (simp add: powr_minus powr_realpow field_simps del: of_nat_Suc)
-    also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^ Suc n))"
-      unfolding ennreal_power[symmetric]
-      using \<open>0 < e\<close>
-      by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
-                    ennreal_power[symmetric])
-    also have "\<dots> = e"
-      by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
-    finally have "measure M (space M) \<le> measure M K + e"
-      using \<open>0 < e\<close> by simp
-    hence "emeasure M (space M) \<le> emeasure M K + e"
-      using \<open>0 < e\<close> by (simp add: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] del: ennreal_plus)
-    moreover have "compact K"
-      unfolding compact_eq_totally_bounded
-    proof safe
-      show "complete K" using \<open>closed K\<close> by (simp add: complete_eq_closed)
-      fix e'::real assume "0 < e'"
-      from nat_approx_posE[OF this] guess n . note n = this
-      let ?k = "from_nat_into X ` {0..k e (Suc n)}"
-      have "finite ?k" by simp
-      moreover have "K \<subseteq> (\<Union>x\<in>?k. ball x e')" unfolding K_def B_def using n by force
-      ultimately show "\<exists>k. finite k \<and> K \<subseteq> (\<Union>x\<in>k. ball x e')" by blast
-    qed
-    ultimately
-    show ?thesis by (auto simp: sU)
-  qed
-  { fix A::"'a set" assume "closed A" hence "A \<in> sets borel" by (simp add: compact_imp_closed)
-    hence [simp]: "A \<in> sets M" by (simp add: sb)
-    have "?inner A"
-    proof (rule approx_inner)
-      fix e::real assume "e > 0"
-      from approx_space[OF this] obtain K where
-        K: "K \<subseteq> space M" "compact K" "emeasure M (space M) \<le> emeasure M K + e"
-        by (auto simp: emeasure_eq_measure)
-      hence [simp]: "K \<in> sets M" by (simp add: sb compact_imp_closed)
-      have "measure M A - measure M (A \<inter> K) = measure M (A - A \<inter> K)"
-        by (subst finite_measure_Diff) auto
-      also have "A - A \<inter> K = A \<union> K - K" by auto
-      also have "measure M \<dots> = measure M (A \<union> K) - measure M K"
-        by (subst finite_measure_Diff) auto
-      also have "\<dots> \<le> measure M (space M) - measure M K"
-        by (simp add: emeasure_eq_measure sU sb finite_measure_mono)
-      also have "\<dots> \<le> e"
-        using K \<open>0 < e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
-      finally have "emeasure M A \<le> emeasure M (A \<inter> K) + ennreal e"
-        using \<open>0<e\<close> by (simp add: emeasure_eq_measure algebra_simps ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
-      moreover have "A \<inter> K \<subseteq> A" "compact (A \<inter> K)" using \<open>closed A\<close> \<open>compact K\<close> by auto
-      ultimately show "\<exists>K \<subseteq> A. compact K \<and> emeasure M A \<le> emeasure M K + ennreal e"
-        by blast
-    qed simp
-    have "?outer A"
-    proof cases
-      assume "A \<noteq> {}"
-      let ?G = "\<lambda>d. {x. infdist x A < d}"
-      {
-        fix d
-        have "?G d = (\<lambda>x. infdist x A) -` {..<d}" by auto
-        also have "open \<dots>"
-          by (intro continuous_open_vimage) (auto intro!: continuous_infdist continuous_ident)
-        finally have "open (?G d)" .
-      } note open_G = this
-      from in_closed_iff_infdist_zero[OF \<open>closed A\<close> \<open>A \<noteq> {}\<close>]
-      have "A = {x. infdist x A = 0}" by auto
-      also have "\<dots> = (\<Inter>i. ?G (1/real (Suc i)))"
-      proof (auto simp del: of_nat_Suc, rule ccontr)
-        fix x
-        assume "infdist x A \<noteq> 0"
-        hence pos: "infdist x A > 0" using infdist_nonneg[of x A] by simp
-        from nat_approx_posE[OF this] guess n .
-        moreover
-        assume "\<forall>i. infdist x A < 1 / real (Suc i)"
-        hence "infdist x A < 1 / real (Suc n)" by auto
-        ultimately show False by simp
-      qed
-      also have "M \<dots> = (INF n. emeasure M (?G (1 / real (Suc n))))"
-      proof (rule INF_emeasure_decseq[symmetric], safe)
-        fix i::nat
-        from open_G[of "1 / real (Suc i)"]
-        show "?G (1 / real (Suc i)) \<in> sets M" by (simp add: sb borel_open)
-      next
-        show "decseq (\<lambda>i. {x. infdist x A < 1 / real (Suc i)})"
-          by (auto intro: less_trans intro!: divide_strict_left_mono
-            simp: decseq_def le_eq_less_or_eq)
-      qed simp
-      finally
-      have "emeasure M A = (INF n. emeasure M {x. infdist x A < 1 / real (Suc n)})" .
-      moreover
-      have "\<dots> \<ge> (INF U:{U. A \<subseteq> U \<and> open U}. emeasure M U)"
-      proof (intro INF_mono)
-        fix m
-        have "?G (1 / real (Suc m)) \<in> {U. A \<subseteq> U \<and> open U}" using open_G by auto
-        moreover have "M (?G (1 / real (Suc m))) \<le> M (?G (1 / real (Suc m)))" by simp
-        ultimately show "\<exists>U\<in>{U. A \<subseteq> U \<and> open U}.
-          emeasure M U \<le> emeasure M {x. infdist x A < 1 / real (Suc m)}"
-          by blast
-      qed
-      moreover
-      have "emeasure M A \<le> (INF U:{U. A \<subseteq> U \<and> open U}. emeasure M U)"
-        by (rule INF_greatest) (auto intro!: emeasure_mono simp: sb)
-      ultimately show ?thesis by simp
-    qed (auto intro!: INF_eqI)
-    note \<open>?inner A\<close> \<open>?outer A\<close> }
-  note closed_in_D = this
-  from \<open>B \<in> sets borel\<close>
-  have "Int_stable (Collect closed)" "Collect closed \<subseteq> Pow UNIV" "B \<in> sigma_sets UNIV (Collect closed)"
-    by (auto simp: Int_stable_def borel_eq_closed)
-  then show "?inner B" "?outer B"
-  proof (induct B rule: sigma_sets_induct_disjoint)
-    case empty
-    { case 1 show ?case by (intro SUP_eqI[symmetric]) auto }
-    { case 2 show ?case by (intro INF_eqI[symmetric]) (auto elim!: meta_allE[of _ "{}"]) }
-  next
-    case (basic B)
-    { case 1 from basic closed_in_D show ?case by auto }
-    { case 2 from basic closed_in_D show ?case by auto }
-  next
-    case (compl B)
-    note inner = compl(2) and outer = compl(3)
-    from compl have [simp]: "B \<in> sets M" by (auto simp: sb borel_eq_closed)
-    case 2
-    have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
-    also have "\<dots> = (INF K:{K. K \<subseteq> B \<and> compact K}. M (space M) -  M K)"
-      by (subst ennreal_SUP_const_minus) (auto simp: less_top[symmetric] inner)
-    also have "\<dots> = (INF U:{U. U \<subseteq> B \<and> compact U}. M (space M - U))"
-      by (rule INF_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
-    also have "\<dots> \<ge> (INF U:{U. U \<subseteq> B \<and> closed U}. M (space M - U))"
-      by (rule INF_superset_mono) (auto simp add: compact_imp_closed)
-    also have "(INF U:{U. U \<subseteq> B \<and> closed U}. M (space M - U)) =
-        (INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U)"
-      unfolding INF_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
-        by (rule INF_cong) (auto simp add: sU Compl_eq_Diff_UNIV [symmetric, simp])
-    finally have
-      "(INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<le> emeasure M (space M - B)" .
-    moreover have
-      "(INF U:{U. space M - B \<subseteq> U \<and> open U}. emeasure M U) \<ge> emeasure M (space M - B)"
-      by (auto simp: sb sU intro!: INF_greatest emeasure_mono)
-    ultimately show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb])
-
-    case 1
-    have "M (space M - B) = M (space M) - emeasure M B" by (auto simp: emeasure_compl)
-    also have "\<dots> = (SUP U: {U. B \<subseteq> U \<and> open U}. M (space M) -  M U)"
-      unfolding outer by (subst ennreal_INF_const_minus) auto
-    also have "\<dots> = (SUP U:{U. B \<subseteq> U \<and> open U}. M (space M - U))"
-      by (rule SUP_cong) (auto simp add: emeasure_compl sb compact_imp_closed)
-    also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> closed K}. emeasure M K)"
-      unfolding SUP_image [of _ "\<lambda>u. space M - u" _, symmetric, unfolded comp_def]
-        by (rule SUP_cong) (auto simp add: sU)
-    also have "\<dots> = (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
-    proof (safe intro!: antisym SUP_least)
-      fix K assume "closed K" "K \<subseteq> space M - B"
-      from closed_in_D[OF \<open>closed K\<close>]
-      have K_inner: "emeasure M K = (SUP K:{Ka. Ka \<subseteq> K \<and> compact Ka}. emeasure M K)" by simp
-      show "emeasure M K \<le> (SUP K:{K. K \<subseteq> space M - B \<and> compact K}. emeasure M K)"
-        unfolding K_inner using \<open>K \<subseteq> space M - B\<close>
-        by (auto intro!: SUP_upper SUP_least)
-    qed (fastforce intro!: SUP_least SUP_upper simp: compact_imp_closed)
-    finally show ?case by (auto intro!: antisym simp: sets_eq_imp_space_eq[OF sb])
-  next
-    case (union D)
-    then have "range D \<subseteq> sets M" by (auto simp: sb borel_eq_closed)
-    with union have M[symmetric]: "(\<Sum>i. M (D i)) = M (\<Union>i. D i)" by (intro suminf_emeasure)
-    also have "(\<lambda>n. \<Sum>i<n. M (D i)) \<longlonglongrightarrow> (\<Sum>i. M (D i))"
-      by (intro summable_LIMSEQ) auto
-    finally have measure_LIMSEQ: "(\<lambda>n. \<Sum>i<n. measure M (D i)) \<longlonglongrightarrow> measure M (\<Union>i. D i)"
-      by (simp add: emeasure_eq_measure measure_nonneg setsum_nonneg)
-    have "(\<Union>i. D i) \<in> sets M" using \<open>range D \<subseteq> sets M\<close> by auto
-
-    case 1
-    show ?case
-    proof (rule approx_inner)
-      fix e::real assume "e > 0"
-      with measure_LIMSEQ
-      have "\<exists>no. \<forall>n\<ge>no. \<bar>(\<Sum>i<n. measure M (D i)) -measure M (\<Union>x. D x)\<bar> < e/2"
-        by (auto simp: lim_sequentially dist_real_def simp del: less_divide_eq_numeral1)
-      hence "\<exists>n0. \<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>x. D x)\<bar> < e/2" by auto
-      then obtain n0 where n0: "\<bar>(\<Sum>i<n0. measure M (D i)) - measure M (\<Union>i. D i)\<bar> < e/2"
-        unfolding choice_iff by blast
-      have "ennreal (\<Sum>i<n0. measure M (D i)) = (\<Sum>i<n0. M (D i))"
-        by (auto simp add: emeasure_eq_measure setsum_nonneg measure_nonneg)
-      also have "\<dots> \<le> (\<Sum>i. M (D i))" by (rule setsum_le_suminf) auto
-      also have "\<dots> = M (\<Union>i. D i)" by (simp add: M)
-      also have "\<dots> = measure M (\<Union>i. D i)" by (simp add: emeasure_eq_measure)
-      finally have n0: "measure M (\<Union>i. D i) - (\<Sum>i<n0. measure M (D i)) < e/2"
-        using n0 by (auto simp: measure_nonneg setsum_nonneg)
-      have "\<forall>i. \<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
-      proof
-        fix i
-        from \<open>0 < e\<close> have "0 < e/(2*Suc n0)" by simp
-        have "emeasure M (D i) = (SUP K:{K. K \<subseteq> (D i) \<and> compact K}. emeasure M K)"
-          using union by blast
-        from SUP_approx_ennreal[OF \<open>0 < e/(2*Suc n0)\<close> _ this]
-        show "\<exists>K. K \<subseteq> D i \<and> compact K \<and> emeasure M (D i) \<le> emeasure M K + e/(2*Suc n0)"
-          by (auto simp: emeasure_eq_measure intro: less_imp_le compact_empty)
-      qed
-      then obtain K where K: "\<And>i. K i \<subseteq> D i" "\<And>i. compact (K i)"
-        "\<And>i. emeasure M (D i) \<le> emeasure M (K i) + e/(2*Suc n0)"
-        unfolding choice_iff by blast
-      let ?K = "\<Union>i\<in>{..<n0}. K i"
-      have "disjoint_family_on K {..<n0}" using K \<open>disjoint_family D\<close>
-        unfolding disjoint_family_on_def by blast
-      hence mK: "measure M ?K = (\<Sum>i<n0. measure M (K i))" using K
-        by (intro finite_measure_finite_Union) (auto simp: sb compact_imp_closed)
-      have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (D i)) + e/2" using n0 by simp
-      also have "(\<Sum>i<n0. measure M (D i)) \<le> (\<Sum>i<n0. measure M (K i) + e/(2*Suc n0))"
-        using K \<open>0 < e\<close>
-        by (auto intro: setsum_mono simp: emeasure_eq_measure measure_nonneg ennreal_plus[symmetric] simp del: ennreal_plus)
-      also have "\<dots> = (\<Sum>i<n0. measure M (K i)) + (\<Sum>i<n0. e/(2*Suc n0))"
-        by (simp add: setsum.distrib)
-      also have "\<dots> \<le> (\<Sum>i<n0. measure M (K i)) +  e / 2" using \<open>0 < e\<close>
-        by (auto simp: field_simps intro!: mult_left_mono)
-      finally
-      have "measure M (\<Union>i. D i) < (\<Sum>i<n0. measure M (K i)) + e / 2 + e / 2"
-        by auto
-      hence "M (\<Union>i. D i) < M ?K + e"
-        using \<open>0<e\<close> by (auto simp: mK emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] simp del: ennreal_plus)
-      moreover
-      have "?K \<subseteq> (\<Union>i. D i)" using K by auto
-      moreover
-      have "compact ?K" using K by auto
-      ultimately
-      have "?K\<subseteq>(\<Union>i. D i) \<and> compact ?K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M ?K + ennreal e" by simp
-      thus "\<exists>K\<subseteq>\<Union>i. D i. compact K \<and> emeasure M (\<Union>i. D i) \<le> emeasure M K + ennreal e" ..
-    qed fact
-    case 2
-    show ?case
-    proof (rule approx_outer[OF \<open>(\<Union>i. D i) \<in> sets M\<close>])
-      fix e::real assume "e > 0"
-      have "\<forall>i::nat. \<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
-      proof
-        fix i::nat
-        from \<open>0 < e\<close> have "0 < e/(2 powr Suc i)" by simp
-        have "emeasure M (D i) = (INF U:{U. (D i) \<subseteq> U \<and> open U}. emeasure M U)"
-          using union by blast
-        from INF_approx_ennreal[OF \<open>0 < e/(2 powr Suc i)\<close> this]
-        show "\<exists>U. D i \<subseteq> U \<and> open U \<and> e/(2 powr Suc i) > emeasure M U - emeasure M (D i)"
-          using \<open>0<e\<close>
-          by (auto simp: emeasure_eq_measure measure_nonneg setsum_nonneg ennreal_less_iff ennreal_plus[symmetric] ennreal_minus
-                         finite_measure_mono sb
-                   simp del: ennreal_plus)
-      qed
-      then obtain U where U: "\<And>i. D i \<subseteq> U i" "\<And>i. open (U i)"
-        "\<And>i. e/(2 powr Suc i) > emeasure M (U i) - emeasure M (D i)"
-        unfolding choice_iff by blast
-      let ?U = "\<Union>i. U i"
-      have "ennreal (measure M ?U - measure M (\<Union>i. D i)) = M ?U - M (\<Union>i. D i)"
-        using U(1,2)
-        by (subst ennreal_minus[symmetric])
-           (auto intro!: finite_measure_mono simp: sb measure_nonneg emeasure_eq_measure)
-      also have "\<dots> = M (?U - (\<Union>i. D i))" using U  \<open>(\<Union>i. D i) \<in> sets M\<close>
-        by (subst emeasure_Diff) (auto simp: sb)
-      also have "\<dots> \<le> M (\<Union>i. U i - D i)" using U  \<open>range D \<subseteq> sets M\<close>
-        by (intro emeasure_mono) (auto simp: sb intro!: sets.countable_nat_UN sets.Diff)
-      also have "\<dots> \<le> (\<Sum>i. M (U i - D i))" using U  \<open>range D \<subseteq> sets M\<close>
-        by (intro emeasure_subadditive_countably) (auto intro!: sets.Diff simp: sb)
-      also have "\<dots> \<le> (\<Sum>i. ennreal e/(2 powr Suc i))" using U \<open>range D \<subseteq> sets M\<close>
-        using \<open>0<e\<close>
-        by (intro suminf_le, subst emeasure_Diff)
-           (auto simp: emeasure_Diff emeasure_eq_measure sb measure_nonneg ennreal_minus
-                       finite_measure_mono divide_ennreal ennreal_less_iff
-                 intro: less_imp_le)
-      also have "\<dots> \<le> (\<Sum>n. ennreal (e * (1 / 2) ^ Suc n))"
-        using \<open>0<e\<close>
-        by (simp add: powr_minus powr_realpow field_simps divide_ennreal del: of_nat_Suc)
-      also have "\<dots> = ennreal e * (\<Sum>n. ennreal ((1 / 2) ^  Suc n))"
-        unfolding ennreal_power[symmetric]
-        using \<open>0 < e\<close>
-        by (simp add: ac_simps ennreal_mult' divide_ennreal[symmetric] divide_ennreal_def
-                      ennreal_power[symmetric])
-      also have "\<dots> = ennreal e"
-        by (subst suminf_ennreal_eq[OF zero_le_power power_half_series]) auto
-      finally have "emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e"
-        using \<open>0<e\<close> by (simp add: emeasure_eq_measure ennreal_plus[symmetric] measure_nonneg del: ennreal_plus)
-      moreover
-      have "(\<Union>i. D i) \<subseteq> ?U" using U by auto
-      moreover
-      have "open ?U" using U by auto
-      ultimately
-      have "(\<Union>i. D i) \<subseteq> ?U \<and> open ?U \<and> emeasure M ?U \<le> emeasure M (\<Union>i. D i) + ennreal e" by simp
-      thus "\<exists>B. (\<Union>i. D i) \<subseteq> B \<and> open B \<and> emeasure M B \<le> emeasure M (\<Union>i. D i) + ennreal e" ..
-    qed
-  qed
-qed
-
-end
-
--- a/src/HOL/Probability/SPMF.thy	Sun Aug 07 12:10:49 2016 +0200
+++ b/src/HOL/Probability/SPMF.thy	Fri Aug 05 18:34:57 2016 +0200
@@ -4,7 +4,6 @@
 
 theory SPMF imports
   Probability_Mass_Function
-  Embed_Measure
   "~~/src/HOL/Library/Complete_Partial_Order2"
   "~~/src/HOL/Library/Rewrite"
 begin
--- a/src/HOL/Probability/Set_Integral.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,602 +0,0 @@
-(*  Title:      HOL/Probability/Set_Integral.thy
-    Author:     Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU)
-
-Notation and useful facts for working with integrals over a set.
-
-TODO: keep all these? Need unicode translations as well.
-*)
-
-theory Set_Integral
-  imports Bochner_Integration Lebesgue_Measure
-begin
-
-(*
-    Notation
-*)
-
-abbreviation "set_borel_measurable M A f \<equiv> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable M"
-
-abbreviation "set_integrable M A f \<equiv> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
-
-abbreviation "set_lebesgue_integral M A f \<equiv> lebesgue_integral M (\<lambda>x. indicator A x *\<^sub>R f x)"
-
-syntax
-"_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
-("(4LINT (_):(_)/|(_)./ _)" [0,60,110,61] 60)
-
-translations
-"LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
-
-abbreviation
-  "set_almost_everywhere A M P \<equiv> AE x in M. x \<in> A \<longrightarrow> P x"
-
-syntax
-  "_set_almost_everywhere" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool"
-("AE _\<in>_ in _./ _" [0,0,0,10] 10)
-
-translations
-  "AE x\<in>A in M. P" == "CONST set_almost_everywhere A M (\<lambda>x. P)"
-
-(*
-    Notation for integration wrt lebesgue measure on the reals:
-
-      LBINT x. f
-      LBINT x : A. f
-
-    TODO: keep all these? Need unicode.
-*)
-
-syntax
-"_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real"
-("(2LBINT _./ _)" [0,60] 60)
-
-translations
-"LBINT x. f" == "CONST lebesgue_integral CONST lborel (\<lambda>x. f)"
-
-syntax
-"_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> real"
-("(3LBINT _:_./ _)" [0,60,61] 60)
-
-translations
-"LBINT x:A. f" == "CONST set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
-
-(*
-    Basic properties
-*)
-
-(*
-lemma indicator_abs_eq: "\<And>A x. \<bar>indicator A x\<bar> = ((indicator A x) :: real)"
-  by (auto simp add: indicator_def)
-*)
-
-lemma set_borel_measurable_sets:
-  fixes f :: "_ \<Rightarrow> _::real_normed_vector"
-  assumes "set_borel_measurable M X f" "B \<in> sets borel" "X \<in> sets M"
-  shows "f -` B \<inter> X \<in> sets M"
-proof -
-  have "f \<in> borel_measurable (restrict_space M X)"
-    using assms by (subst borel_measurable_restrict_space_iff) auto
-  then have "f -` B \<inter> space (restrict_space M X) \<in> sets (restrict_space M X)"
-    by (rule measurable_sets) fact
-  with \<open>X \<in> sets M\<close> show ?thesis
-    by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space)
-qed
-
-lemma set_lebesgue_integral_cong:
-  assumes "A \<in> sets M" and "\<forall>x. x \<in> A \<longrightarrow> f x = g x"
-  shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)"
-  using assms by (auto intro!: integral_cong split: split_indicator simp add: sets.sets_into_space)
-
-lemma set_lebesgue_integral_cong_AE:
-  assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes "AE x \<in> A in M. f x = g x"
-  shows "LINT x:A|M. f x = LINT x:A|M. g x"
-proof-
-  have "AE x in M. indicator A x *\<^sub>R f x = indicator A x *\<^sub>R g x"
-    using assms by auto
-  thus ?thesis by (intro integral_cong_AE) auto
-qed
-
-lemma set_integrable_cong_AE:
-    "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
-    AE x \<in> A in M. f x = g x \<Longrightarrow> A \<in> sets M \<Longrightarrow>
-    set_integrable M A f = set_integrable M A g"
-  by (rule integrable_cong_AE) auto
-
-lemma set_integrable_subset:
-  fixes M A B and f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes "set_integrable M A f" "B \<in> sets M" "B \<subseteq> A"
-  shows "set_integrable M B f"
-proof -
-  have "set_integrable M B (\<lambda>x. indicator A x *\<^sub>R f x)"
-    by (rule integrable_mult_indicator) fact+
-  with \<open>B \<subseteq> A\<close> show ?thesis
-    by (simp add: indicator_inter_arith[symmetric] Int_absorb2)
-qed
-
-(* TODO: integral_cmul_indicator should be named set_integral_const *)
-(* TODO: borel_integrable_atLeastAtMost should be named something like set_integrable_Icc_isCont *)
-
-lemma set_integral_scaleR_right [simp]: "LINT t:A|M. a *\<^sub>R f t = a *\<^sub>R (LINT t:A|M. f t)"
-  by (subst integral_scaleR_right[symmetric]) (auto intro!: integral_cong)
-
-lemma set_integral_mult_right [simp]:
-  fixes a :: "'a::{real_normed_field, second_countable_topology}"
-  shows "LINT t:A|M. a * f t = a * (LINT t:A|M. f t)"
-  by (subst integral_mult_right_zero[symmetric]) (auto intro!: integral_cong)
-
-lemma set_integral_mult_left [simp]:
-  fixes a :: "'a::{real_normed_field, second_countable_topology}"
-  shows "LINT t:A|M. f t * a = (LINT t:A|M. f t) * a"
-  by (subst integral_mult_left_zero[symmetric]) (auto intro!: integral_cong)
-
-lemma set_integral_divide_zero [simp]:
-  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
-  shows "LINT t:A|M. f t / a = (LINT t:A|M. f t) / a"
-  by (subst integral_divide_zero[symmetric], intro integral_cong)
-     (auto split: split_indicator)
-
-lemma set_integrable_scaleR_right [simp, intro]:
-  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a *\<^sub>R f t)"
-  unfolding scaleR_left_commute by (rule integrable_scaleR_right)
-
-lemma set_integrable_scaleR_left [simp, intro]:
-  fixes a :: "_ :: {banach, second_countable_topology}"
-  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t *\<^sub>R a)"
-  using integrable_scaleR_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
-
-lemma set_integrable_mult_right [simp, intro]:
-  fixes a :: "'a::{real_normed_field, second_countable_topology}"
-  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a * f t)"
-  using integrable_mult_right[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
-
-lemma set_integrable_mult_left [simp, intro]:
-  fixes a :: "'a::{real_normed_field, second_countable_topology}"
-  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t * a)"
-  using integrable_mult_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
-
-lemma set_integrable_divide [simp, intro]:
-  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
-  assumes "a \<noteq> 0 \<Longrightarrow> set_integrable M A f"
-  shows "set_integrable M A (\<lambda>t. f t / a)"
-proof -
-  have "integrable M (\<lambda>x. indicator A x *\<^sub>R f x / a)"
-    using assms by (rule integrable_divide_zero)
-  also have "(\<lambda>x. indicator A x *\<^sub>R f x / a) = (\<lambda>x. indicator A x *\<^sub>R (f x / a))"
-    by (auto split: split_indicator)
-  finally show ?thesis .
-qed
-
-lemma set_integral_add [simp, intro]:
-  fixes f g :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes "set_integrable M A f" "set_integrable M A g"
-  shows "set_integrable M A (\<lambda>x. f x + g x)"
-    and "LINT x:A|M. f x + g x = (LINT x:A|M. f x) + (LINT x:A|M. g x)"
-  using assms by (simp_all add: scaleR_add_right)
-
-lemma set_integral_diff [simp, intro]:
-  assumes "set_integrable M A f" "set_integrable M A g"
-  shows "set_integrable M A (\<lambda>x. f x - g x)" and "LINT x:A|M. f x - g x =
-    (LINT x:A|M. f x) - (LINT x:A|M. g x)"
-  using assms by (simp_all add: scaleR_diff_right)
-
-lemma set_integral_reflect:
-  fixes S and f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  shows "(LBINT x : S. f x) = (LBINT x : {x. - x \<in> S}. f (- x))"
-  by (subst lborel_integral_real_affine[where c="-1" and t=0])
-     (auto intro!: integral_cong split: split_indicator)
-
-(* question: why do we have this for negation, but multiplication by a constant
-   requires an integrability assumption? *)
-lemma set_integral_uminus: "set_integrable M A f \<Longrightarrow> LINT x:A|M. - f x = - (LINT x:A|M. f x)"
-  by (subst integral_minus[symmetric]) simp_all
-
-lemma set_integral_complex_of_real:
-  "LINT x:A|M. complex_of_real (f x) = of_real (LINT x:A|M. f x)"
-  by (subst integral_complex_of_real[symmetric])
-     (auto intro!: integral_cong split: split_indicator)
-
-lemma set_integral_mono:
-  fixes f g :: "_ \<Rightarrow> real"
-  assumes "set_integrable M A f" "set_integrable M A g"
-    "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"
-  shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)"
-using assms by (auto intro: integral_mono split: split_indicator)
-
-lemma set_integral_mono_AE:
-  fixes f g :: "_ \<Rightarrow> real"
-  assumes "set_integrable M A f" "set_integrable M A g"
-    "AE x \<in> A in M. f x \<le> g x"
-  shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)"
-using assms by (auto intro: integral_mono_AE split: split_indicator)
-
-lemma set_integrable_abs: "set_integrable M A f \<Longrightarrow> set_integrable M A (\<lambda>x. \<bar>f x\<bar> :: real)"
-  using integrable_abs[of M "\<lambda>x. f x * indicator A x"] by (simp add: abs_mult ac_simps)
-
-lemma set_integrable_abs_iff:
-  fixes f :: "_ \<Rightarrow> real"
-  shows "set_borel_measurable M A f \<Longrightarrow> set_integrable M A (\<lambda>x. \<bar>f x\<bar>) = set_integrable M A f"
-  by (subst (2) integrable_abs_iff[symmetric]) (simp_all add: abs_mult ac_simps)
-
-lemma set_integrable_abs_iff':
-  fixes f :: "_ \<Rightarrow> real"
-  shows "f \<in> borel_measurable M \<Longrightarrow> A \<in> sets M \<Longrightarrow>
-    set_integrable M A (\<lambda>x. \<bar>f x\<bar>) = set_integrable M A f"
-by (intro set_integrable_abs_iff) auto
-
-lemma set_integrable_discrete_difference:
-  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
-  assumes "countable X"
-  assumes diff: "(A - B) \<union> (B - A) \<subseteq> X"
-  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
-  shows "set_integrable M A f \<longleftrightarrow> set_integrable M B f"
-proof (rule integrable_discrete_difference[where X=X])
-  show "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x"
-    using diff by (auto split: split_indicator)
-qed fact+
-
-lemma set_integral_discrete_difference:
-  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
-  assumes "countable X"
-  assumes diff: "(A - B) \<union> (B - A) \<subseteq> X"
-  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
-  shows "set_lebesgue_integral M A f = set_lebesgue_integral M B f"
-proof (rule integral_discrete_difference[where X=X])
-  show "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x"
-    using diff by (auto split: split_indicator)
-qed fact+
-
-lemma set_integrable_Un:
-  fixes f g :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes f_A: "set_integrable M A f" and f_B:  "set_integrable M B f"
-    and [measurable]: "A \<in> sets M" "B \<in> sets M"
-  shows "set_integrable M (A \<union> B) f"
-proof -
-  have "set_integrable M (A - B) f"
-    using f_A by (rule set_integrable_subset) auto
-  from integrable_add[OF this f_B] show ?thesis
-    by (rule integrable_cong[THEN iffD1, rotated 2]) (auto split: split_indicator)
-qed
-
-lemma set_integrable_UN:
-  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> set_integrable M (A i) f"
-    "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M"
-  shows "set_integrable M (\<Union>i\<in>I. A i) f"
-using assms by (induct I) (auto intro!: set_integrable_Un)
-
-lemma set_integral_Un:
-  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes "A \<inter> B = {}"
-  and "set_integrable M A f"
-  and "set_integrable M B f"
-  shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
-by (auto simp add: indicator_union_arith indicator_inter_arith[symmetric]
-      scaleR_add_left assms)
-
-lemma set_integral_cong_set:
-  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes [measurable]: "set_borel_measurable M A f" "set_borel_measurable M B f"
-    and ae: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
-  shows "LINT x:B|M. f x = LINT x:A|M. f x"
-proof (rule integral_cong_AE)
-  show "AE x in M. indicator B x *\<^sub>R f x = indicator A x *\<^sub>R f x"
-    using ae by (auto simp: subset_eq split: split_indicator)
-qed fact+
-
-lemma set_borel_measurable_subset:
-  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes [measurable]: "set_borel_measurable M A f" "B \<in> sets M" and "B \<subseteq> A"
-  shows "set_borel_measurable M B f"
-proof -
-  have "set_borel_measurable M B (\<lambda>x. indicator A x *\<^sub>R f x)"
-    by measurable
-  also have "(\<lambda>x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\<lambda>x. indicator B x *\<^sub>R f x)"
-    using \<open>B \<subseteq> A\<close> by (auto simp: fun_eq_iff split: split_indicator)
-  finally show ?thesis .
-qed
-
-lemma set_integral_Un_AE:
-  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes ae: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)" and [measurable]: "A \<in> sets M" "B \<in> sets M"
-  and "set_integrable M A f"
-  and "set_integrable M B f"
-  shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
-proof -
-  have f: "set_integrable M (A \<union> B) f"
-    by (intro set_integrable_Un assms)
-  then have f': "set_borel_measurable M (A \<union> B) f"
-    by (rule borel_measurable_integrable)
-  have "LINT x:A\<union>B|M. f x = LINT x:(A - A \<inter> B) \<union> (B - A \<inter> B)|M. f x"
-  proof (rule set_integral_cong_set)
-    show "AE x in M. (x \<in> A - A \<inter> B \<union> (B - A \<inter> B)) = (x \<in> A \<union> B)"
-      using ae by auto
-    show "set_borel_measurable M (A - A \<inter> B \<union> (B - A \<inter> B)) f"
-      using f' by (rule set_borel_measurable_subset) auto
-  qed fact
-  also have "\<dots> = (LINT x:(A - A \<inter> B)|M. f x) + (LINT x:(B - A \<inter> B)|M. f x)"
-    by (auto intro!: set_integral_Un set_integrable_subset[OF f])
-  also have "\<dots> = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
-    using ae
-    by (intro arg_cong2[where f="op+"] set_integral_cong_set)
-       (auto intro!: set_borel_measurable_subset[OF f'])
-  finally show ?thesis .
-qed
-
-lemma set_integral_finite_Union:
-  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
-  assumes "finite I" "disjoint_family_on A I"
-    and "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
-  shows "(LINT x:(\<Union>i\<in>I. A i)|M. f x) = (\<Sum>i\<in>I. LINT x:A i|M. f x)"
-  using assms
-  apply induct
-  apply (auto intro!: set_integral_Un set_integrable_Un set_integrable_UN simp: disjoint_family_on_def)
-by (subst set_integral_Un, auto intro: set_integrable_UN)
-
-(* TODO: find a better name? *)
-lemma pos_integrable_to_top:
-  fixes l::real
-  assumes "\<And>i. A i \<in> sets M" "mono A"
-  assumes nneg: "\<And>x i. x \<in> A i \<Longrightarrow> 0 \<le> f x"
-  and intgbl: "\<And>i::nat. set_integrable M (A i) f"
-  and lim: "(\<lambda>i::nat. LINT x:A i|M. f x) \<longlonglongrightarrow> l"
-  shows "set_integrable M (\<Union>i. A i) f"
-  apply (rule integrable_monotone_convergence[where f = "\<lambda>i::nat. \<lambda>x. indicator (A i) x *\<^sub>R f x" and x = l])
-  apply (rule intgbl)
-  prefer 3 apply (rule lim)
-  apply (rule AE_I2)
-  using \<open>mono A\<close> apply (auto simp: mono_def nneg split: split_indicator) []
-proof (rule AE_I2)
-  { fix x assume "x \<in> space M"
-    show "(\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Union>i. A i) x *\<^sub>R f x"
-    proof cases
-      assume "\<exists>i. x \<in> A i"
-      then guess i ..
-      then have *: "eventually (\<lambda>i. x \<in> A i) sequentially"
-        using \<open>x \<in> A i\<close> \<open>mono A\<close> by (auto simp: eventually_sequentially mono_def)
-      show ?thesis
-        apply (intro Lim_eventually)
-        using *
-        apply eventually_elim
-        apply (auto split: split_indicator)
-        done
-    qed auto }
-  then show "(\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R f x) \<in> borel_measurable M"
-    apply (rule borel_measurable_LIMSEQ_real)
-    apply assumption
-    apply (intro borel_measurable_integrable intgbl)
-    done
-qed
-
-(* Proof from Royden Real Analysis, p. 91. *)
-lemma lebesgue_integral_countable_add:
-  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  assumes meas[intro]: "\<And>i::nat. A i \<in> sets M"
-    and disj: "\<And>i j. i \<noteq> j \<Longrightarrow> A i \<inter> A j = {}"
-    and intgbl: "set_integrable M (\<Union>i. A i) f"
-  shows "LINT x:(\<Union>i. A i)|M. f x = (\<Sum>i. (LINT x:(A i)|M. f x))"
-proof (subst integral_suminf[symmetric])
-  show int_A: "\<And>i. set_integrable M (A i) f"
-    using intgbl by (rule set_integrable_subset) auto
-  { fix x assume "x \<in> space M"
-    have "(\<lambda>i. indicator (A i) x *\<^sub>R f x) sums (indicator (\<Union>i. A i) x *\<^sub>R f x)"
-      by (intro sums_scaleR_left indicator_sums) fact }
-  note sums = this
-
-  have norm_f: "\<And>i. set_integrable M (A i) (\<lambda>x. norm (f x))"
-    using int_A[THEN integrable_norm] by auto
-
-  show "AE x in M. summable (\<lambda>i. norm (indicator (A i) x *\<^sub>R f x))"
-    using disj by (intro AE_I2) (auto intro!: summable_mult2 sums_summable[OF indicator_sums])
-
-  show "summable (\<lambda>i. LINT x|M. norm (indicator (A i) x *\<^sub>R f x))"
-  proof (rule summableI_nonneg_bounded)
-    fix n
-    show "0 \<le> LINT x|M. norm (indicator (A n) x *\<^sub>R f x)"
-      using norm_f by (auto intro!: integral_nonneg_AE)
-
-    have "(\<Sum>i<n. LINT x|M. norm (indicator (A i) x *\<^sub>R f x)) =
-      (\<Sum>i<n. set_lebesgue_integral M (A i) (\<lambda>x. norm (f x)))"
-      by (simp add: abs_mult)
-    also have "\<dots> = set_lebesgue_integral M (\<Union>i<n. A i) (\<lambda>x. norm (f x))"
-      using norm_f
-      by (subst set_integral_finite_Union) (auto simp: disjoint_family_on_def disj)
-    also have "\<dots> \<le> set_lebesgue_integral M (\<Union>i. A i) (\<lambda>x. norm (f x))"
-      using intgbl[THEN integrable_norm]
-      by (intro integral_mono set_integrable_UN[of "{..<n}"] norm_f)
-         (auto split: split_indicator)
-    finally show "(\<Sum>i<n. LINT x|M. norm (indicator (A i) x *\<^sub>R f x)) \<le>
-      set_lebesgue_integral M (\<Union>i. A i) (\<lambda>x. norm (f x))"
-      by simp
-  qed
-  show "set_lebesgue_integral M (UNION UNIV A) f = LINT x|M. (\<Sum>i. indicator (A i) x *\<^sub>R f x)"
-    apply (rule integral_cong[OF refl])
-    apply (subst suminf_scaleR_left[OF sums_summable[OF indicator_sums, OF disj], symmetric])
-    using sums_unique[OF indicator_sums[OF disj]]
-    apply auto
-    done
-qed
-
-lemma set_integral_cont_up:
-  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  assumes [measurable]: "\<And>i. A i \<in> sets M" and A: "incseq A"
-  and intgbl: "set_integrable M (\<Union>i. A i) f"
-  shows "(\<lambda>i. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Union>i. A i)|M. f x"
-proof (intro integral_dominated_convergence[where w="\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R norm (f x)"])
-  have int_A: "\<And>i. set_integrable M (A i) f"
-    using intgbl by (rule set_integrable_subset) auto
-  then show "\<And>i. set_borel_measurable M (A i) f" "set_borel_measurable M (\<Union>i. A i) f"
-    "set_integrable M (\<Union>i. A i) (\<lambda>x. norm (f x))"
-    using intgbl integrable_norm[OF intgbl] by auto
-
-  { fix x i assume "x \<in> A i"
-    with A have "(\<lambda>xa. indicator (A xa) x::real) \<longlonglongrightarrow> 1 \<longleftrightarrow> (\<lambda>xa. 1::real) \<longlonglongrightarrow> 1"
-      by (intro filterlim_cong refl)
-         (fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) }
-  then show "AE x in M. (\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Union>i. A i) x *\<^sub>R f x"
-    by (intro AE_I2 tendsto_intros) (auto split: split_indicator)
-qed (auto split: split_indicator)
-
-(* Can the int0 hypothesis be dropped? *)
-lemma set_integral_cont_down:
-  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
-  assumes [measurable]: "\<And>i. A i \<in> sets M" and A: "decseq A"
-  and int0: "set_integrable M (A 0) f"
-  shows "(\<lambda>i::nat. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Inter>i. A i)|M. f x"
-proof (rule integral_dominated_convergence)
-  have int_A: "\<And>i. set_integrable M (A i) f"
-    using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
-  show "set_integrable M (A 0) (\<lambda>x. norm (f x))"
-    using int0[THEN integrable_norm] by simp
-  have "set_integrable M (\<Inter>i. A i) f"
-    using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
-  with int_A show "set_borel_measurable M (\<Inter>i. A i) f" "\<And>i. set_borel_measurable M (A i) f"
-    by auto
-  show "\<And>i. AE x in M. norm (indicator (A i) x *\<^sub>R f x) \<le> indicator (A 0) x *\<^sub>R norm (f x)"
-    using A by (auto split: split_indicator simp: decseq_def)
-  { fix x i assume "x \<in> space M" "x \<notin> A i"
-    with A have "(\<lambda>i. indicator (A i) x::real) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<lambda>i. 0::real) \<longlonglongrightarrow> 0"
-      by (intro filterlim_cong refl)
-         (auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) }
-  then show "AE x in M. (\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Inter>i. A i) x *\<^sub>R f x"
-    by (intro AE_I2 tendsto_intros) (auto split: split_indicator)
-qed
-
-lemma set_integral_at_point:
-  fixes a :: real
-  assumes "set_integrable M {a} f"
-  and [simp]: "{a} \<in> sets M" and "(emeasure M) {a} \<noteq> \<infinity>"
-  shows "(LINT x:{a} | M. f x) = f a * measure M {a}"
-proof-
-  have "set_lebesgue_integral M {a} f = set_lebesgue_integral M {a} (%x. f a)"
-    by (intro set_lebesgue_integral_cong) simp_all
-  then show ?thesis using assms by simp
-qed
-
-
-abbreviation complex_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
-  "complex_integrable M f \<equiv> integrable M f"
-
-abbreviation complex_lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" ("integral\<^sup>C") where
-  "integral\<^sup>C M f == integral\<^sup>L M f"
-
-syntax
-  "_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex"
- ("\<integral>\<^sup>C _. _ \<partial>_" [60,61] 110)
-
-translations
-  "\<integral>\<^sup>Cx. f \<partial>M" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"
-
-syntax
-  "_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
-  ("(3CLINT _|_. _)" [0,110,60] 60)
-
-translations
-  "CLINT x|M. f" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"
-
-lemma complex_integrable_cnj [simp]:
-  "complex_integrable M (\<lambda>x. cnj (f x)) \<longleftrightarrow> complex_integrable M f"
-proof
-  assume "complex_integrable M (\<lambda>x. cnj (f x))"
-  then have "complex_integrable M (\<lambda>x. cnj (cnj (f x)))"
-    by (rule integrable_cnj)
-  then show "complex_integrable M f"
-    by simp
-qed simp
-
-lemma complex_of_real_integrable_eq:
-  "complex_integrable M (\<lambda>x. complex_of_real (f x)) \<longleftrightarrow> integrable M f"
-proof
-  assume "complex_integrable M (\<lambda>x. complex_of_real (f x))"
-  then have "integrable M (\<lambda>x. Re (complex_of_real (f x)))"
-    by (rule integrable_Re)
-  then show "integrable M f"
-    by simp
-qed simp
-
-
-abbreviation complex_set_integrable :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
-  "complex_set_integrable M A f \<equiv> set_integrable M A f"
-
-abbreviation complex_set_lebesgue_integral :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" where
-  "complex_set_lebesgue_integral M A f \<equiv> set_lebesgue_integral M A f"
-
-syntax
-"_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
-("(4CLINT _:_|_. _)" [0,60,110,61] 60)
-
-translations
-"CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\<lambda>x. f)"
-
-(*
-lemma cmod_mult: "cmod ((a :: real) * (x :: complex)) = \<bar>a\<bar> * cmod x"
-  apply (simp add: norm_mult)
-  by (subst norm_mult, auto)
-*)
-
-lemma borel_integrable_atLeastAtMost':
-  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
-  assumes f: "continuous_on {a..b} f"
-  shows "set_integrable lborel {a..b} f" (is "integrable _ ?f")
-  by (intro borel_integrable_compact compact_Icc f)
-
-lemma integral_FTC_atLeastAtMost:
-  fixes f :: "real \<Rightarrow> 'a :: euclidean_space"
-  assumes "a \<le> b"
-    and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (F has_vector_derivative f x) (at x within {a .. b})"
-    and f: "continuous_on {a .. b} f"
-  shows "integral\<^sup>L lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x) = F b - F a"
-proof -
-  let ?f = "\<lambda>x. indicator {a .. b} x *\<^sub>R f x"
-  have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV"
-    using borel_integrable_atLeastAtMost'[OF f] by (rule has_integral_integral_lborel)
-  moreover
-  have "(f has_integral F b - F a) {a .. b}"
-    by (intro fundamental_theorem_of_calculus ballI assms) auto
-  then have "(?f has_integral F b - F a) {a .. b}"
-    by (subst has_integral_cong[where g=f]) auto
-  then have "(?f has_integral F b - F a) UNIV"
-    by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto
-  ultimately show "integral\<^sup>L lborel ?f = F b - F a"
-    by (rule has_integral_unique)
-qed
-
-lemma set_borel_integral_eq_integral:
-  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
-  assumes "set_integrable lborel S f"
-  shows "f integrable_on S" "LINT x : S | lborel. f x = integral S f"
-proof -
-  let ?f = "\<lambda>x. indicator S x *\<^sub>R f x"
-  have "(?f has_integral LINT x : S | lborel. f x) UNIV"
-    by (rule has_integral_integral_lborel) fact
-  hence 1: "(f has_integral (set_lebesgue_integral lborel S f)) S"
-    apply (subst has_integral_restrict_univ [symmetric])
-    apply (rule has_integral_eq)
-    by auto
-  thus "f integrable_on S"
-    by (auto simp add: integrable_on_def)
-  with 1 have "(f has_integral (integral S f)) S"
-    by (intro integrable_integral, auto simp add: integrable_on_def)
-  thus "LINT x : S | lborel. f x = integral S f"
-    by (intro has_integral_unique [OF 1])
-qed
-
-lemma set_borel_measurable_continuous:
-  fixes f :: "_ \<Rightarrow> _::real_normed_vector"
-  assumes "S \<in> sets borel" "continuous_on S f"
-  shows "set_borel_measurable borel S f"
-proof -
-  have "(\<lambda>x. if x \<in> S then f x else 0) \<in> borel_measurable borel"
-    by (intro assms borel_measurable_continuous_on_if continuous_on_const)
-  also have "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. indicator S x *\<^sub>R f x)"
-    by auto
-  finally show ?thesis .
-qed
-
-lemma set_measurable_continuous_on_ivl:
-  assumes "continuous_on {a..b} (f :: real \<Rightarrow> real)"
-  shows "set_borel_measurable borel {a..b} f"
-  by (rule set_borel_measurable_continuous[OF _ assms]) simp
-
-end
-
--- a/src/HOL/Probability/Sigma_Algebra.thy	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,2262 +0,0 @@
-(*  Title:      HOL/Probability/Sigma_Algebra.thy
-    Author:     Stefan Richter, Markus Wenzel, TU München
-    Author:     Johannes Hölzl, TU München
-    Plus material from the Hurd/Coble measure theory development,
-    translated by Lawrence Paulson.
-*)
-
-section \<open>Describing measurable sets\<close>
-
-theory Sigma_Algebra
-imports
-  Complex_Main
-  "~~/src/HOL/Library/Countable_Set"
-  "~~/src/HOL/Library/FuncSet"
-  "~~/src/HOL/Library/Indicator_Function"
-  "~~/src/HOL/Library/Extended_Nonnegative_Real"
-  "~~/src/HOL/Library/Disjoint_Sets"
-begin
-
-text \<open>Sigma algebras are an elementary concept in measure
-  theory. To measure --- that is to integrate --- functions, we first have
-  to measure sets. Unfortunately, when dealing with a large universe,
-  it is often not possible to consistently assign a measure to every
-  subset. Therefore it is necessary to define the set of measurable
-  subsets of the universe. A sigma algebra is such a set that has
-  three very natural and desirable properties.\<close>
-
-subsection \<open>Families of sets\<close>
-
-locale subset_class =
-  fixes \<Omega> :: "'a set" and M :: "'a set set"
-  assumes space_closed: "M \<subseteq> Pow \<Omega>"
-
-lemma (in subset_class) sets_into_space: "x \<in> M \<Longrightarrow> x \<subseteq> \<Omega>"
-  by (metis PowD contra_subsetD space_closed)
-
-subsubsection \<open>Semiring of sets\<close>
-
-locale semiring_of_sets = subset_class +
-  assumes empty_sets[iff]: "{} \<in> M"
-  assumes Int[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"
-  assumes Diff_cover:
-    "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> \<exists>C\<subseteq>M. finite C \<and> disjoint C \<and> a - b = \<Union>C"
-
-lemma (in semiring_of_sets) finite_INT[intro]:
-  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
-  shows "(\<Inter>i\<in>I. A i) \<in> M"
-  using assms by (induct rule: finite_ne_induct) auto
-
-lemma (in semiring_of_sets) Int_space_eq1 [simp]: "x \<in> M \<Longrightarrow> \<Omega> \<inter> x = x"
-  by (metis Int_absorb1 sets_into_space)
-
-lemma (in semiring_of_sets) Int_space_eq2 [simp]: "x \<in> M \<Longrightarrow> x \<inter> \<Omega> = x"
-  by (metis Int_absorb2 sets_into_space)
-
-lemma (in semiring_of_sets) sets_Collect_conj:
-  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
-  shows "{x\<in>\<Omega>. Q x \<and> P x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. Q x \<and> P x} = {x\<in>\<Omega>. Q x} \<inter> {x\<in>\<Omega>. P x}"
-    by auto
-  with assms show ?thesis by auto
-qed
-
-lemma (in semiring_of_sets) sets_Collect_finite_All':
-  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S" "S \<noteq> {}"
-  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (\<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
-    using \<open>S \<noteq> {}\<close> by auto
-  with assms show ?thesis by auto
-qed
-
-locale ring_of_sets = semiring_of_sets +
-  assumes Un [intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"
-
-lemma (in ring_of_sets) finite_Union [intro]:
-  "finite X \<Longrightarrow> X \<subseteq> M \<Longrightarrow> \<Union>X \<in> M"
-  by (induct set: finite) (auto simp add: Un)
-
-lemma (in ring_of_sets) finite_UN[intro]:
-  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
-  shows "(\<Union>i\<in>I. A i) \<in> M"
-  using assms by induct auto
-
-lemma (in ring_of_sets) Diff [intro]:
-  assumes "a \<in> M" "b \<in> M" shows "a - b \<in> M"
-  using Diff_cover[OF assms] by auto
-
-lemma ring_of_setsI:
-  assumes space_closed: "M \<subseteq> Pow \<Omega>"
-  assumes empty_sets[iff]: "{} \<in> M"
-  assumes Un[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"
-  assumes Diff[intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a - b \<in> M"
-  shows "ring_of_sets \<Omega> M"
-proof
-  fix a b assume ab: "a \<in> M" "b \<in> M"
-  from ab show "\<exists>C\<subseteq>M. finite C \<and> disjoint C \<and> a - b = \<Union>C"
-    by (intro exI[of _ "{a - b}"]) (auto simp: disjoint_def)
-  have "a \<inter> b = a - (a - b)" by auto
-  also have "\<dots> \<in> M" using ab by auto
-  finally show "a \<inter> b \<in> M" .
-qed fact+
-
-lemma ring_of_sets_iff: "ring_of_sets \<Omega> M \<longleftrightarrow> M \<subseteq> Pow \<Omega> \<and> {} \<in> M \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a \<union> b \<in> M) \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a - b \<in> M)"
-proof
-  assume "ring_of_sets \<Omega> M"
-  then interpret ring_of_sets \<Omega> M .
-  show "M \<subseteq> Pow \<Omega> \<and> {} \<in> M \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a \<union> b \<in> M) \<and> (\<forall>a\<in>M. \<forall>b\<in>M. a - b \<in> M)"
-    using space_closed by auto
-qed (auto intro!: ring_of_setsI)
-
-lemma (in ring_of_sets) insert_in_sets:
-  assumes "{x} \<in> M" "A \<in> M" shows "insert x A \<in> M"
-proof -
-  have "{x} \<union> A \<in> M" using assms by (rule Un)
-  thus ?thesis by auto
-qed
-
-lemma (in ring_of_sets) sets_Collect_disj:
-  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
-  shows "{x\<in>\<Omega>. Q x \<or> P x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. Q x \<or> P x} = {x\<in>\<Omega>. Q x} \<union> {x\<in>\<Omega>. P x}"
-    by auto
-  with assms show ?thesis by auto
-qed
-
-lemma (in ring_of_sets) sets_Collect_finite_Ex:
-  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
-  shows "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} = (\<Union>i\<in>S. {x\<in>\<Omega>. P i x})"
-    by auto
-  with assms show ?thesis by auto
-qed
-
-locale algebra = ring_of_sets +
-  assumes top [iff]: "\<Omega> \<in> M"
-
-lemma (in algebra) compl_sets [intro]:
-  "a \<in> M \<Longrightarrow> \<Omega> - a \<in> M"
-  by auto
-
-lemma algebra_iff_Un:
-  "algebra \<Omega> M \<longleftrightarrow>
-    M \<subseteq> Pow \<Omega> \<and>
-    {} \<in> M \<and>
-    (\<forall>a \<in> M. \<Omega> - a \<in> M) \<and>
-    (\<forall>a \<in> M. \<forall> b \<in> M. a \<union> b \<in> M)" (is "_ \<longleftrightarrow> ?Un")
-proof
-  assume "algebra \<Omega> M"
-  then interpret algebra \<Omega> M .
-  show ?Un using sets_into_space by auto
-next
-  assume ?Un
-  then have "\<Omega> \<in> M" by auto
-  interpret ring_of_sets \<Omega> M
-  proof (rule ring_of_setsI)
-    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
-      using \<open>?Un\<close> by auto
-    fix a b assume a: "a \<in> M" and b: "b \<in> M"
-    then show "a \<union> b \<in> M" using \<open>?Un\<close> by auto
-    have "a - b = \<Omega> - ((\<Omega> - a) \<union> b)"
-      using \<Omega> a b by auto
-    then show "a - b \<in> M"
-      using a b  \<open>?Un\<close> by auto
-  qed
-  show "algebra \<Omega> M" proof qed fact
-qed
-
-lemma algebra_iff_Int:
-     "algebra \<Omega> M \<longleftrightarrow>
-       M \<subseteq> Pow \<Omega> & {} \<in> M &
-       (\<forall>a \<in> M. \<Omega> - a \<in> M) &
-       (\<forall>a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)" (is "_ \<longleftrightarrow> ?Int")
-proof
-  assume "algebra \<Omega> M"
-  then interpret algebra \<Omega> M .
-  show ?Int using sets_into_space by auto
-next
-  assume ?Int
-  show "algebra \<Omega> M"
-  proof (unfold algebra_iff_Un, intro conjI ballI)
-    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
-      using \<open>?Int\<close> by auto
-    from \<open>?Int\<close> show "\<And>a. a \<in> M \<Longrightarrow> \<Omega> - a \<in> M" by auto
-    fix a b assume M: "a \<in> M" "b \<in> M"
-    hence "a \<union> b = \<Omega> - ((\<Omega> - a) \<inter> (\<Omega> - b))"
-      using \<Omega> by blast
-    also have "... \<in> M"
-      using M \<open>?Int\<close> by auto
-    finally show "a \<union> b \<in> M" .
-  qed
-qed
-
-lemma (in algebra) sets_Collect_neg:
-  assumes "{x\<in>\<Omega>. P x} \<in> M"
-  shows "{x\<in>\<Omega>. \<not> P x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<not> P x} = \<Omega> - {x\<in>\<Omega>. P x}" by auto
-  with assms show ?thesis by auto
-qed
-
-lemma (in algebra) sets_Collect_imp:
-  "{x\<in>\<Omega>. P x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x \<longrightarrow> P x} \<in> M"
-  unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg)
-
-lemma (in algebra) sets_Collect_const:
-  "{x\<in>\<Omega>. P} \<in> M"
-  by (cases P) auto
-
-lemma algebra_single_set:
-  "X \<subseteq> S \<Longrightarrow> algebra S { {}, X, S - X, S }"
-  by (auto simp: algebra_iff_Int)
-
-subsubsection \<open>Restricted algebras\<close>
-
-abbreviation (in algebra)
-  "restricted_space A \<equiv> (op \<inter> A) ` M"
-
-lemma (in algebra) restricted_algebra:
-  assumes "A \<in> M" shows "algebra A (restricted_space A)"
-  using assms by (auto simp: algebra_iff_Int)
-
-subsubsection \<open>Sigma Algebras\<close>
-
-locale sigma_algebra = algebra +
-  assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
-
-lemma (in algebra) is_sigma_algebra:
-  assumes "finite M"
-  shows "sigma_algebra \<Omega> M"
-proof
-  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> M"
-  then have "(\<Union>i. A i) = (\<Union>s\<in>M \<inter> range A. s)"
-    by auto
-  also have "(\<Union>s\<in>M \<inter> range A. s) \<in> M"
-    using \<open>finite M\<close> by auto
-  finally show "(\<Union>i. A i) \<in> M" .
-qed
-
-lemma countable_UN_eq:
-  fixes A :: "'i::countable \<Rightarrow> 'a set"
-  shows "(range A \<subseteq> M \<longrightarrow> (\<Union>i. A i) \<in> M) \<longleftrightarrow>
-    (range (A \<circ> from_nat) \<subseteq> M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> M)"
-proof -
-  let ?A' = "A \<circ> from_nat"
-  have *: "(\<Union>i. ?A' i) = (\<Union>i. A i)" (is "?l = ?r")
-  proof safe
-    fix x i assume "x \<in> A i" thus "x \<in> ?l"
-      by (auto intro!: exI[of _ "to_nat i"])
-  next
-    fix x i assume "x \<in> ?A' i" thus "x \<in> ?r"
-      by (auto intro!: exI[of _ "from_nat i"])
-  qed
-  have **: "range ?A' = range A"
-    using surj_from_nat
-    by (auto simp: image_comp [symmetric] intro!: imageI)
-  show ?thesis unfolding * ** ..
-qed
-
-lemma (in sigma_algebra) countable_Union [intro]:
-  assumes "countable X" "X \<subseteq> M" shows "\<Union>X \<in> M"
-proof cases
-  assume "X \<noteq> {}"
-  hence "\<Union>X = (\<Union>n. from_nat_into X n)"
-    using assms by (auto intro: from_nat_into) (metis from_nat_into_surj)
-  also have "\<dots> \<in> M" using assms
-    by (auto intro!: countable_nat_UN) (metis \<open>X \<noteq> {}\<close> from_nat_into set_mp)
-  finally show ?thesis .
-qed simp
-
-lemma (in sigma_algebra) countable_UN[intro]:
-  fixes A :: "'i::countable \<Rightarrow> 'a set"
-  assumes "A`X \<subseteq> M"
-  shows  "(\<Union>x\<in>X. A x) \<in> M"
-proof -
-  let ?A = "\<lambda>i. if i \<in> X then A i else {}"
-  from assms have "range ?A \<subseteq> M" by auto
-  with countable_nat_UN[of "?A \<circ> from_nat"] countable_UN_eq[of ?A M]
-  have "(\<Union>x. ?A x) \<in> M" by auto
-  moreover have "(\<Union>x. ?A x) = (\<Union>x\<in>X. A x)" by (auto split: if_split_asm)
-  ultimately show ?thesis by simp
-qed
-
-lemma (in sigma_algebra) countable_UN':
-  fixes A :: "'i \<Rightarrow> 'a set"
-  assumes X: "countable X"
-  assumes A: "A`X \<subseteq> M"
-  shows  "(\<Union>x\<in>X. A x) \<in> M"
-proof -
-  have "(\<Union>x\<in>X. A x) = (\<Union>i\<in>to_nat_on X ` X. A (from_nat_into X i))"
-    using X by auto
-  also have "\<dots> \<in> M"
-    using A X
-    by (intro countable_UN) auto
-  finally show ?thesis .
-qed
-
-lemma (in sigma_algebra) countable_UN'':
-  "\<lbrakk> countable X; \<And>x y. x \<in> X \<Longrightarrow> A x \<in> M \<rbrakk> \<Longrightarrow> (\<Union>x\<in>X. A x) \<in> M"
-by(erule countable_UN')(auto)
-
-lemma (in sigma_algebra) countable_INT [intro]:
-  fixes A :: "'i::countable \<Rightarrow> 'a set"
-  assumes A: "A`X \<subseteq> M" "X \<noteq> {}"
-  shows "(\<Inter>i\<in>X. A i) \<in> M"
-proof -
-  from A have "\<forall>i\<in>X. A i \<in> M" by fast
-  hence "\<Omega> - (\<Union>i\<in>X. \<Omega> - A i) \<in> M" by blast
-  moreover
-  have "(\<Inter>i\<in>X. A i) = \<Omega> - (\<Union>i\<in>X. \<Omega> - A i)" using space_closed A
-    by blast
-  ultimately show ?thesis by metis
-qed
-
-lemma (in sigma_algebra) countable_INT':
-  fixes A :: "'i \<Rightarrow> 'a set"
-  assumes X: "countable X" "X \<noteq> {}"
-  assumes A: "A`X \<subseteq> M"
-  shows  "(\<Inter>x\<in>X. A x) \<in> M"
-proof -
-  have "(\<Inter>x\<in>X. A x) = (\<Inter>i\<in>to_nat_on X ` X. A (from_nat_into X i))"
-    using X by auto
-  also have "\<dots> \<in> M"
-    using A X
-    by (intro countable_INT) auto
-  finally show ?thesis .
-qed
-
-lemma (in sigma_algebra) countable_INT'':
-  "UNIV \<in> M \<Longrightarrow> countable I \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i \<in> M) \<Longrightarrow> (\<Inter>i\<in>I. F i) \<in> M"
-  by (cases "I = {}") (auto intro: countable_INT')
-
-lemma (in sigma_algebra) countable:
-  assumes "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> M" "countable A"
-  shows "A \<in> M"
-proof -
-  have "(\<Union>a\<in>A. {a}) \<in> M"
-    using assms by (intro countable_UN') auto
-  also have "(\<Union>a\<in>A. {a}) = A" by auto
-  finally show ?thesis by auto
-qed
-
-lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)"
-  by (auto simp: ring_of_sets_iff)
-
-lemma algebra_Pow: "algebra sp (Pow sp)"
-  by (auto simp: algebra_iff_Un)
-
-lemma sigma_algebra_iff:
-  "sigma_algebra \<Omega> M \<longleftrightarrow>
-    algebra \<Omega> M \<and> (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
-  by (simp add: sigma_algebra_def sigma_algebra_axioms_def)
-
-lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)"
-  by (auto simp: sigma_algebra_iff algebra_iff_Int)
-
-lemma (in sigma_algebra) sets_Collect_countable_All:
-  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
-  shows "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} = (\<Inter>i. {x\<in>\<Omega>. P i x})" by auto
-  with assms show ?thesis by auto
-qed
-
-lemma (in sigma_algebra) sets_Collect_countable_Ex:
-  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
-  shows "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} = (\<Union>i. {x\<in>\<Omega>. P i x})" by auto
-  with assms show ?thesis by auto
-qed
-
-lemma (in sigma_algebra) sets_Collect_countable_Ex':
-  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
-  assumes "countable I"
-  shows "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<exists>i\<in>I. P i x} = (\<Union>i\<in>I. {x\<in>\<Omega>. P i x})" by auto
-  with assms show ?thesis
-    by (auto intro!: countable_UN')
-qed
-
-lemma (in sigma_algebra) sets_Collect_countable_All':
-  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
-  assumes "countable I"
-  shows "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<forall>i\<in>I. P i x} = (\<Inter>i\<in>I. {x\<in>\<Omega>. P i x}) \<inter> \<Omega>" by auto
-  with assms show ?thesis
-    by (cases "I = {}") (auto intro!: countable_INT')
-qed
-
-lemma (in sigma_algebra) sets_Collect_countable_Ex1':
-  assumes "\<And>i. i \<in> I \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M"
-  assumes "countable I"
-  shows "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} \<in> M"
-proof -
-  have "{x\<in>\<Omega>. \<exists>!i\<in>I. P i x} = {x\<in>\<Omega>. \<exists>i\<in>I. P i x \<and> (\<forall>j\<in>I. P j x \<longrightarrow> i = j)}"
-    by auto
-  with assms show ?thesis
-    by (auto intro!: sets_Collect_countable_All' sets_Collect_countable_Ex' sets_Collect_conj sets_Collect_imp sets_Collect_const)
-qed
-
-lemmas (in sigma_algebra) sets_Collect =
-  sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const
-  sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All
-
-lemma (in sigma_algebra) sets_Collect_countable_Ball:
-  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
-  shows "{x\<in>\<Omega>. \<forall>i::'i::countable\<in>X. P i x} \<in> M"
-  unfolding Ball_def by (intro sets_Collect assms)
-
-lemma (in sigma_algebra) sets_Collect_countable_Bex:
-  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
-  shows "{x\<in>\<Omega>. \<exists>i::'i::countable\<in>X. P i x} \<in> M"
-  unfolding Bex_def by (intro sets_Collect assms)
-
-lemma sigma_algebra_single_set:
-  assumes "X \<subseteq> S"
-  shows "sigma_algebra S { {}, X, S - X, S }"
-  using algebra.is_sigma_algebra[OF algebra_single_set[OF \<open>X \<subseteq> S\<close>]] by simp
-
-subsubsection \<open>Binary Unions\<close>
-
-definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
-  where "binary a b =  (\<lambda>x. b)(0 := a)"
-
-lemma range_binary_eq: "range(binary a b) = {a,b}"
-  by (auto simp add: binary_def)
-
-lemma Un_range_binary: "a \<union> b = (\<Union>i::nat. binary a b i)"
-  by (simp add: range_binary_eq cong del: strong_SUP_cong)
-
-lemma Int_range_binary: "a \<inter> b = (\<Inter>i::nat. binary a b i)"
-  by (simp add: range_binary_eq cong del: strong_INF_cong)
-
-lemma sigma_algebra_iff2:
-     "sigma_algebra \<Omega> M \<longleftrightarrow>
-       M \<subseteq> Pow \<Omega> \<and>
-       {} \<in> M \<and> (\<forall>s \<in> M. \<Omega> - s \<in> M) \<and>
-       (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
-  by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
-         algebra_iff_Un Un_range_binary)
-
-subsubsection \<open>Initial Sigma Algebra\<close>
-
-text \<open>Sigma algebras can naturally be created as the closure of any set of
-  M with regard to the properties just postulated.\<close>
-
-inductive_set sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
-  for sp :: "'a set" and A :: "'a set set"
-  where
-    Basic[intro, simp]: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"
-  | Empty: "{} \<in> sigma_sets sp A"
-  | Compl: "a \<in> sigma_sets sp A \<Longrightarrow> sp - a \<in> sigma_sets sp A"
-  | Union: "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Union>i. a i) \<in> sigma_sets sp A"
-
-lemma (in sigma_algebra) sigma_sets_subset:
-  assumes a: "a \<subseteq> M"
-  shows "sigma_sets \<Omega> a \<subseteq> M"
-proof
-  fix x
-  assume "x \<in> sigma_sets \<Omega> a"
-  from this show "x \<in> M"
-    by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
-qed
-
-lemma sigma_sets_into_sp: "A \<subseteq> Pow sp \<Longrightarrow> x \<in> sigma_sets sp A \<Longrightarrow> x \<subseteq> sp"
-  by (erule sigma_sets.induct, auto)
-
-lemma sigma_algebra_sigma_sets:
-     "a \<subseteq> Pow \<Omega> \<Longrightarrow> sigma_algebra \<Omega> (sigma_sets \<Omega> a)"
-  by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp
-           intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl)
-
-lemma sigma_sets_least_sigma_algebra:
-  assumes "A \<subseteq> Pow S"
-  shows "sigma_sets S A = \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
-proof safe
-  fix B X assume "A \<subseteq> B" and sa: "sigma_algebra S B"
-    and X: "X \<in> sigma_sets S A"
-  from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF \<open>A \<subseteq> B\<close>] X
-  show "X \<in> B" by auto
-next
-  fix X assume "X \<in> \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
-  then have [intro!]: "\<And>B. A \<subseteq> B \<Longrightarrow> sigma_algebra S B \<Longrightarrow> X \<in> B"
-     by simp
-  have "A \<subseteq> sigma_sets S A" using assms by auto
-  moreover have "sigma_algebra S (sigma_sets S A)"
-    using assms by (intro sigma_algebra_sigma_sets[of A]) auto
-  ultimately show "X \<in> sigma_sets S A" by auto
-qed
-
-lemma sigma_sets_top: "sp \<in> sigma_sets sp A"
-  by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty)
-
-lemma sigma_sets_Un:
-  "a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"
-apply (simp add: Un_range_binary range_binary_eq)
-apply (rule Union, simp add: binary_def)
-done
-
-lemma sigma_sets_Inter:
-  assumes Asb: "A \<subseteq> Pow sp"
-  shows "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Inter>i. a i) \<in> sigma_sets sp A"
-proof -
-  assume ai: "\<And>i::nat. a i \<in> sigma_sets sp A"
-  hence "\<And>i::nat. sp-(a i) \<in> sigma_sets sp A"
-    by (rule sigma_sets.Compl)
-  hence "(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
-    by (rule sigma_sets.Union)
-  hence "sp-(\<Union>i. sp-(a i)) \<in> sigma_sets sp A"
-    by (rule sigma_sets.Compl)
-  also have "sp-(\<Union>i. sp-(a i)) = sp Int (\<Inter>i. a i)"
-    by auto
-  also have "... = (\<Inter>i. a i)" using ai
-    by (blast dest: sigma_sets_into_sp [OF Asb])
-  finally show ?thesis .
-qed
-
-lemma sigma_sets_INTER:
-  assumes Asb: "A \<subseteq> Pow sp"
-      and ai: "\<And>i::nat. i \<in> S \<Longrightarrow> a i \<in> sigma_sets sp A" and non: "S \<noteq> {}"
-  shows "(\<Inter>i\<in>S. a i) \<in> sigma_sets sp A"
-proof -
-  from ai have "\<And>i. (if i\<in>S then a i else sp) \<in> sigma_sets sp A"
-    by (simp add: sigma_sets.intros(2-) sigma_sets_top)
-  hence "(\<Inter>i. (if i\<in>S then a i else sp)) \<in> sigma_sets sp A"
-    by (rule sigma_sets_Inter [OF Asb])
-  also have "(\<Inter>i. (if i\<in>S then a i else sp)) = (\<Inter>i\<in>S. a i)"
-    by auto (metis ai non sigma_sets_into_sp subset_empty subset_iff Asb)+
-  finally show ?thesis .
-qed
-
-lemma sigma_sets_UNION:
-  "countable B \<Longrightarrow> (\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets X A) \<Longrightarrow> (\<Union>B) \<in> sigma_sets X A"
-  apply (cases "B = {}")
-  apply (simp add: sigma_sets.Empty)
-  using from_nat_into [of B] range_from_nat_into [of B] sigma_sets.Union [of "from_nat_into B" X A]
-  apply simp
-  apply auto
-  apply (metis Sup_bot_conv(1) Union_empty \<open>\<lbrakk>B \<noteq> {}; countable B\<rbrakk> \<Longrightarrow> range (from_nat_into B) = B\<close>)
-  done
-
-lemma (in sigma_algebra) sigma_sets_eq:
-     "sigma_sets \<Omega> M = M"
-proof
-  show "M \<subseteq> sigma_sets \<Omega> M"
-    by (metis Set.subsetI sigma_sets.Basic)
-  next
-  show "sigma_sets \<Omega> M \<subseteq> M"
-    by (metis sigma_sets_subset subset_refl)
-qed
-
-lemma sigma_sets_eqI:
-  assumes A: "\<And>a. a \<in> A \<Longrightarrow> a \<in> sigma_sets M B"
-  assumes B: "\<And>b. b \<in> B \<Longrightarrow> b \<in> sigma_sets M A"
-  shows "sigma_sets M A = sigma_sets M B"
-proof (intro set_eqI iffI)
-  fix a assume "a \<in> sigma_sets M A"
-  from this A show "a \<in> sigma_sets M B"
-    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
-next
-  fix b assume "b \<in> sigma_sets M B"
-  from this B show "b \<in> sigma_sets M A"
-    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
-qed
-
-lemma sigma_sets_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
-proof
-  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert \<open>A \<subseteq> B\<close>, auto intro: sigma_sets.intros(2-))
-qed
-
-lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
-proof
-  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert \<open>A \<subseteq> sigma_sets X B\<close>, auto intro: sigma_sets.intros(2-))
-qed
-
-lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
-proof
-  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert \<open>A \<subseteq> B\<close>, auto intro: sigma_sets.intros(2-))
-qed
-
-lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"
-  by (auto intro: sigma_sets.Basic)
-
-lemma (in sigma_algebra) restriction_in_sets:
-  fixes A :: "nat \<Rightarrow> 'a set"
-  assumes "S \<in> M"
-  and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` M" (is "_ \<subseteq> ?r")
-  shows "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
-proof -
-  { fix i have "A i \<in> ?r" using * by auto
-    hence "\<exists>B. A i = B \<inter> S \<and> B \<in> M" by auto
-    hence "A i \<subseteq> S" "A i \<in> M" using \<open>S \<in> M\<close> by auto }
-  thus "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
-    by (auto intro!: image_eqI[of _ _ "(\<Union>i. A i)"])
-qed
-
-lemma (in sigma_algebra) restricted_sigma_algebra:
-  assumes "S \<in> M"
-  shows "sigma_algebra S (restricted_space S)"
-  unfolding sigma_algebra_def sigma_algebra_axioms_def
-proof safe
-  show "algebra S (restricted_space S)" using restricted_algebra[OF assms] .
-next
-  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> restricted_space S"
-  from restriction_in_sets[OF assms this[simplified]]
-  show "(\<Union>i. A i) \<in> restricted_space S" by simp
-qed
-
-lemma sigma_sets_Int:
-  assumes "A \<in> sigma_sets sp st" "A \<subseteq> sp"
-  shows "op \<inter> A ` sigma_sets sp st = sigma_sets A (op \<inter> A ` st)"
-proof (intro equalityI subsetI)
-  fix x assume "x \<in> op \<inter> A ` sigma_sets sp st"
-  then obtain y where "y \<in> sigma_sets sp st" "x = y \<inter> A" by auto
-  then have "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
-  proof (induct arbitrary: x)
-    case (Compl a)
-    then show ?case
-      by (force intro!: sigma_sets.Compl simp: Diff_Int_distrib ac_simps)
-  next
-    case (Union a)
-    then show ?case
-      by (auto intro!: sigma_sets.Union
-               simp add: UN_extend_simps simp del: UN_simps)
-  qed (auto intro!: sigma_sets.intros(2-))
-  then show "x \<in> sigma_sets A (op \<inter> A ` st)"
-    using \<open>A \<subseteq> sp\<close> by (simp add: Int_absorb2)
-next
-  fix x assume "x \<in> sigma_sets A (op \<inter> A ` st)"
-  then show "x \<in> op \<inter> A ` sigma_sets sp st"
-  proof induct
-    case (Compl a)
-    then obtain x where "a = A \<inter> x" "x \<in> sigma_sets sp st" by auto
-    then show ?case using \<open>A \<subseteq> sp\<close>
-      by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)
-  next
-    case (Union a)
-    then have "\<forall>i. \<exists>x. x \<in> sigma_sets sp st \<and> a i = A \<inter> x"
-      by (auto simp: image_iff Bex_def)
-    from choice[OF this] guess f ..
-    then show ?case
-      by (auto intro!: bexI[of _ "(\<Union>x. f x)"] sigma_sets.Union
-               simp add: image_iff)
-  qed (auto intro!: sigma_sets.intros(2-))
-qed
-
-lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}"
-proof (intro set_eqI iffI)
-  fix a assume "a \<in> sigma_sets A {}" then show "a \<in> {{}, A}"
-    by induct blast+
-qed (auto intro: sigma_sets.Empty sigma_sets_top)
-
-lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}"
-proof (intro set_eqI iffI)
-  fix x assume "x \<in> sigma_sets A {A}"
-  then show "x \<in> {{}, A}"
-    by induct blast+
-next
-  fix x assume "x \<in> {{}, A}"
-  then show "x \<in> sigma_sets A {A}"
-    by (auto intro: sigma_sets.Empty sigma_sets_top)
-qed
-
-lemma sigma_sets_sigma_sets_eq:
-  "M \<subseteq> Pow S \<Longrightarrow> sigma_sets S (sigma_sets S M) = sigma_sets S M"
-  by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto
-
-lemma sigma_sets_singleton:
-  assumes "X \<subseteq> S"
-  shows "sigma_sets S { X } = { {}, X, S - X, S }"
-proof -
-  interpret sigma_algebra S "{ {}, X, S - X, S }"
-    by (rule sigma_algebra_single_set) fact
-  have "sigma_sets S { X } \<subseteq> sigma_sets S { {}, X, S - X, S }"
-    by (rule sigma_sets_subseteq) simp
-  moreover have "\<dots> = { {}, X, S - X, S }"
-    using sigma_sets_eq by simp
-  moreover
-  { fix A assume "A \<in> { {}, X, S - X, S }"
-    then have "A \<in> sigma_sets S { X }"
-      by (auto intro: sigma_sets.intros(2-) sigma_sets_top) }
-  ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }"
-    by (intro antisym) auto
-  with sigma_sets_eq show ?thesis by simp
-qed
-
-lemma restricted_sigma:
-  assumes S: "S \<in> sigma_sets \<Omega> M" and M: "M \<subseteq> Pow \<Omega>"
-  shows "algebra.restricted_space (sigma_sets \<Omega> M) S =
-    sigma_sets S (algebra.restricted_space M S)"
-proof -
-  from S sigma_sets_into_sp[OF M]
-  have "S \<in> sigma_sets \<Omega> M" "S \<subseteq> \<Omega>" by auto
-  from sigma_sets_Int[OF this]
-  show ?thesis by simp
-qed
-
-lemma sigma_sets_vimage_commute:
-  assumes X: "X \<in> \<Omega> \<rightarrow> \<Omega>'"
-  shows "{X -` A \<inter> \<Omega> |A. A \<in> sigma_sets \<Omega>' M'}
-       = sigma_sets \<Omega> {X -` A \<inter> \<Omega> |A. A \<in> M'}" (is "?L = ?R")
-proof
-  show "?L \<subseteq> ?R"
-  proof clarify
-    fix A assume "A \<in> sigma_sets \<Omega>' M'"
-    then show "X -` A \<inter> \<Omega> \<in> ?R"
-    proof induct
-      case Empty then show ?case
-        by (auto intro!: sigma_sets.Empty)
-    next
-      case (Compl B)
-      have [simp]: "X -` (\<Omega>' - B) \<inter> \<Omega> = \<Omega> - (X -` B \<inter> \<Omega>)"
-        by (auto simp add: funcset_mem [OF X])
-      with Compl show ?case
-        by (auto intro!: sigma_sets.Compl)
-    next
-      case (Union F)
-      then show ?case
-        by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps
-                 intro!: sigma_sets.Union)
-    qed auto
-  qed
-  show "?R \<subseteq> ?L"
-  proof clarify
-    fix A assume "A \<in> ?R"
-    then show "\<exists>B. A = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'"
-    proof induct
-      case (Basic B) then show ?case by auto
-    next
-      case Empty then show ?case
-        by (auto intro!: sigma_sets.Empty exI[of _ "{}"])
-    next
-      case (Compl B)
-      then obtain A where A: "B = X -` A \<inter> \<Omega>" "A \<in> sigma_sets \<Omega>' M'" by auto
-      then have [simp]: "\<Omega> - B = X -` (\<Omega>' - A) \<inter> \<Omega>"
-        by (auto simp add: funcset_mem [OF X])
-      with A(2) show ?case
-        by (auto intro: sigma_sets.Compl)
-    next
-      case (Union F)
-      then have "\<forall>i. \<exists>B. F i = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'" by auto
-      from choice[OF this] guess A .. note A = this
-      with A show ?case
-        by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union)
-    qed
-  qed
-qed
-
-lemma (in ring_of_sets) UNION_in_sets:
-  fixes A:: "nat \<Rightarrow> 'a set"
-  assumes A: "range A \<subseteq> M"
-  shows  "(\<Union>i\<in>{0..<n}. A i) \<in> M"
-proof (induct n)
-  case 0 show ?case by simp
-next
-  case (Suc n)
-  thus ?case
-    by (simp add: atLeastLessThanSuc) (metis A Un UNIV_I image_subset_iff)
-qed
-
-lemma (in ring_of_sets) range_disjointed_sets:
-  assumes A: "range A \<subseteq> M"
-  shows  "range (disjointed A) \<subseteq> M"
-proof (auto simp add: disjointed_def)
-  fix n
-  show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> M" using UNION_in_sets
-    by (metis A Diff UNIV_I image_subset_iff)
-qed
-
-lemma (in algebra) range_disjointed_sets':
-  "range A \<subseteq> M \<Longrightarrow> range (disjointed A) \<subseteq> M"
-  using range_disjointed_sets .
-
-lemma sigma_algebra_disjoint_iff:
-  "sigma_algebra \<Omega> M \<longleftrightarrow> algebra \<Omega> M \<and>
-    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
-proof (auto simp add: sigma_algebra_iff)
-  fix A :: "nat \<Rightarrow> 'a set"
-  assume M: "algebra \<Omega> M"
-     and A: "range A \<subseteq> M"
-     and UnA: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M"
-  hence "range (disjointed A) \<subseteq> M \<longrightarrow>
-         disjoint_family (disjointed A) \<longrightarrow>
-         (\<Union>i. disjointed A i) \<in> M" by blast
-  hence "(\<Union>i. disjointed A i) \<in> M"
-    by (simp add: algebra.range_disjointed_sets'[of \<Omega>] M A disjoint_family_disjointed)
-  thus "(\<Union>i::nat. A i) \<in> M" by (simp add: UN_disjointed_eq)
-qed
-
-subsubsection \<open>Ring generated by a semiring\<close>
-
-definition (in semiring_of_sets)
-  "generated_ring = { \<Union>C | C. C \<subseteq> M \<and> finite C \<and> disjoint C }"
-
-lemma (in semiring_of_sets) generated_ringE[elim?]:
-  assumes "a \<in> generated_ring"
-  obtains C where "finite C" "disjoint C" "C \<subseteq> M" "a = \<Union>C"
-  using assms unfolding generated_ring_def by auto
-
-lemma (in semiring_of_sets) generated_ringI[intro?]:
-  assumes "finite C" "disjoint C" "C \<subseteq> M" "a = \<Union>C"
-  shows "a \<in> generated_ring"
-  using assms unfolding generated_ring_def by auto
-
-lemma (in semiring_of_sets) generated_ringI_Basic:
-  "A \<in> M \<Longrightarrow> A \<in> generated_ring"
-  by (rule generated_ringI[of "{A}"]) (auto simp: disjoint_def)
-
-lemma (in semiring_of_sets) generated_ring_disjoint_Un[intro]:
-  assumes a: "a \<in> generated_ring" and b: "b \<in> generated_ring"
-  and "a \<inter> b = {}"
-  shows "a \<union> b \<in> generated_ring"
-proof -
-  from a guess Ca .. note Ca = this
-  from b guess Cb .. note Cb = this
-  show ?thesis
-  proof
-    show "disjoint (Ca \<union> Cb)"
-      using \<open>a \<inter> b = {}\<close> Ca Cb by (auto intro!: disjoint_union)
-  qed (insert Ca Cb, auto)
-qed
-
-lemma (in semiring_of_sets) generated_ring_empty: "{} \<in> generated_ring"
-  by (auto simp: generated_ring_def disjoint_def)
-
-lemma (in semiring_of_sets) generated_ring_disjoint_Union:
-  assumes "finite A" shows "A \<subseteq> generated_ring \<Longrightarrow> disjoint A \<Longrightarrow> \<Union>A \<in> generated_ring"
-  using assms by (induct A) (auto simp: disjoint_def intro!: generated_ring_disjoint_Un generated_ring_empty)
-
-lemma (in semiring_of_sets) generated_ring_disjoint_UNION:
-  "finite I \<Longrightarrow> disjoint (A ` I) \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> UNION I A \<in> generated_ring"
-  by (intro generated_ring_disjoint_Union) auto
-
-lemma (in semiring_of_sets) generated_ring_Int:
-  assumes a: "a \<in> generated_ring" and b: "b \<in> generated_ring"
-  shows "a \<inter> b \<in> generated_ring"
-proof -
-  from a guess Ca .. note Ca = this
-  from b guess Cb .. note Cb = this
-  define C where "C = (\<lambda>(a,b). a \<inter> b)` (Ca\<times>Cb)"
-  show ?thesis
-  proof
-    show "disjoint C"
-    proof (simp add: disjoint_def C_def, intro ballI impI)
-      fix a1 b1 a2 b2 assume sets: "a1 \<in> Ca" "b1 \<in> Cb" "a2 \<in> Ca" "b2 \<in> Cb"
-      assume "a1 \<inter> b1 \<noteq> a2 \<inter> b2"
-      then have "a1 \<noteq> a2 \<or> b1 \<noteq> b2" by auto
-      then show "(a1 \<inter> b1) \<inter> (a2 \<inter> b2) = {}"
-      proof
-        assume "a1 \<noteq> a2"
-        with sets Ca have "a1 \<inter> a2 = {}"
-          by (auto simp: disjoint_def)
-        then show ?thesis by auto
-      next
-        assume "b1 \<noteq> b2"
-        with sets Cb have "b1 \<inter> b2 = {}"
-          by (auto simp: disjoint_def)
-        then show ?thesis by auto
-      qed
-    qed
-  qed (insert Ca Cb, auto simp: C_def)
-qed
-
-lemma (in semiring_of_sets) generated_ring_Inter:
-  assumes "finite A" "A \<noteq> {}" shows "A \<subseteq> generated_ring \<Longrightarrow> \<Inter>A \<in> generated_ring"
-  using assms by (induct A rule: finite_ne_induct) (auto intro: generated_ring_Int)
-
-lemma (in semiring_of_sets) generated_ring_INTER:
-  "finite I \<Longrightarrow> I \<noteq> {} \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> A i \<in> generated_ring) \<Longrightarrow> INTER I A \<in> generated_ring"
-  by (intro generated_ring_Inter) auto
-
-lemma (in semiring_of_sets) generating_ring:
-  "ring_of_sets \<Omega> generated_ring"
-proof (rule ring_of_setsI)
-  let ?R = generated_ring
-  show "?R \<subseteq> Pow \<Omega>"
-    using sets_into_space by (auto simp: generated_ring_def generated_ring_empty)
-  show "{} \<in> ?R" by (rule generated_ring_empty)
-
-  { fix a assume a: "a \<in> ?R" then guess Ca .. note Ca = this
-    fix b assume b: "b \<in> ?R" then guess Cb .. note Cb = this
-
-    show "a - b \<in> ?R"
-    proof cases
-      assume "Cb = {}" with Cb \<open>a \<in> ?R\<close> show ?thesis
-        by simp
-    next
-      assume "Cb \<noteq> {}"
-      with Ca Cb have "a - b = (\<Union>a'\<in>Ca. \<Inter>b'\<in>Cb. a' - b')" by auto
-      also have "\<dots> \<in> ?R"
-      proof (intro generated_ring_INTER generated_ring_disjoint_UNION)
-        fix a b assume "a \<in> Ca" "b \<in> Cb"
-        with Ca Cb Diff_cover[of a b] show "a - b \<in> ?R"
-          by (auto simp add: generated_ring_def)
-            (metis DiffI Diff_eq_empty_iff empty_iff)
-      next
-        show "disjoint ((\<lambda>a'. \<Inter>b'\<in>Cb. a' - b')`Ca)"
-          using Ca by (auto simp add: disjoint_def \<open>Cb \<noteq> {}\<close>)
-      next
-        show "finite Ca" "finite Cb" "Cb \<noteq> {}" by fact+
-      qed
-      finally show "a - b \<in> ?R" .
-    qed }
-  note Diff = this
-
-  fix a b assume sets: "a \<in> ?R" "b \<in> ?R"
-  have "a \<union> b = (a - b) \<union> (a \<inter> b) \<union> (b - a)" by auto
-  also have "\<dots> \<in> ?R"
-    by (intro sets generated_ring_disjoint_Un generated_ring_Int Diff) auto
-  finally show "a \<union> b \<in> ?R" .
-qed
-
-lemma (in semiring_of_sets) sigma_sets_generated_ring_eq: "sigma_sets \<Omega> generated_ring = sigma_sets \<Omega> M"
-proof
-  interpret M: sigma_algebra \<Omega> "sigma_sets \<Omega> M"
-    using space_closed by (rule sigma_algebra_sigma_sets)
-  show "sigma_sets \<Omega> generated_ring \<subseteq> sigma_sets \<Omega> M"
-    by (blast intro!: sigma_sets_mono elim: generated_ringE)
-qed (auto intro!: generated_ringI_Basic sigma_sets_mono)
-
-subsubsection \<open>A Two-Element Series\<close>
-
-definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set"
-  where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
-
-lemma range_binaryset_eq: "range(binaryset A B) = {A,B,{}}"
-  apply (simp add: binaryset_def)
-  apply (rule set_eqI)
-  apply (auto simp add: image_iff)
-  done
-
-lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
-  by (simp add: range_binaryset_eq cong del: strong_SUP_cong)
-
-subsubsection \<open>Closed CDI\<close>
-
-definition closed_cdi where
-  "closed_cdi \<Omega> M \<longleftrightarrow>
-   M \<subseteq> Pow \<Omega> &
-   (\<forall>s \<in> M. \<Omega> - s \<in> M) &
-   (\<forall>A. (range A \<subseteq> M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
-        (\<Union>i. A i) \<in> M) &
-   (\<forall>A. (range A \<subseteq> M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
-
-inductive_set
-  smallest_ccdi_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
-  for \<Omega> M
-  where
-    Basic [intro]:
-      "a \<in> M \<Longrightarrow> a \<in> smallest_ccdi_sets \<Omega> M"
-  | Compl [intro]:
-      "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> \<Omega> - a \<in> smallest_ccdi_sets \<Omega> M"
-  | Inc:
-      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
-       \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets \<Omega> M"
-  | Disj:
-      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> disjoint_family A
-       \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets \<Omega> M"
-
-lemma (in subset_class) smallest_closed_cdi1: "M \<subseteq> smallest_ccdi_sets \<Omega> M"
-  by auto
-
-lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets \<Omega> M \<subseteq> Pow \<Omega>"
-  apply (rule subsetI)
-  apply (erule smallest_ccdi_sets.induct)
-  apply (auto intro: range_subsetD dest: sets_into_space)
-  done
-
-lemma (in subset_class) smallest_closed_cdi2: "closed_cdi \<Omega> (smallest_ccdi_sets \<Omega> M)"
-  apply (auto simp add: closed_cdi_def smallest_ccdi_sets)
-  apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
-  done
-
-lemma closed_cdi_subset: "closed_cdi \<Omega> M \<Longrightarrow> M \<subseteq> Pow \<Omega>"
-  by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Compl: "closed_cdi \<Omega> M \<Longrightarrow> s \<in> M \<Longrightarrow> \<Omega> - s \<in> M"
-  by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Inc:
-  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow> (\<Union>i. A i) \<in> M"
-  by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Disj:
-  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
-  by (simp add: closed_cdi_def)
-
-lemma closed_cdi_Un:
-  assumes cdi: "closed_cdi \<Omega> M" and empty: "{} \<in> M"
-      and A: "A \<in> M" and B: "B \<in> M"
-      and disj: "A \<inter> B = {}"
-    shows "A \<union> B \<in> M"
-proof -
-  have ra: "range (binaryset A B) \<subseteq> M"
-   by (simp add: range_binaryset_eq empty A B)
- have di:  "disjoint_family (binaryset A B)" using disj
-   by (simp add: disjoint_family_on_def binaryset_def Int_commute)
- from closed_cdi_Disj [OF cdi ra di]
- show ?thesis
-   by (simp add: UN_binaryset_eq)
-qed
-
-lemma (in algebra) smallest_ccdi_sets_Un:
-  assumes A: "A \<in> smallest_ccdi_sets \<Omega> M" and B: "B \<in> smallest_ccdi_sets \<Omega> M"
-      and disj: "A \<inter> B = {}"
-    shows "A \<union> B \<in> smallest_ccdi_sets \<Omega> M"
-proof -
-  have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets \<Omega> M)"
-    by (simp add: range_binaryset_eq  A B smallest_ccdi_sets.Basic)
-  have di:  "disjoint_family (binaryset A B)" using disj
-    by (simp add: disjoint_family_on_def binaryset_def Int_commute)
-  from Disj [OF ra di]
-  show ?thesis
-    by (simp add: UN_binaryset_eq)
-qed
-
-lemma (in algebra) smallest_ccdi_sets_Int1:
-  assumes a: "a \<in> M"
-  shows "b \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"
-proof (induct rule: smallest_ccdi_sets.induct)
-  case (Basic x)
-  thus ?case
-    by (metis a Int smallest_ccdi_sets.Basic)
-next
-  case (Compl x)
-  have "a \<inter> (\<Omega> - x) = \<Omega> - ((\<Omega> - a) \<union> (a \<inter> x))"
-    by blast
-  also have "... \<in> smallest_ccdi_sets \<Omega> M"
-    by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
-           Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un
-           smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl)
-  finally show ?case .
-next
-  case (Inc A)
-  have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
-    by blast
-  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc
-    by blast
-  moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
-    by (simp add: Inc)
-  moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc
-    by blast
-  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"
-    by (rule smallest_ccdi_sets.Inc)
-  show ?case
-    by (metis 1 2)
-next
-  case (Disj A)
-  have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
-    by blast
-  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj
-    by blast
-  moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj
-    by (auto simp add: disjoint_family_on_def)
-  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"
-    by (rule smallest_ccdi_sets.Disj)
-  show ?case
-    by (metis 1 2)
-qed
-
-
-lemma (in algebra) smallest_ccdi_sets_Int:
-  assumes b: "b \<in> smallest_ccdi_sets \<Omega> M"
-  shows "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"
-proof (induct rule: smallest_ccdi_sets.induct)
-  case (Basic x)
-  thus ?case
-    by (metis b smallest_ccdi_sets_Int1)
-next
-  case (Compl x)
-  have "(\<Omega> - x) \<inter> b = \<Omega> - (x \<inter> b \<union> (\<Omega> - b))"
-    by blast
-  also have "... \<in> smallest_ccdi_sets \<Omega> M"
-    by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
-           smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
-  finally show ?case .
-next
-  case (Inc A)
-  have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
-    by blast
-  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc
-    by blast
-  moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
-    by (simp add: Inc)
-  moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc
-    by blast
-  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"
-    by (rule smallest_ccdi_sets.Inc)
-  show ?case
-    by (metis 1 2)
-next
-  case (Disj A)
-  have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
-    by blast
-  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj
-    by blast
-  moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj
-    by (auto simp add: disjoint_family_on_def)
-  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"
-    by (rule smallest_ccdi_sets.Disj)
-  show ?case
-    by (metis 1 2)
-qed
-
-lemma (in algebra) sigma_property_disjoint_lemma:
-  assumes sbC: "M \<subseteq> C"
-      and ccdi: "closed_cdi \<Omega> C"
-  shows "sigma_sets \<Omega> M \<subseteq> C"
-proof -
-  have "smallest_ccdi_sets \<Omega> M \<in> {B . M \<subseteq> B \<and> sigma_algebra \<Omega> B}"
-    apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
-            smallest_ccdi_sets_Int)
-    apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
-    apply (blast intro: smallest_ccdi_sets.Disj)
-    done
-  hence "sigma_sets (\<Omega>) (M) \<subseteq> smallest_ccdi_sets \<Omega> M"
-    by clarsimp
-       (drule sigma_algebra.sigma_sets_subset [where a="M"], auto)
-  also have "...  \<subseteq> C"
-    proof
-      fix x
-      assume x: "x \<in> smallest_ccdi_sets \<Omega> M"
-      thus "x \<in> C"
-        proof (induct rule: smallest_ccdi_sets.induct)
-          case (Basic x)
-          thus ?case
-            by (metis Basic subsetD sbC)
-        next
-          case (Compl x)
-          thus ?case
-            by (blast intro: closed_cdi_Compl [OF ccdi, simplified])
-        next
-          case (Inc A)
-          thus ?case
-               by (auto intro: closed_cdi_Inc [OF ccdi, simplified])
-        next
-          case (Disj A)
-          thus ?case
-               by (auto intro: closed_cdi_Disj [OF ccdi, simplified])
-        qed
-    qed
-  finally show ?thesis .
-qed
-
-lemma (in algebra) sigma_property_disjoint:
-  assumes sbC: "M \<subseteq> C"
-      and compl: "!!s. s \<in> C \<inter> sigma_sets (\<Omega>) (M) \<Longrightarrow> \<Omega> - s \<in> C"
-      and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)
-                     \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
-                     \<Longrightarrow> (\<Union>i. A i) \<in> C"
-      and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)
-                      \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"
-  shows "sigma_sets (\<Omega>) (M) \<subseteq> C"
-proof -
-  have "sigma_sets (\<Omega>) (M) \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"
-    proof (rule sigma_property_disjoint_lemma)
-      show "M \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"
-        by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
-    next
-      show "closed_cdi \<Omega> (C \<inter> sigma_sets (\<Omega>) (M))"
-        by (simp add: closed_cdi_def compl inc disj)
-           (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
-             IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
-    qed
-  thus ?thesis
-    by blast
-qed
-
-subsubsection \<open>Dynkin systems\<close>
-
-locale dynkin_system = subset_class +
-  assumes space: "\<Omega> \<in> M"
-    and   compl[intro!]: "\<And>A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
-    and   UN[intro!]: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
-                           \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
-
-lemma (in dynkin_system) empty[intro, simp]: "{} \<in> M"
-  using space compl[of "\<Omega>"] by simp
-
-lemma (in dynkin_system) diff:
-  assumes sets: "D \<in> M" "E \<in> M" and "D \<subseteq> E"
-  shows "E - D \<in> M"
-proof -
-  let ?f = "\<lambda>x. if x = 0 then D else if x = Suc 0 then \<Omega> - E else {}"
-  have "range ?f = {D, \<Omega> - E, {}}"
-    by (auto simp: image_iff)
-  moreover have "D \<union> (\<Omega> - E) = (\<Union>i. ?f i)"
-    by (auto simp: image_iff split: if_split_asm)
-  moreover
-  have "disjoint_family ?f" unfolding disjoint_family_on_def
-    using \<open>D \<in> M\<close>[THEN sets_into_space] \<open>D \<subseteq> E\<close> by auto
-  ultimately have "\<Omega> - (D \<union> (\<Omega> - E)) \<in> M"
-    using sets by auto
-  also have "\<Omega> - (D \<union> (\<Omega> - E)) = E - D"
-    using assms sets_into_space by auto
-  finally show ?thesis .
-qed
-
-lemma dynkin_systemI:
-  assumes "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>" "\<Omega> \<in> M"
-  assumes "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
-  assumes "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
-          \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
-  shows "dynkin_system \<Omega> M"
-  using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def)
-
-lemma dynkin_systemI':
-  assumes 1: "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>"
-  assumes empty: "{} \<in> M"
-  assumes Diff: "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
-  assumes 2: "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
-          \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
-  shows "dynkin_system \<Omega> M"
-proof -
-  from Diff[OF empty] have "\<Omega> \<in> M" by auto
-  from 1 this Diff 2 show ?thesis
-    by (intro dynkin_systemI) auto
-qed
-
-lemma dynkin_system_trivial:
-  shows "dynkin_system A (Pow A)"
-  by (rule dynkin_systemI) auto
-
-lemma sigma_algebra_imp_dynkin_system:
-  assumes "sigma_algebra \<Omega> M" shows "dynkin_system \<Omega> M"
-proof -
-  interpret sigma_algebra \<Omega> M by fact
-  show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI)
-qed
-
-subsubsection "Intersection sets systems"
-
-definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)"
-
-lemma (in algebra) Int_stable: "Int_stable M"
-  unfolding Int_stable_def by auto
-
-lemma Int_stableI:
-  "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A) \<Longrightarrow> Int_stable A"
-  unfolding Int_stable_def by auto
-
-lemma Int_stableD:
-  "Int_stable M \<Longrightarrow> a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"
-  unfolding Int_stable_def by auto
-
-lemma (in dynkin_system) sigma_algebra_eq_Int_stable:
-  "sigma_algebra \<Omega> M \<longleftrightarrow> Int_stable M"
-proof
-  assume "sigma_algebra \<Omega> M" then show "Int_stable M"
-    unfolding sigma_algebra_def using algebra.Int_stable by auto
-next
-  assume "Int_stable M"
-  show "sigma_algebra \<Omega> M"
-    unfolding sigma_algebra_disjoint_iff algebra_iff_Un
-  proof (intro conjI ballI allI impI)
-    show "M \<subseteq> Pow (\<Omega>)" using sets_into_space by auto
-  next
-    fix A B assume "A \<in> M" "B \<in> M"
-    then have "A \<union> B = \<Omega> - ((\<Omega> - A) \<inter> (\<Omega> - B))"
-              "\<Omega> - A \<in> M" "\<Omega> - B \<in> M"
-      using sets_into_space by auto
-    then show "A \<union> B \<in> M"
-      using \<open>Int_stable M\<close> unfolding Int_stable_def by auto
-  qed auto
-qed
-
-subsubsection "Smallest Dynkin systems"
-
-definition dynkin where
-  "dynkin \<Omega> M =  (\<Inter>{D. dynkin_system \<Omega> D \<and> M \<subseteq> D})"
-
-lemma dynkin_system_dynkin:
-  assumes "M \<subseteq> Pow (\<Omega>)"
-  shows "dynkin_system \<Omega> (dynkin \<Omega> M)"
-proof (rule dynkin_systemI)
-  fix A assume "A \<in> dynkin \<Omega> M"
-  moreover
-  { fix D assume "A \<in> D" and d: "dynkin_system \<Omega> D"
-    then have "A \<subseteq> \<Omega>" by (auto simp: dynkin_system_def subset_class_def) }
-  moreover have "{D. dynkin_system \<Omega> D \<and> M \<subseteq> D} \<noteq> {}"
-    using assms dynkin_system_trivial by fastforce
-  ultimately show "A \<subseteq> \<Omega>"
-    unfolding dynkin_def using assms
-    by auto
-next
-  show "\<Omega> \<in> dynkin \<Omega> M"
-    unfolding dynkin_def using dynkin_system.space by fastforce
-next
-  fix A assume "A \<in> dynkin \<Omega> M"
-  then show "\<Omega> - A \<in> dynkin \<Omega> M"
-    unfolding dynkin_def using dynkin_system.compl by force
-next
-  fix A :: "nat \<Rightarrow> 'a set"
-  assume A: "disjoint_family A" "range A \<subseteq> dynkin \<Omega> M"
-  show "(\<Union>i. A i) \<in> dynkin \<Omega> M" unfolding dynkin_def
-  proof (simp, safe)
-    fix D assume "dynkin_system \<Omega> D" "M \<subseteq> D"
-    with A have "(\<Union>i. A i) \<in> D"
-      by (intro dynkin_system.UN) (auto simp: dynkin_def)
-    then show "(\<Union>i. A i) \<in> D" by auto
-  qed
-qed
-
-lemma dynkin_Basic[intro]: "A \<in> M \<Longrightarrow> A \<in> dynkin \<Omega> M"
-  unfolding dynkin_def by auto
-
-lemma (in dynkin_system) restricted_dynkin_system:
-  assumes "D \<in> M"
-  shows "dynkin_system \<Omega> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
-proof (rule dynkin_systemI, simp_all)
-  have "\<Omega> \<inter> D = D"
-    using \<open>D \<in> M\<close> sets_into_space by auto
-  then show "\<Omega> \<inter> D \<in> M"
-    using \<open>D \<in> M\<close> by auto
-next
-  fix A assume "A \<subseteq> \<Omega> \<and> A \<inter> D \<in> M"
-  moreover have "(\<Omega> - A) \<inter> D = (\<Omega> - (A \<inter> D)) - (\<Omega> - D)"
-    by auto
-  ultimately show "\<Omega> - A \<subseteq> \<Omega> \<and> (\<Omega> - A) \<inter> D \<in> M"
-    using  \<open>D \<in> M\<close> by (auto intro: diff)
-next
-  fix A :: "nat \<Rightarrow> 'a set"
-  assume "disjoint_family A" "range A \<subseteq> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
-  then have "\<And>i. A i \<subseteq> \<Omega>" "disjoint_family (\<lambda>i. A i \<inter> D)"
-    "range (\<lambda>i. A i \<inter> D) \<subseteq> M" "(\<Union>x. A x) \<inter> D = (\<Union>x. A x \<inter> D)"
-    by ((fastforce simp: disjoint_family_on_def)+)
-  then show "(\<Union>x. A x) \<subseteq> \<Omega> \<and> (\<Union>x. A x) \<inter> D \<in> M"
-    by (auto simp del: UN_simps)
-qed
-
-lemma (in dynkin_system) dynkin_subset:
-  assumes "N \<subseteq> M"
-  shows "dynkin \<Omega> N \<subseteq> M"
-proof -
-  have "dynkin_system \<Omega> M" ..
-  then have "dynkin_system \<Omega> M"
-    using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp
-  with \<open>N \<subseteq> M\<close> show ?thesis by (auto simp add: dynkin_def)
-qed
-
-lemma sigma_eq_dynkin:
-  assumes sets: "M \<subseteq> Pow \<Omega>"
-  assumes "Int_stable M"
-  shows "sigma_sets \<Omega> M = dynkin \<Omega> M"
-proof -
-  have "dynkin \<Omega> M \<subseteq> sigma_sets (\<Omega>) (M)"
-    using sigma_algebra_imp_dynkin_system
-    unfolding dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto
-  moreover
-  interpret dynkin_system \<Omega> "dynkin \<Omega> M"
-    using dynkin_system_dynkin[OF sets] .
-  have "sigma_algebra \<Omega> (dynkin \<Omega> M)"
-    unfolding sigma_algebra_eq_Int_stable Int_stable_def
-  proof (intro ballI)
-    fix A B assume "A \<in> dynkin \<Omega> M" "B \<in> dynkin \<Omega> M"
-    let ?D = "\<lambda>E. {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> E \<in> dynkin \<Omega> M}"
-    have "M \<subseteq> ?D B"
-    proof
-      fix E assume "E \<in> M"
-      then have "M \<subseteq> ?D E" "E \<in> dynkin \<Omega> M"
-        using sets_into_space \<open>Int_stable M\<close> by (auto simp: Int_stable_def)
-      then have "dynkin \<Omega> M \<subseteq> ?D E"
-        using restricted_dynkin_system \<open>E \<in> dynkin \<Omega> M\<close>
-        by (intro dynkin_system.dynkin_subset) simp_all
-      then have "B \<in> ?D E"
-        using \<open>B \<in> dynkin \<Omega> M\<close> by auto
-      then have "E \<inter> B \<in> dynkin \<Omega> M"
-        by (subst Int_commute) simp
-      then show "E \<in> ?D B"
-        using sets \<open>E \<in> M\<close> by auto
-    qed
-    then have "dynkin \<Omega> M \<subseteq> ?D B"
-      using restricted_dynkin_system \<open>B \<in> dynkin \<Omega> M\<close>
-      by (intro dynkin_system.dynkin_subset) simp_all
-    then show "A \<inter> B \<in> dynkin \<Omega> M"
-      using \<open>A \<in> dynkin \<Omega> M\<close> sets_into_space by auto
-  qed
-  from sigma_algebra.sigma_sets_subset[OF this, of "M"]
-  have "sigma_sets (\<Omega>) (M) \<subseteq> dynkin \<Omega> M" by auto
-  ultimately have "sigma_sets (\<Omega>) (M) = dynkin \<Omega> M" by auto
-  then show ?thesis
-    by (auto simp: dynkin_def)
-qed
-
-lemma (in dynkin_system) dynkin_idem:
-  "dynkin \<Omega> M = M"
-proof -
-  have "dynkin \<Omega> M = M"
-  proof
-    show "M \<subseteq> dynkin \<Omega> M"
-      using dynkin_Basic by auto
-    show "dynkin \<Omega> M \<subseteq> M"
-      by (intro dynkin_subset) auto
-  qed
-  then show ?thesis
-    by (auto simp: dynkin_def)
-qed
-
-lemma (in dynkin_system) dynkin_lemma:
-  assumes "Int_stable E"
-  and E: "E \<subseteq> M" "M \<subseteq> sigma_sets \<Omega> E"
-  shows "sigma_sets \<Omega> E = M"
-proof -
-  have "E \<subseteq> Pow \<Omega>"
-    using E sets_into_space by force
-  then have *: "sigma_sets \<Omega> E = dynkin \<Omega> E"
-    using \<open>Int_stable E\<close> by (rule sigma_eq_dynkin)
-  then have "dynkin \<Omega> E = M"
-    using assms dynkin_subset[OF E(1)] by simp
-  with * show ?thesis
-    using assms by (auto simp: dynkin_def)
-qed
-
-subsubsection \<open>Induction rule for intersection-stable generators\<close>
-
-text \<open>The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras
-generated by a generator closed under intersection.\<close>
-
-lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:
-  assumes "Int_stable G"
-    and closed: "G \<subseteq> Pow \<Omega>"
-    and A: "A \<in> sigma_sets \<Omega> G"
-  assumes basic: "\<And>A. A \<in> G \<Longrightarrow> P A"
-    and empty: "P {}"
-    and compl: "\<And>A. A \<in> sigma_sets \<Omega> G \<Longrightarrow> P A \<Longrightarrow> P (\<Omega> - A)"
-    and union: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> sigma_sets \<Omega> G \<Longrightarrow> (\<And>i. P (A i)) \<Longrightarrow> P (\<Union>i::nat. A i)"
-  shows "P A"
-proof -
-  let ?D = "{ A \<in> sigma_sets \<Omega> G. P A }"
-  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> G"
-    using closed by (rule sigma_algebra_sigma_sets)
-  from compl[OF _ empty] closed have space: "P \<Omega>" by simp
-  interpret dynkin_system \<Omega> ?D
-    by standard (auto dest: sets_into_space intro!: space compl union)
-  have "sigma_sets \<Omega> G = ?D"
-    by (rule dynkin_lemma) (auto simp: basic \<open>Int_stable G\<close>)
-  with A show ?thesis by auto
-qed
-
-subsection \<open>Measure type\<close>
-
-definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
-  "positive M \<mu> \<longleftrightarrow> \<mu> {} = 0"
-
-definition countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
-  "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
-    (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
-
-definition measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> bool" where
-  "measure_space \<Omega> A \<mu> \<longleftrightarrow> sigma_algebra \<Omega> A \<and> positive A \<mu> \<and> countably_additive A \<mu>"
-
-typedef 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
-proof
-  have "sigma_algebra UNIV {{}, UNIV}"
-    by (auto simp: sigma_algebra_iff2)
-  then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
-    by (auto simp: measure_space_def positive_def countably_additive_def)
-qed
-
-definition space :: "'a measure \<Rightarrow> 'a set" where
-  "space M = fst (Rep_measure M)"
-
-definition sets :: "'a measure \<Rightarrow> 'a set set" where
-  "sets M = fst (snd (Rep_measure M))"
-
-definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ennreal" where
-  "emeasure M = snd (snd (Rep_measure M))"
-
-definition measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where
-  "measure M A = enn2real (emeasure M A)"
-
-declare [[coercion sets]]
-
-declare [[coercion measure]]
-
-declare [[coercion emeasure]]
-
-lemma measure_space: "measure_space (space M) (sets M) (emeasure M)"
-  by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse)
-
-interpretation sets: sigma_algebra "space M" "sets M" for M :: "'a measure"
-  using measure_space[of M] by (auto simp: measure_space_def)
-
-definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ennreal) \<Rightarrow> 'a measure" where
-  "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
-    \<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"
-
-abbreviation "sigma \<Omega> A \<equiv> measure_of \<Omega> A (\<lambda>x. 0)"
-
-lemma measure_space_0: "A \<subseteq> Pow \<Omega> \<Longrightarrow> measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>x. 0)"
-  unfolding measure_space_def
-  by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)
-
-lemma sigma_algebra_trivial: "sigma_algebra \<Omega> {{}, \<Omega>}"
-by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\<Omega>}"])+
-
-lemma measure_space_0': "measure_space \<Omega> {{}, \<Omega>} (\<lambda>x. 0)"
-by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial)
-
-lemma measure_space_closed:
-  assumes "measure_space \<Omega> M \<mu>"
-  shows "M \<subseteq> Pow \<Omega>"
-proof -
-  interpret sigma_algebra \<Omega> M using assms by(simp add: measure_space_def)
-  show ?thesis by(rule space_closed)
-qed
-
-lemma (in ring_of_sets) positive_cong_eq:
-  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> positive M \<mu>' = positive M \<mu>"
-  by (auto simp add: positive_def)
-
-lemma (in sigma_algebra) countably_additive_eq:
-  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> countably_additive M \<mu>' = countably_additive M \<mu>"
-  unfolding countably_additive_def
-  by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)
-
-lemma measure_space_eq:
-  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a"
-  shows "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
-proof -
-  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" using closed by (rule sigma_algebra_sigma_sets)
-  from positive_cong_eq[OF eq, of "\<lambda>i. i"] countably_additive_eq[OF eq, of "\<lambda>i. i"] show ?thesis
-    by (auto simp: measure_space_def)
-qed
-
-lemma measure_of_eq:
-  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "(\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a)"
-  shows "measure_of \<Omega> A \<mu> = measure_of \<Omega> A \<mu>'"
-proof -
-  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
-    using assms by (rule measure_space_eq)
-  with eq show ?thesis
-    by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure])
-qed
-
-lemma
-  shows space_measure_of_conv: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)
-  and sets_measure_of_conv:
-  "sets (measure_of \<Omega> A \<mu>) = (if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>})" (is ?sets)
-  and emeasure_measure_of_conv:
-  "emeasure (measure_of \<Omega> A \<mu>) =
-  (\<lambda>B. if B \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> B else 0)" (is ?emeasure)
-proof -
-  have "?space \<and> ?sets \<and> ?emeasure"
-  proof(cases "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>")
-    case True
-    from measure_space_closed[OF this] sigma_sets_superset_generator[of A \<Omega>]
-    have "A \<subseteq> Pow \<Omega>" by simp
-    hence "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)
-      (\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"
-      by(rule measure_space_eq) auto
-    with True \<open>A \<subseteq> Pow \<Omega>\<close> show ?thesis
-      by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse)
-  next
-    case False thus ?thesis
-      by(cases "A \<subseteq> Pow \<Omega>")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0')
-  qed
-  thus ?space ?sets ?emeasure by simp_all
-qed
-
-lemma [simp]:
-  assumes A: "A \<subseteq> Pow \<Omega>"
-  shows sets_measure_of: "sets (measure_of \<Omega> A \<mu>) = sigma_sets \<Omega> A"
-    and space_measure_of: "space (measure_of \<Omega> A \<mu>) = \<Omega>"
-using assms
-by(simp_all add: sets_measure_of_conv space_measure_of_conv)
-
-lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \<Omega> M \<mu>) = M"
-  using space_closed by (auto intro!: sigma_sets_eq)
-
-lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \<Omega> M \<mu>) = \<Omega>"
-  by (rule space_measure_of_conv)
-
-lemma measure_of_subset: "M \<subseteq> Pow \<Omega> \<Longrightarrow> M' \<subseteq> M \<Longrightarrow> sets (measure_of \<Omega> M' \<mu>) \<subseteq> sets (measure_of \<Omega> M \<mu>')"
-  by (auto intro!: sigma_sets_subseteq)
-
-lemma emeasure_sigma: "emeasure (sigma \<Omega> A) = (\<lambda>x. 0)"
-  unfolding measure_of_def emeasure_def
-  by (subst Abs_measure_inverse)
-     (auto simp: measure_space_def positive_def countably_additive_def
-           intro!: sigma_algebra_sigma_sets sigma_algebra_trivial)
-
-lemma sigma_sets_mono'':
-  assumes "A \<in> sigma_sets C D"
-  assumes "B \<subseteq> D"
-  assumes "D \<subseteq> Pow C"
-  shows "sigma_sets A B \<subseteq> sigma_sets C D"
-proof
-  fix x assume "x \<in> sigma_sets A B"
-  thus "x \<in> sigma_sets C D"
-  proof induct
-    case (Basic a) with assms have "a \<in> D" by auto
-    thus ?case ..
-  next
-    case Empty show ?case by (rule sigma_sets.Empty)
-  next
-    from assms have "A \<in> sets (sigma C D)" by (subst sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
-    moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
-    ultimately have "A - a \<in> sets (sigma C D)" ..
-    thus ?case by (subst (asm) sets_measure_of[OF \<open>D \<subseteq> Pow C\<close>])
-  next
-    case (Union a)
-    thus ?case by (intro sigma_sets.Union)
-  qed
-qed
-
-lemma in_measure_of[intro, simp]: "M \<subseteq> Pow \<Omega> \<Longrightarrow> A \<in> M \<Longrightarrow> A \<in> sets (measure_of \<Omega> M \<mu>)"
-  by auto
-
-lemma space_empty_iff: "space N = {} \<longleftrightarrow> sets N = {{}}"
-  by (metis Pow_empty Sup_bot_conv(1) cSup_singleton empty_iff
-            sets.sigma_sets_eq sets.space_closed sigma_sets_top subset_singletonD)
-
-subsubsection \<open>Constructing simple @{typ "'a measure"}\<close>
-
-lemma emeasure_measure_of:
-  assumes M: "M = measure_of \<Omega> A \<mu>"
-  assumes ms: "A \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>" "countably_additive (sets M) \<mu>"
-  assumes X: "X \<in> sets M"
-  shows "emeasure M X = \<mu> X"
-proof -
-  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" by (rule sigma_algebra_sigma_sets) fact
-  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"
-    using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets)
-  thus ?thesis using X ms
-    by(simp add: M emeasure_measure_of_conv sets_measure_of_conv)
-qed
-
-lemma emeasure_measure_of_sigma:
-  assumes ms: "sigma_algebra \<Omega> M" "positive M \<mu>" "countably_additive M \<mu>"
-  assumes A: "A \<in> M"
-  shows "emeasure (measure_of \<Omega> M \<mu>) A = \<mu> A"
-proof -
-  interpret sigma_algebra \<Omega> M by fact
-  have "measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
-    using ms sigma_sets_eq by (simp add: measure_space_def)
-  thus ?thesis by(simp add: emeasure_measure_of_conv A)
-qed
-
-lemma measure_cases[cases type: measure]:
-  obtains (measure) \<Omega> A \<mu> where "x = Abs_measure (\<Omega>, A, \<mu>)" "\<forall>a\<in>-A. \<mu> a = 0" "measure_space \<Omega> A \<mu>"
-  by atomize_elim (cases x, auto)
-
-lemma sets_le_imp_space_le: "sets A \<subseteq> sets B \<Longrightarrow> space A \<subseteq> space B"
-  by (auto dest: sets.sets_into_space)
-
-lemma sets_eq_imp_space_eq: "sets M = sets M' \<Longrightarrow> space M = space M'"
-  by (auto intro!: antisym sets_le_imp_space_le)
-
-lemma emeasure_notin_sets: "A \<notin> sets M \<Longrightarrow> emeasure M A = 0"
-  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
-
-lemma emeasure_neq_0_sets: "emeasure M A \<noteq> 0 \<Longrightarrow> A \<in> sets M"
-  using emeasure_notin_sets[of A M] by blast
-
-lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"
-  by (simp add: measure_def emeasure_notin_sets zero_ennreal.rep_eq)
-
-lemma measure_eqI:
-  fixes M N :: "'a measure"
-  assumes "sets M = sets N" and eq: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A = emeasure N A"
-  shows "M = N"
-proof (cases M N rule: measure_cases[case_product measure_cases])
-  case (measure_measure \<Omega> A \<mu> \<Omega>' A' \<mu>')
-  interpret M: sigma_algebra \<Omega> A using measure_measure by (auto simp: measure_space_def)
-  interpret N: sigma_algebra \<Omega>' A' using measure_measure by (auto simp: measure_space_def)
-  have "A = sets M" "A' = sets N"
-    using measure_measure by (simp_all add: sets_def Abs_measure_inverse)
-  with \<open>sets M = sets N\<close> have AA': "A = A'" by simp
-  moreover from M.top N.top M.space_closed N.space_closed AA' have "\<Omega> = \<Omega>'" by auto
-  moreover { fix B have "\<mu> B = \<mu>' B"
-    proof cases
-      assume "B \<in> A"
-      with eq \<open>A = sets M\<close> have "emeasure M B = emeasure N B" by simp
-      with measure_measure show "\<mu> B = \<mu>' B"
-        by (simp add: emeasure_def Abs_measure_inverse)
-    next
-      assume "B \<notin> A"
-      with \<open>A = sets M\<close> \<open>A' = sets N\<close> \<open>A = A'\<close> have "B \<notin> sets M" "B \<notin> sets N"
-        by auto
-      then have "emeasure M B = 0" "emeasure N B = 0"
-        by (simp_all add: emeasure_notin_sets)
-      with measure_measure show "\<mu> B = \<mu>' B"
-        by (simp add: emeasure_def Abs_measure_inverse)
-    qed }
-  then have "\<mu> = \<mu>'" by auto
-  ultimately show "M = N"
-    by (simp add: measure_measure)
-qed
-
-lemma sigma_eqI:
-  assumes [simp]: "M \<subseteq> Pow \<Omega>" "N \<subseteq> Pow \<Omega>" "sigma_sets \<Omega> M = sigma_sets \<Omega> N"
-  shows "sigma \<Omega> M = sigma \<Omega> N"
-  by (rule measure_eqI) (simp_all add: emeasure_sigma)
-
-subsubsection \<open>Measurable functions\<close>
-
-definition measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set" (infixr "\<rightarrow>\<^sub>M" 60) where
-  "measurable A B = {f \<in> space A \<rightarrow> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
-
-lemma measurableI:
-  "(\<And>x. x \<in> space M \<Longrightarrow> f x \<in> space N) \<Longrightarrow> (\<And>A. A \<in> sets N \<Longrightarrow> f -` A \<inter> space M \<in> sets M) \<Longrightarrow>
-    f \<in> measurable M N"
-  by (auto simp: measurable_def)
-
-lemma measurable_space:
-  "f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"
-   unfolding measurable_def by auto
-
-lemma measurable_sets:
-  "f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
-   unfolding measurable_def by auto
-
-lemma measurable_sets_Collect:
-  assumes f: "f \<in> measurable M N" and P: "{x\<in>space N. P x} \<in> sets N" shows "{x\<in>space M. P (f x)} \<in> sets M"
-proof -
-  have "f -` {x \<in> space N. P x} \<inter> space M = {x\<in>space M. P (f x)}"
-    using measurable_space[OF f] by auto
-  with measurable_sets[OF f P] show ?thesis
-    by simp
-qed
-
-lemma measurable_sigma_sets:
-  assumes B: "sets N = sigma_sets \<Omega> A" "A \<subseteq> Pow \<Omega>"
-      and f: "f \<in> space M \<rightarrow> \<Omega>"
-      and ba: "\<And>y. y \<in> A \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
-  shows "f \<in> measurable M N"
-proof -
-  interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)
-  from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force
-
-  { fix X assume "X \<in> sigma_sets \<Omega> A"
-    then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"
-      proof induct
-        case (Basic a) then show ?case
-          by (auto simp add: ba) (metis B(2) subsetD PowD)
-      next
-        case (Compl a)
-        have [simp]: "f -` \<Omega> \<inter> space M = space M"
-          by (auto simp add: funcset_mem [OF f])
-        then show ?case
-          by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl)
-      next
-        case (Union a)
-        then show ?case
-          by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast
-      qed auto }
-  with f show ?thesis
-    by (auto simp add: measurable_def B \<Omega>)
-qed
-
-lemma measurable_measure_of:
-  assumes B: "N \<subseteq> Pow \<Omega>"
-      and f: "f \<in> space M \<rightarrow> \<Omega>"
-      and ba: "\<And>y. y \<in> N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
-  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>)"
-proof -
-  have "sets (measure_of \<Omega> N \<mu>) = sigma_sets \<Omega> N"
-    using B by (rule sets_measure_of)
-  from this assms show ?thesis by (rule measurable_sigma_sets)
-qed
-
-lemma measurable_iff_measure_of:
-  assumes "N \<subseteq> Pow \<Omega>" "f \<in> space M \<rightarrow> \<Omega>"
-  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>) \<longleftrightarrow> (\<forall>A\<in>N. f -` A \<inter> space M \<in> sets M)"
-  by (metis assms in_measure_of measurable_measure_of assms measurable_sets)
-
-lemma measurable_cong_sets:
-  assumes sets: "sets M = sets M'" "sets N = sets N'"
-  shows "measurable M N = measurable M' N'"
-  using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)
-
-lemma measurable_cong:
-  assumes "\<And>w. w \<in> space M \<Longrightarrow> f w = g w"
-  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
-  unfolding measurable_def using assms
-  by (simp cong: vimage_inter_cong Pi_cong)
-
-lemma measurable_cong':
-  assumes "\<And>w. w \<in> space M =simp=> f w = g w"
-  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
-  unfolding measurable_def using assms
-  by (simp cong: vimage_inter_cong Pi_cong add: simp_implies_def)
-
-lemma measurable_cong_strong:
-  "M = N \<Longrightarrow> M' = N' \<Longrightarrow> (\<And>w. w \<in> space M \<Longrightarrow> f w = g w) \<Longrightarrow>
-    f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable N N'"
-  by (metis measurable_cong)
-
-lemma measurable_compose:
-  assumes f: "f \<in> measurable M N" and g: "g \<in> measurable N L"
-  shows "(\<lambda>x. g (f x)) \<in> measurable M L"
-proof -
-  have "\<And>A. (\<lambda>x. g (f x)) -` A \<inter> space M = f -` (g -` A \<inter> space N) \<inter> space M"
-    using measurable_space[OF f] by auto
-  with measurable_space[OF f] measurable_space[OF g] show ?thesis
-    by (auto intro: measurable_sets[OF f] measurable_sets[OF g]
-             simp del: vimage_Int simp add: measurable_def)
-qed
-
-lemma measurable_comp:
-  "f \<in> measurable M N \<Longrightarrow> g \<in> measurable N L \<Longrightarrow> g \<circ> f \<in> measurable M L"
-  using measurable_compose[of f M N g L] by (simp add: comp_def)
-
-lemma measurable_const:
-  "c \<in> space M' \<Longrightarrow> (\<lambda>x. c) \<in> measurable M M'"
-  by (auto simp add: measurable_def)
-
-lemma measurable_ident: "id \<in> measurable M M"
-  by (auto simp add: measurable_def)
-
-lemma measurable_id: "(\<lambda>x. x) \<in> measurable M M"
-  by (simp add: measurable_def)
-
-lemma measurable_ident_sets:
-  assumes eq: "sets M = sets M'" shows "(\<lambda>x. x) \<in> measurable M M'"
-  using measurable_ident[of M]
-  unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] .
-
-lemma sets_Least:
-  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> M"
-  shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"
-proof -
-  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
-    proof cases
-      assume i: "(LEAST j. False) = i"
-      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
-        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
-        by (simp add: set_eq_iff, safe)
-           (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)
-      with meas show ?thesis
-        by (auto intro!: sets.Int)
-    next
-      assume i: "(LEAST j. False) \<noteq> i"
-      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
-        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
-      proof (simp add: set_eq_iff, safe)
-        fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"
-        have "\<exists>j. P j x"
-          by (rule ccontr) (insert neq, auto)
-        then show "P (LEAST j. P j x) x" by (rule LeastI_ex)
-      qed (auto dest: Least_le intro!: Least_equality)
-      with meas show ?thesis
-        by auto
-    qed }
-  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
-    by (intro sets.countable_UN) auto
-  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
-    (\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma measurable_mono1:
-  "M' \<subseteq> Pow \<Omega> \<Longrightarrow> M \<subseteq> M' \<Longrightarrow>
-    measurable (measure_of \<Omega> M \<mu>) N \<subseteq> measurable (measure_of \<Omega> M' \<mu>') N"
-  using measure_of_subset[of M' \<Omega> M] by (auto simp add: measurable_def)
-
-subsubsection \<open>Counting space\<close>
-
-definition count_space :: "'a set \<Rightarrow> 'a measure" where
-  "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then of_nat (card A) else \<infinity>)"
-
-lemma
-  shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"
-    and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"
-  using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]
-  by (auto simp: count_space_def)
-
-lemma measurable_count_space_eq1[simp]:
-  "f \<in> measurable (count_space A) M \<longleftrightarrow> f \<in> A \<rightarrow> space M"
- unfolding measurable_def by simp
-
-lemma measurable_compose_countable':
-  assumes f: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. f i x) \<in> measurable M N"
-  and g: "g \<in> measurable M (count_space I)" and I: "countable I"
-  shows "(\<lambda>x. f (g x) x) \<in> measurable M N"
-  unfolding measurable_def
-proof safe
-  fix x assume "x \<in> space M" then show "f (g x) x \<in> space N"
-    using measurable_space[OF f] g[THEN measurable_space] by auto
-next
-  fix A assume A: "A \<in> sets N"
-  have "(\<lambda>x. f (g x) x) -` A \<inter> space M = (\<Union>i\<in>I. (g -` {i} \<inter> space M) \<inter> (f i -` A \<inter> space M))"
-    using measurable_space[OF g] by auto
-  also have "\<dots> \<in> sets M"
-    using f[THEN measurable_sets, OF _ A] g[THEN measurable_sets]
-    by (auto intro!: sets.countable_UN' I intro: sets.Int[OF measurable_sets measurable_sets])
-  finally show "(\<lambda>x. f (g x) x) -` A \<inter> space M \<in> sets M" .
-qed
-
-lemma measurable_count_space_eq_countable:
-  assumes "countable A"
-  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
-proof -
-  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
-    with \<open>countable A\<close> have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
-      by (auto dest: countable_subset)
-    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
-    ultimately have "f -` X \<inter> space M \<in> sets M"
-      using \<open>X \<subseteq> A\<close> by (auto intro!: sets.countable_UN' simp del: UN_simps) }
-  then show ?thesis
-    unfolding measurable_def by auto
-qed
-
-lemma measurable_count_space_eq2:
-  "finite A \<Longrightarrow> f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
-  by (intro measurable_count_space_eq_countable countable_finite)
-
-lemma measurable_count_space_eq2_countable:
-  fixes f :: "'a => 'c::countable"
-  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
-  by (intro measurable_count_space_eq_countable countableI_type)
-
-lemma measurable_compose_countable:
-  assumes f: "\<And>i::'i::countable. (\<lambda>x. f i x) \<in> measurable M N" and g: "g \<in> measurable M (count_space UNIV)"
-  shows "(\<lambda>x. f (g x) x) \<in> measurable M N"
-  by (rule measurable_compose_countable'[OF assms]) auto
-
-lemma measurable_count_space_const:
-  "(\<lambda>x. c) \<in> measurable M (count_space UNIV)"
-  by (simp add: measurable_const)
-
-lemma measurable_count_space:
-  "f \<in> measurable (count_space A) (count_space UNIV)"
-  by simp
-
-lemma measurable_compose_rev:
-  assumes f: "f \<in> measurable L N" and g: "g \<in> measurable M L"
-  shows "(\<lambda>x. f (g x)) \<in> measurable M N"
-  using measurable_compose[OF g f] .
-
-lemma measurable_empty_iff:
-  "space N = {} \<Longrightarrow> f \<in> measurable M N \<longleftrightarrow> space M = {}"
-  by (auto simp add: measurable_def Pi_iff)
-
-subsubsection \<open>Extend measure\<close>
-
-definition "extend_measure \<Omega> I G \<mu> =
-  (if (\<exists>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>') \<and> \<not> (\<forall>i\<in>I. \<mu> i = 0)
-      then measure_of \<Omega> (G`I) (SOME \<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>')
-      else measure_of \<Omega> (G`I) (\<lambda>_. 0))"
-
-lemma space_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> space (extend_measure \<Omega> I G \<mu>) = \<Omega>"
-  unfolding extend_measure_def by simp
-
-lemma sets_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> sets (extend_measure \<Omega> I G \<mu>) = sigma_sets \<Omega> (G`I)"
-  unfolding extend_measure_def by simp
-
-lemma emeasure_extend_measure:
-  assumes M: "M = extend_measure \<Omega> I G \<mu>"
-    and eq: "\<And>i. i \<in> I \<Longrightarrow> \<mu>' (G i) = \<mu> i"
-    and ms: "G ` I \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
-    and "i \<in> I"
-  shows "emeasure M (G i) = \<mu> i"
-proof cases
-  assume *: "(\<forall>i\<in>I. \<mu> i = 0)"
-  with M have M_eq: "M = measure_of \<Omega> (G`I) (\<lambda>_. 0)"
-   by (simp add: extend_measure_def)
-  from measure_space_0[OF ms(1)] ms \<open>i\<in>I\<close>
-  have "emeasure M (G i) = 0"
-    by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure)
-  with \<open>i\<in>I\<close> * show ?thesis
-    by simp
-next
-  define P where "P \<mu>' \<longleftrightarrow> (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'" for \<mu>'
-  assume "\<not> (\<forall>i\<in>I. \<mu> i = 0)"
-  moreover
-  have "measure_space (space M) (sets M) \<mu>'"
-    using ms unfolding measure_space_def by auto standard
-  with ms eq have "\<exists>\<mu>'. P \<mu>'"
-    unfolding P_def
-    by (intro exI[of _ \<mu>']) (auto simp add: M space_extend_measure sets_extend_measure)
-  ultimately have M_eq: "M = measure_of \<Omega> (G`I) (Eps P)"
-    by (simp add: M extend_measure_def P_def[symmetric])
-
-  from \<open>\<exists>\<mu>'. P \<mu>'\<close> have P: "P (Eps P)" by (rule someI_ex)
-  show "emeasure M (G i) = \<mu> i"
-  proof (subst emeasure_measure_of[OF M_eq])
-    have sets_M: "sets M = sigma_sets \<Omega> (G`I)"
-      using M_eq ms by (auto simp: sets_extend_measure)
-    then show "G i \<in> sets M" using \<open>i \<in> I\<close> by auto
-    show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \<mu> i"
-      using P \<open>i\<in>I\<close> by (auto simp add: sets_M measure_space_def P_def)
-  qed fact
-qed
-
-lemma emeasure_extend_measure_Pair:
-  assumes M: "M = extend_measure \<Omega> {(i, j). I i j} (\<lambda>(i, j). G i j) (\<lambda>(i, j). \<mu> i j)"
-    and eq: "\<And>i j. I i j \<Longrightarrow> \<mu>' (G i j) = \<mu> i j"
-    and ms: "\<And>i j. I i j \<Longrightarrow> G i j \<in> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
-    and "I i j"
-  shows "emeasure M (G i j) = \<mu> i j"
-  using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) \<open>I i j\<close>
-  by (auto simp: subset_eq)
-
-subsection \<open>The smallest $\sigma$-algebra regarding a function\<close>
-
-definition
-  "vimage_algebra X f M = sigma X {f -` A \<inter> X | A. A \<in> sets M}"
-
-lemma space_vimage_algebra[simp]: "space (vimage_algebra X f M) = X"
-  unfolding vimage_algebra_def by (rule space_measure_of) auto
-
-lemma sets_vimage_algebra: "sets (vimage_algebra X f M) = sigma_sets X {f -` A \<inter> X | A. A \<in> sets M}"
-  unfolding vimage_algebra_def by (rule sets_measure_of) auto
-
-lemma sets_vimage_algebra2:
-  "f \<in> X \<rightarrow> space M \<Longrightarrow> sets (vimage_algebra X f M) = {f -` A \<inter> X | A. A \<in> sets M}"
-  using sigma_sets_vimage_commute[of f X "space M" "sets M"]
-  unfolding sets_vimage_algebra sets.sigma_sets_eq by simp
-
-lemma sets_vimage_algebra_cong: "sets M = sets N \<Longrightarrow> sets (vimage_algebra X f M) = sets (vimage_algebra X f N)"
-  by (simp add: sets_vimage_algebra)
-
-lemma vimage_algebra_cong:
-  assumes "X = Y"
-  assumes "\<And>x. x \<in> Y \<Longrightarrow> f x = g x"
-  assumes "sets M = sets N"
-  shows "vimage_algebra X f M = vimage_algebra Y g N"
-  by (auto simp: vimage_algebra_def assms intro!: arg_cong2[where f=sigma])
-
-lemma in_vimage_algebra: "A \<in> sets M \<Longrightarrow> f -` A \<inter> X \<in> sets (vimage_algebra X f M)"
-  by (auto simp: vimage_algebra_def)
-
-lemma sets_image_in_sets:
-  assumes N: "space N = X"
-  assumes f: "f \<in> measurable N M"
-  shows "sets (vimage_algebra X f M) \<subseteq> sets N"
-  unfolding sets_vimage_algebra N[symmetric]
-  by (rule sets.sigma_sets_subset) (auto intro!: measurable_sets f)
-
-lemma measurable_vimage_algebra1: "f \<in> X \<rightarrow> space M \<Longrightarrow> f \<in> measurable (vimage_algebra X f M) M"
-  unfolding measurable_def by (auto intro: in_vimage_algebra)
-
-lemma measurable_vimage_algebra2:
-  assumes g: "g \<in> space N \<rightarrow> X" and f: "(\<lambda>x. f (g x)) \<in> measurable N M"
-  shows "g \<in> measurable N (vimage_algebra X f M)"
-  unfolding vimage_algebra_def
-proof (rule measurable_measure_of)
-  fix A assume "A \<in> {f -` A \<inter> X | A. A \<in> sets M}"
-  then obtain Y where Y: "Y \<in> sets M" and A: "A = f -` Y \<inter> X"
-    by auto
-  then have "g -` A \<inter> space N = (\<lambda>x. f (g x)) -` Y \<inter> space N"
-    using g by auto
-  also have "\<dots> \<in> sets N"
-    using f Y by (rule measurable_sets)
-  finally show "g -` A \<inter> space N \<in> sets N" .
-qed (insert g, auto)
-
-lemma vimage_algebra_sigma:
-  assumes X: "X \<subseteq> Pow \<Omega>'" and f: "f \<in> \<Omega> \<rightarrow> \<Omega>'"
-  shows "vimage_algebra \<Omega> f (sigma \<Omega>' X) = sigma \<Omega> {f -` A \<inter> \<Omega> | A. A \<in> X }" (is "?V = ?S")
-proof (rule measure_eqI)
-  have \<Omega>: "{f -` A \<inter> \<Omega> |A. A \<in> X} \<subseteq> Pow \<Omega>" by auto
-  show "sets ?V = sets ?S"
-    using sigma_sets_vimage_commute[OF f, of X]
-    by (simp add: space_measure_of_conv f sets_vimage_algebra2 \<Omega> X)
-qed (simp add: vimage_algebra_def emeasure_sigma)
-
-lemma vimage_algebra_vimage_algebra_eq:
-  assumes *: "f \<in> X \<rightarrow> Y" "g \<in> Y \<rightarrow> space M"
-  shows "vimage_algebra X f (vimage_algebra Y g M) = vimage_algebra X (\<lambda>x. g (f x)) M"
-    (is "?VV = ?V")
-proof (rule measure_eqI)
-  have "(\<lambda>x. g (f x)) \<in> X \<rightarrow> space M" "\<And>A. A \<inter> f -` Y \<inter> X = A \<inter> X"
-    using * by auto
-  with * show "sets ?VV = sets ?V"
-    by (simp add: sets_vimage_algebra2 ex_simps[symmetric] vimage_comp comp_def del: ex_simps)
-qed (simp add: vimage_algebra_def emeasure_sigma)
-
-subsubsection \<open>Restricted Space Sigma Algebra\<close>
-
-definition restrict_space where
-  "restrict_space M \<Omega> = measure_of (\<Omega> \<inter> space M) ((op \<inter> \<Omega>) ` sets M) (emeasure M)"
-
-lemma space_restrict_space: "space (restrict_space M \<Omega>) = \<Omega> \<inter> space M"
-  using sets.sets_into_space unfolding restrict_space_def by (subst space_measure_of) auto
-
-lemma space_restrict_space2: "\<Omega> \<in> sets M \<Longrightarrow> space (restrict_space M \<Omega>) = \<Omega>"
-  by (simp add: space_restrict_space sets.sets_into_space)
-
-lemma sets_restrict_space: "sets (restrict_space M \<Omega>) = (op \<inter> \<Omega>) ` sets M"
-  unfolding restrict_space_def
-proof (subst sets_measure_of)
-  show "op \<inter> \<Omega> ` sets M \<subseteq> Pow (\<Omega> \<inter> space M)"
-    by (auto dest: sets.sets_into_space)
-  have "sigma_sets (\<Omega> \<inter> space M) {((\<lambda>x. x) -` X) \<inter> (\<Omega> \<inter> space M) | X. X \<in> sets M} =
-    (\<lambda>X. X \<inter> (\<Omega> \<inter> space M)) ` sets M"
-    by (subst sigma_sets_vimage_commute[symmetric, where \<Omega>' = "space M"])
-       (auto simp add: sets.sigma_sets_eq)
-  moreover have "{((\<lambda>x. x) -` X) \<inter> (\<Omega> \<inter> space M) | X. X \<in> sets M} = (\<lambda>X. X \<inter> (\<Omega> \<inter> space M)) `  sets M"
-    by auto
-  moreover have "(\<lambda>X. X \<inter> (\<Omega> \<inter> space M)) `  sets M = (op \<inter> \<Omega>) ` sets M"
-    by (intro image_cong) (auto dest: sets.sets_into_space)
-  ultimately show "sigma_sets (\<Omega> \<inter> space M) (op \<inter> \<Omega> ` sets M) = op \<inter> \<Omega> ` sets M"
-    by simp
-qed
-
-lemma restrict_space_sets_cong:
-  "A = B \<Longrightarrow> sets M = sets N \<Longrightarrow> sets (restrict_space M A) = sets (restrict_space N B)"
-  by (auto simp: sets_restrict_space)
-
-lemma sets_restrict_space_count_space :
-  "sets (restrict_space (count_space A) B) = sets (count_space (A \<inter> B))"
-by(auto simp add: sets_restrict_space)
-
-lemma sets_restrict_UNIV[simp]: "sets (restrict_space M UNIV) = sets M"
-  by (auto simp add: sets_restrict_space)
-
-lemma sets_restrict_restrict_space:
-  "sets (restrict_space (restrict_space M A) B) = sets (restrict_space M (A \<inter> B))"
-  unfolding sets_restrict_space image_comp by (intro image_cong) auto
-
-lemma sets_restrict_space_iff:
-  "\<Omega> \<inter> space M \<in> sets M \<Longrightarrow> A \<in> sets (restrict_space M \<Omega>) \<longleftrightarrow> (A \<subseteq> \<Omega> \<and> A \<in> sets M)"
-proof (subst sets_restrict_space, safe)
-  fix A assume "\<Omega> \<inter> space M \<in> sets M" and A: "A \<in> sets M"
-  then have "(\<Omega> \<inter> space M) \<inter> A \<in> sets M"
-    by rule
-  also have "(\<Omega> \<inter> space M) \<inter> A = \<Omega> \<inter> A"
-    using sets.sets_into_space[OF A] by auto
-  finally show "\<Omega> \<inter> A \<in> sets M"
-    by auto
-qed auto
-
-lemma sets_restrict_space_cong: "sets M = sets N \<Longrightarrow> sets (restrict_space M \<Omega>) = sets (restrict_space N \<Omega>)"
-  by (simp add: sets_restrict_space)
-
-lemma restrict_space_eq_vimage_algebra:
-  "\<Omega> \<subseteq> space M \<Longrightarrow> sets (restrict_space M \<Omega>) = sets (vimage_algebra \<Omega> (\<lambda>x. x) M)"
-  unfolding restrict_space_def
-  apply (subst sets_measure_of)
-  apply (auto simp add: image_subset_iff dest: sets.sets_into_space) []
-  apply (auto simp add: sets_vimage_algebra intro!: arg_cong2[where f=sigma_sets])
-  done
-
-lemma sets_Collect_restrict_space_iff:
-  assumes "S \<in> sets M"
-  shows "{x\<in>space (restrict_space M S). P x} \<in> sets (restrict_space M S) \<longleftrightarrow> {x\<in>space M. x \<in> S \<and> P x} \<in> sets M"
-proof -
-  have "{x\<in>S. P x} = {x\<in>space M. x \<in> S \<and> P x}"
-    using sets.sets_into_space[OF assms] by auto
-  then show ?thesis
-    by (subst sets_restrict_space_iff) (auto simp add: space_restrict_space assms)
-qed
-
-lemma measurable_restrict_space1:
-  assumes f: "f \<in> measurable M N"
-  shows "f \<in> measurable (restrict_space M \<Omega>) N"
-  unfolding measurable_def
-proof (intro CollectI conjI ballI)
-  show sp: "f \<in> space (restrict_space M \<Omega>) \<rightarrow> space N"
-    using measurable_space[OF f] by (auto simp: space_restrict_space)
-
-  fix A assume "A \<in> sets N"
-  have "f -` A \<inter> space (restrict_space M \<Omega>) = (f -` A \<inter> space M) \<inter> (\<Omega> \<inter> space M)"
-    by (auto simp: space_restrict_space)
-  also have "\<dots> \<in> sets (restrict_space M \<Omega>)"
-    unfolding sets_restrict_space
-    using measurable_sets[OF f \<open>A \<in> sets N\<close>] by blast
-  finally show "f -` A \<inter> space (restrict_space M \<Omega>) \<in> sets (restrict_space M \<Omega>)" .
-qed
-
-lemma measurable_restrict_space2_iff:
-  "f \<in> measurable M (restrict_space N \<Omega>) \<longleftrightarrow> (f \<in> measurable M N \<and> f \<in> space M \<rightarrow> \<Omega>)"
-proof -
-  have "\<And>A. f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f -` \<Omega> \<inter> f -` A \<inter> space M = f -` A \<inter> space M"
-    by auto
-  then show ?thesis
-    by (auto simp: measurable_def space_restrict_space Pi_Int[symmetric] sets_restrict_space)
-qed
-
-lemma measurable_restrict_space2:
-  "f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> f \<in> measurable M (restrict_space N \<Omega>)"
-  by (simp add: measurable_restrict_space2_iff)
-
-lemma measurable_piecewise_restrict:
-  assumes I: "countable C"
-    and X: "\<And>\<Omega>. \<Omega> \<in> C \<Longrightarrow> \<Omega> \<inter> space M \<in> sets M" "space M \<subseteq> \<Union>C"
-    and f: "\<And>\<Omega>. \<Omega> \<in> C \<Longrightarrow> f \<in> measurable (restrict_space M \<Omega>) N"
-  shows "f \<in> measurable M N"
-proof (rule measurableI)
-  fix x assume "x \<in> space M"
-  with X obtain \<Omega> where "\<Omega> \<in> C" "x \<in> \<Omega>" "x \<in> space M" by auto
-  then show "f x \<in> space N"
-    by (auto simp: space_restrict_space intro: f measurable_space)
-next
-  fix A assume A: "A \<in> sets N"
-  have "f -` A \<inter> space M = (\<Union>\<Omega>\<in>C. (f -` A \<inter> (\<Omega> \<inter> space M)))"
-    using X by (auto simp: subset_eq)
-  also have "\<dots> \<in> sets M"
-    using measurable_sets[OF f A] X I
-    by (intro sets.countable_UN') (auto simp: sets_restrict_space_iff space_restrict_space)
-  finally show "f -` A \<inter> space M \<in> sets M" .
-qed
-
-lemma measurable_piecewise_restrict_iff:
-  "countable C \<Longrightarrow> (\<And>\<Omega>. \<Omega> \<in> C \<Longrightarrow> \<Omega> \<inter> space M \<in> sets M) \<Longrightarrow> space M \<subseteq> (\<Union>C) \<Longrightarrow>
-    f \<in> measurable M N \<longleftrightarrow> (\<forall>\<Omega>\<in>C. f \<in> measurable (restrict_space M \<Omega>) N)"
-  by (auto intro: measurable_piecewise_restrict measurable_restrict_space1)
-
-lemma measurable_If_restrict_space_iff:
-  "{x\<in>space M. P x} \<in> sets M \<Longrightarrow>
-    (\<lambda>x. if P x then f x else g x) \<in> measurable M N \<longleftrightarrow>
-    (f \<in> measurable (restrict_space M {x. P x}) N \<and> g \<in> measurable (restrict_space M {x. \<not> P x}) N)"
-  by (subst measurable_piecewise_restrict_iff[where C="{{x. P x}, {x. \<not> P x}}"])
-     (auto simp: Int_def sets.sets_Collect_neg space_restrict_space conj_commute[of _ "x \<in> space M" for x]
-           cong: measurable_cong')
-
-lemma measurable_If:
-  "f \<in> measurable M M' \<Longrightarrow> g \<in> measurable M M' \<Longrightarrow> {x\<in>space M. P x} \<in> sets M \<Longrightarrow>
-    (\<lambda>x. if P x then f x else g x) \<in> measurable M M'"
-  unfolding measurable_If_restrict_space_iff by (auto intro: measurable_restrict_space1)
-
-lemma measurable_If_set:
-  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
-  assumes P: "A \<inter> space M \<in> sets M"
-  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
-proof (rule measurable_If[OF measure])
-  have "{x \<in> space M. x \<in> A} = A \<inter> space M" by auto
-  thus "{x \<in> space M. x \<in> A} \<in> sets M" using \<open>A \<inter> space M \<in> sets M\<close> by auto
-qed
-
-lemma measurable_restrict_space_iff:
-  "\<Omega> \<inter> space M \<in> sets M \<Longrightarrow> c \<in> space N \<Longrightarrow>
-    f \<in> measurable (restrict_space M \<Omega>) N \<longleftrightarrow> (\<lambda>x. if x \<in> \<Omega> then f x else c) \<in> measurable M N"
-  by (subst measurable_If_restrict_space_iff)
-     (simp_all add: Int_def conj_commute measurable_const)
-
-lemma restrict_space_singleton: "{x} \<in> sets M \<Longrightarrow> sets (restrict_space M {x}) = sets (count_space {x})"
-  using sets_restrict_space_iff[of "{x}" M]
-  by (auto simp add: sets_restrict_space_iff dest!: subset_singletonD)
-
-lemma measurable_restrict_countable:
-  assumes X[intro]: "countable X"
-  assumes sets[simp]: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
-  assumes space[simp]: "\<And>x. x \<in> X \<Longrightarrow> f x \<in> space N"
-  assumes f: "f \<in> measurable (restrict_space M (- X)) N"
-  shows "f \<in> measurable M N"
-  using f sets.countable[OF sets X]
-  by (intro measurable_piecewise_restrict[where M=M and C="{- X} \<union> ((\<lambda>x. {x}) ` X)"])
-     (auto simp: Diff_Int_distrib2 Compl_eq_Diff_UNIV Int_insert_left sets.Diff restrict_space_singleton
-           simp del: sets_count_space  cong: measurable_cong_sets)
-
-lemma measurable_discrete_difference:
-  assumes f: "f \<in> measurable M N"
-  assumes X: "countable X" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" "\<And>x. x \<in> X \<Longrightarrow> g x \<in> space N"
-  assumes eq: "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
-  shows "g \<in> measurable M N"
-  by (rule measurable_restrict_countable[OF X])
-     (auto simp: eq[symmetric] space_restrict_space cong: measurable_cong' intro: f measurable_restrict_space1)
-
-end
--- a/src/HOL/Probability/measurable.ML	Sun Aug 07 12:10:49 2016 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,280 +0,0 @@
-(*  Title:      HOL/Probability/measurable.ML
-    Author:     Johannes Hölzl <hoelzl@in.tum.de>
-
-Measurability prover.
-*)
-
-signature MEASURABLE = 
-sig
-  type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
-
-  datatype level = Concrete | Generic
-
-  val dest_thm_attr : attribute context_parser
-  val cong_thm_attr : attribute context_parser
-  val measurable_thm_attr : bool * (bool * level) -> attribute
-
-  val add_del_cong_thm : bool -> thm -> Context.generic -> Context.generic ;
-
-  val get_all : Context.generic -> thm list
-  val get_dest : Context.generic -> thm list
-  val get_cong : Context.generic -> thm list
-
-  val measurable_tac : Proof.context -> thm list -> tactic
-
-  val simproc : Proof.context -> cterm -> thm option
-
-  val add_preprocessor : string -> preprocessor -> Context.generic -> Context.generic
-  val del_preprocessor : string -> Context.generic -> Context.generic
-  val add_local_cong : thm -> Proof.context -> Proof.context
-
-  val prepare_facts : Proof.context -> thm list -> (thm list * Proof.context)
-end ;
-
-structure Measurable : MEASURABLE =
-struct
-
-type preprocessor = thm -> Proof.context -> (thm list * Proof.context)
-
-datatype level = Concrete | Generic;
-
-fun eq_measurable_thms ((th1, d1), (th2, d2)) = 
-  d1 = d2 andalso Thm.eq_thm_prop (th1, th2) ;
-
-fun merge_dups (xs:(string * preprocessor) list) ys =
-  xs @ (filter (fn (name, _) => is_none (find_first (fn (name', _) => name' = name) xs)) ys) 
-
-structure Data = Generic_Data
-(
-  type T = {
-    measurable_thms : (thm * (bool * level)) Item_Net.T,
-    dest_thms : thm Item_Net.T,
-    cong_thms : thm Item_Net.T,
-    preprocessors : (string * preprocessor) list }
-  val empty: T = {
-    measurable_thms = Item_Net.init eq_measurable_thms (single o Thm.prop_of o fst),
-    dest_thms = Thm.full_rules,
-    cong_thms = Thm.full_rules,
-    preprocessors = [] };
-  val extend = I;
-  fun merge ({measurable_thms = t1, dest_thms = dt1, cong_thms = ct1, preprocessors = i1 },
-      {measurable_thms = t2, dest_thms = dt2, cong_thms = ct2, preprocessors = i2 }) : T = {
-    measurable_thms = Item_Net.merge (t1, t2),
-    dest_thms = Item_Net.merge (dt1, dt2),
-    cong_thms = Item_Net.merge (ct1, ct2),
-    preprocessors = merge_dups i1 i2 
-    };
-);
-
-val debug =
-  Attrib.setup_config_bool @{binding measurable_debug} (K false)
-
-val split =
-  Attrib.setup_config_bool @{binding measurable_split} (K true)
-
-fun map_data f1 f2 f3 f4
-  {measurable_thms = t1,    dest_thms = t2,    cong_thms = t3,    preprocessors = t4 } =
-  {measurable_thms = f1 t1, dest_thms = f2 t2, cong_thms = f3 t3, preprocessors = f4 t4}
-
-fun map_measurable_thms f = map_data f I I I
-fun map_dest_thms f = map_data I f I I
-fun map_cong_thms f = map_data I I f I
-fun map_preprocessors f = map_data I I I f
-
-fun generic_add_del map : attribute context_parser =
-  Scan.lift
-    (Args.add >> K Item_Net.update || Args.del >> K Item_Net.remove || Scan.succeed Item_Net.update) >>
-    (fn f => Thm.declaration_attribute (Data.map o map o f))
-
-val dest_thm_attr = generic_add_del map_dest_thms
-
-val cong_thm_attr = generic_add_del map_cong_thms
-
-fun del_thm th net =
-  let
-    val thms = net |> Item_Net.content |> filter (fn (th', _) => Thm.eq_thm (th, th'))
-  in fold Item_Net.remove thms net end ;
-
-fun measurable_thm_attr (do_add, d) = Thm.declaration_attribute
-  (Data.map o map_measurable_thms o (if do_add then Item_Net.update o rpair d else del_thm))
-
-val get_dest = Item_Net.content o #dest_thms o Data.get;
-
-val get_cong = Item_Net.content o #cong_thms o Data.get;
-val add_cong = Data.map o map_cong_thms o Item_Net.update;
-val del_cong = Data.map o map_cong_thms o Item_Net.remove;
-fun add_del_cong_thm true = add_cong
-  | add_del_cong_thm false = del_cong
-
-fun add_preprocessor name f = Data.map (map_preprocessors (fn xs => xs @ [(name, f)]))
-fun del_preprocessor name = Data.map (map_preprocessors (filter (fn (n, _) => n <> name)))
-val add_local_cong = Context.proof_map o add_cong
-
-val get_preprocessors = Context.Proof #> Data.get #> #preprocessors ;
-
-fun is_too_generic thm =
-  let 
-    val concl = Thm.concl_of thm
-    val concl' = HOLogic.dest_Trueprop concl handle TERM _ => concl
-  in is_Var (head_of concl') end
-
-val get_thms = Data.get #> #measurable_thms #> Item_Net.content ;
-
-val get_all = get_thms #> map fst ;
-
-fun debug_tac ctxt msg f =
-  if Config.get ctxt debug then print_tac ctxt (msg ()) THEN f else f
-
-fun nth_hol_goal thm i =
-  HOLogic.dest_Trueprop (Logic.strip_imp_concl (strip_all_body (nth (Thm.prems_of thm) (i - 1))))
-
-fun dest_measurable_fun t =
-  (case t of
-    (Const (@{const_name "Set.member"}, _) $ f $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => f
-  | _ => raise (TERM ("not a measurability predicate", [t])))
-
-fun not_measurable_prop n thm =
-  if length (Thm.prems_of thm) < n then false
-  else
-    (case nth_hol_goal thm n of
-      (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "sets"}, _) $ _)) => false
-    | (Const (@{const_name "Set.member"}, _) $ _ $ (Const (@{const_name "measurable"}, _) $ _ $ _)) => false
-    | _ => true)
-    handle TERM _ => true;
-
-fun indep (Bound i) t b = i < b orelse t <= i
-  | indep (f $ t) top bot = indep f top bot andalso indep t top bot
-  | indep (Abs (_,_,t)) top bot = indep t (top + 1) (bot + 1)
-  | indep _ _ _ = true;
-
-fun cnt_prefixes ctxt (Abs (n, T, t)) =
-    let
-      fun is_countable ty = Sign.of_sort (Proof_Context.theory_of ctxt) (ty, @{sort countable})
-      fun cnt_walk (Abs (ns, T, t)) Ts =
-          map (fn (t', t'') => (Abs (ns, T, t'), t'')) (cnt_walk t (T::Ts))
-        | cnt_walk (f $ g) Ts = let
-            val n = length Ts - 1
-          in
-            map (fn (f', t) => (f' $ g, t)) (cnt_walk f Ts) @
-            map (fn (g', t) => (f $ g', t)) (cnt_walk g Ts) @
-            (if is_countable (type_of1 (Ts, g)) andalso loose_bvar1 (g, n)
-                andalso indep g n 0 andalso g <> Bound n
-              then [(f $ Bound (n + 1), incr_boundvars (~ n) g)]
-              else [])
-          end
-        | cnt_walk _ _ = []
-    in map (fn (t1, t2) => let
-        val T1 = type_of1 ([T], t2)
-        val T2 = type_of1 ([T], t)
-      in ([SOME (Abs (n, T1, Abs (n, T, t1))), NONE, NONE, SOME (Abs (n, T, t2))],
-        [SOME T1, SOME T, SOME T2])
-      end) (cnt_walk t [T])
-    end
-  | cnt_prefixes _ _ = []
-
-fun apply_dests thm dests =
-  let
-    fun apply thm th' =
-      let
-        val th'' = thm RS th'
-      in [th''] @ loop th'' end
-      handle (THM _) => []
-    and loop thm =
-      flat (map (apply thm) dests)
-  in
-    [thm] @ ([thm RS @{thm measurable_compose_rev}] handle (THM _) => []) @ loop thm
-  end
-
-fun prepare_facts ctxt facts = 
-  let
-    val dests = get_dest (Context.Proof ctxt)
-    fun prep_dest thm =
-      (if is_too_generic thm then [] else apply_dests thm dests) ;
-    val preprocessors = (("std", prep_dest #> pair) :: get_preprocessors ctxt) ;
-    fun preprocess_thm (thm, raw) =
-      if raw then pair [thm] else fold_map (fn (_, proc) => proc thm) preprocessors #>> flat
-    
-    fun sel lv (th, (raw, lv')) = if lv = lv' then SOME (th, raw) else NONE ;
-    fun get lv = ctxt |> Context.Proof |> get_thms |> rev |> map_filter (sel lv) ;
-    val pre_thms = map (Simplifier.norm_hhf ctxt #> rpair false) facts @ get Concrete @ get Generic
-
-    val (thms, ctxt) = fold_map preprocess_thm pre_thms ctxt |>> flat
-  in (thms, ctxt) end
-
-fun measurable_tac ctxt facts =
-  let
-    fun debug_fact msg thm () =
-      msg ^ " " ^ Pretty.unformatted_string_of (Syntax.pretty_term ctxt (Thm.prop_of thm))
-
-    fun IF' c t i = COND (c i) (t i) no_tac
-
-    fun r_tac msg =
-      if Config.get ctxt debug
-      then FIRST' o
-        map (fn thm => resolve_tac ctxt [thm]
-          THEN' K (debug_tac ctxt (debug_fact (msg ^ " resolved using") thm) all_tac))
-      else resolve_tac ctxt
-
-    val elem_congI = @{lemma "A = B \<Longrightarrow> x \<in> B \<Longrightarrow> x \<in> A" by simp}
-
-    val (thms, ctxt) = prepare_facts ctxt facts
-
-    fun is_sets_eq (Const (@{const_name "HOL.eq"}, _) $
-          (Const (@{const_name "sets"}, _) $ _) $
-          (Const (@{const_name "sets"}, _) $ _)) = true
-      | is_sets_eq (Const (@{const_name "HOL.eq"}, _) $
-          (Const (@{const_name "measurable"}, _) $ _ $ _) $
-          (Const (@{const_name "measurable"}, _) $ _ $ _)) = true
-      | is_sets_eq _ = false
-
-    val cong_thms = get_cong (Context.Proof ctxt) @
-      filter (fn thm => Thm.concl_of thm |> HOLogic.dest_Trueprop |> is_sets_eq handle TERM _ => false) facts
-
-    fun sets_cong_tac i =
-      Subgoal.FOCUS (fn {context = ctxt', prems = prems, ...} => (
-        let
-          val ctxt'' = Simplifier.add_prems prems ctxt'
-        in
-          r_tac "cong intro" [elem_congI]
-          THEN' SOLVED' (fn i => REPEAT_DETERM (
-              ((r_tac "cong solve" (cong_thms @ [@{thm refl}])
-                ORELSE' IF' (fn i => fn thm => Thm.nprems_of thm > i)
-                  (SOLVED' (asm_full_simp_tac ctxt''))) i)))
-        end) 1) ctxt i
-        THEN flexflex_tac ctxt
-
-    val simp_solver_tac = 
-      IF' not_measurable_prop (debug_tac ctxt (K "simp ") o SOLVED' (asm_full_simp_tac ctxt))
-
-    val split_countable_tac =
-      Subgoal.FOCUS (fn {context = ctxt, ...} => SUBGOAL (fn (t, i) =>
-        let
-          val f = dest_measurable_fun (HOLogic.dest_Trueprop t)
-          fun inst (ts, Ts) =
-            Thm.instantiate'
-              (map (Option.map (Thm.ctyp_of ctxt)) Ts)
-              (map (Option.map (Thm.cterm_of ctxt)) ts)
-              @{thm measurable_compose_countable}
-        in r_tac "case_prod countable" (cnt_prefixes ctxt f |> map inst) i end
-        handle TERM _ => no_tac) 1)
-
-    val splitter = if Config.get ctxt split then split_countable_tac ctxt else K no_tac
-
-    val single_step_tac =
-      simp_solver_tac
-      ORELSE' r_tac "step" thms
-      ORELSE' splitter
-      ORELSE' (CHANGED o sets_cong_tac)
-      ORELSE' (K (debug_tac ctxt (K "backtrack") no_tac))
-
-  in debug_tac ctxt (K "start") (REPEAT (single_step_tac 1)) end;
-
-fun simproc ctxt redex =
-  let
-    val t = HOLogic.mk_Trueprop (Thm.term_of redex);
-    fun tac {context = ctxt, prems = _ } =
-      SOLVE (measurable_tac ctxt (Simplifier.prems_of ctxt));
-  in try (fn () => Goal.prove ctxt [] [] t tac RS @{thm Eq_TrueI}) () end;
-
-end
-