merged
authorAndreas Lochbihler
Wed, 10 Oct 2012 15:17:18 +0200
changeset 49809 39db47ed6d54
parent 49808 418991ce7567 (current diff)
parent 49805 9a2a53be24a2 (diff)
child 49810 53f14f62cca2
merged
--- a/src/HOL/Probability/Binary_Product_Measure.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Binary_Product_Measure.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -33,15 +33,22 @@
       {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
       (\<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+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 space_closed[of A] space_closed[of B] by auto
+
 lemma space_pair_measure:
   "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
-  unfolding pair_measure_def using space_closed[of A] space_closed[of B]
-  by (intro space_measure_of) auto
+  unfolding pair_measure_def using pair_measure_closed[of A B]
+  by (rule space_measure_of)
 
 lemma sets_pair_measure:
   "sets (A \<Otimes>\<^isub>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 space_closed[of A] space_closed[of B]
-  by (intro sets_measure_of) auto
+  unfolding pair_measure_def using pair_measure_closed[of A B]
+  by (rule sets_measure_of)
+
+lemma sets_pair_measure_cong[cong]:
+  "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^isub>M M2) = sets (M1' \<Otimes>\<^isub>M M2')"
+  unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
 
 lemma pair_measureI[intro, simp]:
   "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
@@ -54,21 +61,30 @@
   unfolding pair_measure_def using 1 2
   by (intro measurable_measure_of) (auto dest: sets_into_space)
 
+lemma measurable_Pair:
+  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>\<^isub>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 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_pair:
+  assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
+  shows "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+  using measurable_Pair[OF assms] by simp
+
 lemma measurable_fst[intro!, simp]: "fst \<in> measurable (M1 \<Otimes>\<^isub>M M2) M1"
-  unfolding measurable_def
-proof safe
-  fix A assume A: "A \<in> sets M1"
-  from this[THEN sets_into_space] have "fst -` A \<inter> space M1 \<times> space M2 = A \<times> space M2" by auto
-  with A show "fst -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure)
-qed (simp add: space_pair_measure)
+  by (auto simp: fst_vimage_eq_Times space_pair_measure sets_into_space times_Int_times measurable_def)
 
 lemma measurable_snd[intro!, simp]: "snd \<in> measurable (M1 \<Otimes>\<^isub>M M2) M2"
-  unfolding measurable_def
-proof safe
-  fix A assume A: "A \<in> sets M2"
-  from this[THEN sets_into_space] have "snd -` A \<inter> space M1 \<times> space M2 = space M1 \<times> A" by auto
-  with A show "snd -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure)
-qed (simp add: space_pair_measure)
+  by (auto simp: snd_vimage_eq_Times space_pair_measure sets_into_space times_Int_times measurable_def)
 
 lemma measurable_fst': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. fst (f x)) \<in> measurable M N"
   using measurable_comp[OF _ measurable_fst] by (auto simp: comp_def)
@@ -84,55 +100,29 @@
 
 lemma measurable_pair_iff:
   "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
-proof safe
-  assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
-  show "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
-  proof (rule measurable_pair_measureI)
-    show "f \<in> space M \<rightarrow> space M1 \<times> space M2"
-      using f s by (auto simp: mem_Times_iff measurable_def comp_def)
-    fix A B assume "A \<in> sets M1" "B \<in> sets M2"
-    moreover have "(fst \<circ> f) -` A \<inter> space M \<in> sets M" "(snd \<circ> f) -` B \<inter> space M \<in> sets M"
-      using f `A \<in> sets M1` s `B \<in> sets M2` by (auto simp: measurable_sets)
-    moreover have "f -` (A \<times> B) \<inter> space M = ((fst \<circ> f) -` A \<inter> space M) \<inter> ((snd \<circ> f) -` B \<inter> space M)"
-      by (auto simp: vimage_Times)
-    ultimately show "f -` (A \<times> B) \<inter> space M \<in> sets M" by auto
-  qed
-qed auto
+  using measurable_pair[of f M M1 M2] by auto
 
-lemma measurable_pair:
-  "(fst \<circ> f) \<in> measurable M M1 \<Longrightarrow> (snd \<circ> f) \<in> measurable M M2 \<Longrightarrow> f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
-  unfolding measurable_pair_iff by simp
+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>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
-proof (rule measurable_pair_measureI)
-  fix A B assume "A \<in> sets M2" "B \<in> sets M1"
-  moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) = B \<times> A"
-    by (auto dest: sets_into_space simp: space_pair_measure)
-  ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-    by auto
-qed (auto simp add: space_pair_measure)
+  by (auto intro!: measurable_Pair simp: measurable_split_conv)
 
 lemma measurable_pair_swap:
   assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
-proof -
-  have "(\<lambda>x. f (case x of (x, y) \<Rightarrow> (y, x))) = (\<lambda>(x, y). f (y, x))" by auto
-  then show ?thesis
-    using measurable_comp[OF measurable_pair_swap' f] by (simp add: comp_def)
-qed
+  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>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) M"
   using measurable_pair_swap[of "\<lambda>(x,y). f (y, x)"]
   by (auto intro!: measurable_pair_swap)
 
+lemma measurable_ident[intro, simp]: "(\<lambda>x. x) \<in> measurable M M"
+  unfolding measurable_def by auto
+
 lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^isub>M M2)"
-proof (rule measurable_pair_measureI)
-  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
-  moreover then have "Pair x -` (A \<times> B) \<inter> space M2 = (if x \<in> A then B else {})"
-    by (auto dest: sets_into_space simp: space_pair_measure)
-  ultimately show "Pair x -` (A \<times> B) \<inter> space M2 \<in> sets M2"
-    by auto
-qed (auto simp add: space_pair_measure)
+  by (auto intro!: measurable_Pair)
 
 lemma sets_Pair1: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "Pair x -` A \<in> sets M2"
 proof -
@@ -144,8 +134,7 @@
 qed
 
 lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^isub>M M2)"
-  using measurable_comp[OF measurable_Pair1' measurable_pair_swap', of y M2 M1]
-  by (simp add: comp_def)
+  by (auto intro!: measurable_Pair)
 
 lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
 proof -
@@ -172,46 +161,115 @@
   unfolding Int_stable_def
   by safe (auto simp add: times_Int_times)
 
-lemma finite_measure_cut_measurable:
-  assumes "sigma_finite_measure M1" "finite_measure M2"
-  assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-  shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
+lemma (in finite_measure) finite_measure_cut_measurable:
+  assumes "Q \<in> sets (N \<Otimes>\<^isub>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_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 top show ?case
+    by (auto intro!: measurable_If simp: space_pair_measure)
+next
+  case (union F)
+  moreover then have "\<And>x. emeasure M (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
+    unfolding sets_pair_measure[symmetric]
+    by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1)
+  ultimately show ?case
+    by (auto simp: vimage_UN)
+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>\<^isub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")
 proof -
-  interpret M1: sigma_finite_measure M1 by fact
-  interpret M2: finite_measure M2 by fact
-  let ?\<Omega> = "space M1 \<times> space M2" and ?D = "{A\<in>sets (M1 \<Otimes>\<^isub>M M2). ?s A \<in> borel_measurable M1}"
-  note space_pair_measure[simp]
-  interpret dynkin_system ?\<Omega> ?D
-  proof (intro dynkin_systemI)
-    fix A assume "A \<in> ?D" then show "A \<subseteq> ?\<Omega>"
-      using sets_into_space[of A "M1 \<Otimes>\<^isub>M M2"] by simp
-  next
-    from top show "?\<Omega> \<in> ?D"
-      by (auto simp add: if_distrib intro!: measurable_If)
-  next
-    fix A assume "A \<in> ?D"
-    with sets_into_space have "\<And>x. emeasure M2 (Pair x -` (?\<Omega> - A)) =
-        (if x \<in> space M1 then emeasure M2 (space M2) - ?s A x else 0)"
-      by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1)
-    with `A \<in> ?D` top show "?\<Omega> - A \<in> ?D"
-      by (auto intro!: measurable_If)
-  next
-    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> ?D"
-    moreover then have "\<And>x. emeasure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
-      by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1)
-    ultimately show "(\<Union>i. F i) \<in> ?D"
-      by (auto simp: vimage_UN intro!: borel_measurable_psuminf)
+  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_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_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 `Q \<in> sets (N \<Otimes>\<^isub>M M)` 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 `disjoint_family F`]) auto
+    qed
+    also have "(\<Union>i. ?C x i) = Pair x -` Q"
+      using F sets_into_space[OF `Q \<in> sets (N \<Otimes>\<^isub>M M)`]
+      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 `Q \<in> sets (N \<Otimes>\<^isub>M M)` F_sets
+    by auto
+qed
+
+lemma (in sigma_finite_measure) emeasure_pair_measure:
+  assumes "X \<in> sets (N \<Otimes>\<^isub>M M)"
+  shows "emeasure (N \<Otimes>\<^isub>M M) X = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ 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>\<^isub>M M)) ?\<mu>"
+    by (auto simp: positive_def positive_integral_positive)
+  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>\<^isub>M M)) ?\<mu>"
+  proof (rule countably_additiveI)
+    fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^isub>M M)" "disjoint_family F"
+    from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^isub>M M)" "(\<Union>i. F i) \<in> sets (N \<Otimes>\<^isub>M M)" by auto
+    moreover from F have "\<And>i. (\<lambda>x. emeasure M (Pair x -` F i)) \<in> borel_measurable N"
+      by (intro measurable_emeasure_Pair) 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: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1
+               intro!: positive_integral_cong positive_integral_indicator[symmetric])
   qed
-  let ?G = "{a \<times> b | a b. a \<in> sets M1 \<and> b \<in> sets M2}"
-  have "sigma_sets ?\<Omega> ?G = ?D"
-  proof (rule dynkin_lemma)
-    show "?G \<subseteq> ?D"
-      by (auto simp: if_distrib Int_def[symmetric] intro!: measurable_If)
-  qed (auto simp: sets_pair_measure  Int_stable_pair_measure_generator)
-  with `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` have "Q \<in> ?D"
-    by (simp add: sets_pair_measure[symmetric])
-  then show "?s Q \<in> borel_measurable M1" by simp
+  show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)"
+    using space_closed[of N] space_closed[of M] by auto
+qed fact
+
+lemma (in sigma_finite_measure) emeasure_pair_measure_alt:
+  assumes X: "X \<in> sets (N \<Otimes>\<^isub>M M)"
+  shows "emeasure (N  \<Otimes>\<^isub>M M) X = (\<integral>\<^isup>+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!: positive_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>\<^isub>M M) (A \<times> B) = emeasure N A * emeasure M B"
+proof -
+  have "emeasure (N \<Otimes>\<^isub>M M) (A \<times> B) = (\<integral>\<^isup>+x. emeasure M B * indicator A x \<partial>N)"
+    using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt)
+  also have "\<dots> = emeasure M B * emeasure N A"
+    using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator)
+  finally show ?thesis
+    by (simp add: ac_simps)
 qed
 
 subsection {* Binary products of $\sigma$-finite emeasure spaces *}
@@ -219,114 +277,21 @@
 locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
   for M1 :: "'a measure" and M2 :: "'b measure"
 
-lemma sets_pair_measure_cong[cong]:
-  "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^isub>M M2) = sets (M1' \<Otimes>\<^isub>M M2')"
-  unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
-
 lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
-  assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
-proof -
-  from M2.sigma_finite_disjoint guess F . note F = this
-  then have F_sets: "\<And>i. F i \<in> sets M2" by auto
-  have M1: "sigma_finite_measure M1" ..
-  let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"
-  { fix i
-    have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
-      using F sets_into_space by auto
-    let ?R = "density M2 (indicator (F i))"
-    have "(\<lambda>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))) \<in> borel_measurable M1"
-    proof (intro finite_measure_cut_measurable[OF M1])
-      show "finite_measure ?R"
-        using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)
-      show "(space M1 \<times> space ?R) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R)"
-        using Q by (simp add: Int)
-    qed
-    moreover have "\<And>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))
-        = emeasure M2 (F i \<inter> Pair x -` (space M1 \<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 M1 \<times> space ?R \<inter> Q) = ?C x i"
-      using sets_into_space[OF Q] by (auto simp: space_pair_measure)
-    ultimately have "(\<lambda>x. emeasure M2 (?C x i)) \<in> borel_measurable M1"
-      by simp }
-  moreover
-  { fix x
-    have "(\<Sum>i. emeasure M2 (?C x i)) = emeasure M2 (\<Union>i. ?C x i)"
-    proof (intro suminf_emeasure)
-      show "range (?C x) \<subseteq> sets M2"
-        using F `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` 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 `disjoint_family F`]) auto
-    qed
-    also have "(\<Union>i. ?C x i) = Pair x -` Q"
-      using F sets_into_space[OF `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
-      by (auto simp: space_pair_measure)
-    finally have "emeasure M2 (Pair x -` Q) = (\<Sum>i. emeasure M2 (?C x i))"
-      by simp }
-  ultimately show ?thesis using `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` F_sets
-    by auto
-qed
+  "Q \<in> sets (M1 \<Otimes>\<^isub>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>\<^isub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
 proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
   have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
     using Q measurable_pair_swap' by (auto intro: measurable_sets)
-  note Q.measurable_emeasure_Pair1[OF this]
+  note M1.measurable_emeasure_Pair[OF this]
   moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1)) = (\<lambda>x. (x, y)) -` Q"
     using Q[THEN sets_into_space] by (auto simp: space_pair_measure)
   ultimately show ?thesis by simp
 qed
 
-lemma (in pair_sigma_finite) emeasure_pair_measure:
-  assumes "X \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-  shows "emeasure (M1 \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator X (x, y) \<partial>M2 \<partial>M1)" (is "_ = ?\<mu> X")
-proof (rule emeasure_measure_of[OF pair_measure_def])
-  show "positive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>"
-    by (auto simp: positive_def positive_integral_positive)
-  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 (M1 \<Otimes>\<^isub>M M2)) ?\<mu>"
-  proof (rule countably_additiveI)
-    fix F :: "nat \<Rightarrow> ('a \<times> 'b) set" assume F: "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" "disjoint_family F"
-    from F have *: "\<And>i. F i \<in> sets (M1 \<Otimes>\<^isub>M M2)" "(\<Union>i. F i) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by auto
-    moreover from F have "\<And>i. (\<lambda>x. emeasure M2 (Pair x -` F i)) \<in> borel_measurable M1"
-      by (intro measurable_emeasure_Pair1) 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 M2"
-      using F by (auto simp: sets_Pair1)
-    ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"
-      by (auto simp add: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1
-               intro!: positive_integral_cong positive_integral_indicator[symmetric])
-  qed
-  show "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2} \<subseteq> Pow (space M1 \<times> space M2)"
-    using space_closed[of M1] space_closed[of M2] by auto
-qed fact
-
-lemma (in pair_sigma_finite) emeasure_pair_measure_alt:
-  assumes X: "X \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-  shows "emeasure (M1  \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+x. emeasure M2 (Pair x -` X) \<partial>M1)"
-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!: positive_integral_cong simp: emeasure_pair_measure sets_Pair1)
-qed
-
-lemma (in pair_sigma_finite) emeasure_pair_measure_Times:
-  assumes A: "A \<in> sets M1" and B: "B \<in> sets M2"
-  shows "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = emeasure M1 A * emeasure M2 B"
-proof -
-  have "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+x. emeasure M2 B * indicator A x \<partial>M1)"
-    using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt)
-  also have "\<dots> = emeasure M2 B * emeasure M1 A"
-    using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator)
-  finally show ?thesis
-    by (simp add: ac_simps)
-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>
@@ -366,7 +331,7 @@
   qed
 qed
 
-sublocale pair_sigma_finite \<subseteq> sigma_finite_measure "M1 \<Otimes>\<^isub>M M2"
+sublocale pair_sigma_finite \<subseteq> P: sigma_finite_measure "M1 \<Otimes>\<^isub>M M2"
 proof
   from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
   show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
@@ -396,7 +361,6 @@
 lemma (in pair_sigma_finite) distr_pair_swap:
   "M1 \<Otimes>\<^isub>M M2 = distr (M2 \<Otimes>\<^isub>M M1) (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
 proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
   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
@@ -408,7 +372,7 @@
     then show "sets ?D = sigma_sets (space ?P) ?E"
       by simp
   next
-    show "range F \<subseteq> ?E" "incseq F" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
+    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"
@@ -416,7 +380,7 @@
     have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^isub>M M1) = B \<times> A"
       using sets_into_space[OF A] sets_into_space[OF B] by (auto simp: space_pair_measure)
     with A B show "emeasure (M1 \<Otimes>\<^isub>M M2) X = emeasure ?D X"
-      by (simp add: emeasure_pair_measure_Times Q.emeasure_pair_measure_Times emeasure_distr
+      by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr
                     measurable_pair_swap' ac_simps)
   qed
 qed
@@ -426,138 +390,14 @@
   shows "emeasure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
     (is "_ = ?\<nu> A")
 proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
   have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))) = (\<lambda>x. (x, y)) -` A"
     using 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']
-                 Q.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
+                 M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
 qed
 
-section "Fubinis theorem"
-
-lemma (in pair_sigma_finite) simple_function_cut:
-  assumes f: "simple_function (M1 \<Otimes>\<^isub>M M2) f" "\<And>x. 0 \<le> f x"
-  shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
-proof -
-  have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
-    using f(1) by (rule borel_measurable_simple_function)
-  let ?F = "\<lambda>z. f -` {z} \<inter> space (M1 \<Otimes>\<^isub>M M2)"
-  let ?F' = "\<lambda>x z. Pair x -` ?F z"
-  { fix x assume "x \<in> space M1"
-    have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
-      by (auto simp: indicator_def)
-    have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space (M1 \<Otimes>\<^isub>M M2)" using `x \<in> space M1`
-      by (simp add: space_pair_measure)
-    moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
-      by (rule sets_Pair1[OF measurable_sets]) auto
-    ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
-      apply (rule_tac simple_function_cong[THEN iffD2, OF _])
-      apply (rule simple_function_indicator_representation[OF f(1)])
-      using `x \<in> space M1` by auto }
-  note M2_sf = this
-  { fix x assume x: "x \<in> space M1"
-    then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space (M1 \<Otimes>\<^isub>M M2). z * emeasure M2 (?F' x z))"
-      unfolding positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
-      unfolding simple_integral_def
-    proof (safe intro!: setsum_mono_zero_cong_left)
-      from f(1) show "finite (f ` space (M1 \<Otimes>\<^isub>M M2))" by (rule simple_functionD)
-    next
-      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space (M1 \<Otimes>\<^isub>M M2)"
-        using `x \<in> space M1` by (auto simp: space_pair_measure)
-    next
-      fix x' y assume "(x', y) \<in> space (M1 \<Otimes>\<^isub>M M2)"
-        "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
-      then have *: "?F' x (f (x', y)) = {}"
-        by (force simp: space_pair_measure)
-      show  "f (x', y) * emeasure M2 (?F' x (f (x', y))) = 0"
-        unfolding * by simp
-    qed (simp add: vimage_compose[symmetric] comp_def
-                   space_pair_measure) }
-  note eq = this
-  moreover have "\<And>z. ?F z \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-    by (auto intro!: f_borel borel_measurable_vimage)
-  moreover then have "\<And>z. (\<lambda>x. emeasure M2 (?F' x z)) \<in> borel_measurable M1"
-    by (auto intro!: measurable_emeasure_Pair1 simp del: vimage_Int)
-  moreover have *: "\<And>i x. 0 \<le> emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))"
-    using f(1)[THEN simple_functionD(2)] f(2) by (intro emeasure_nonneg)
-  moreover { fix i assume "i \<in> f`space (M1 \<Otimes>\<^isub>M M2)"
-    with * have "\<And>x. 0 \<le> i * emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))"
-      using f(2) by auto }
-  ultimately
-  show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" using f(2)
-    by (auto simp del: vimage_Int cong: measurable_cong intro!: setsum_cong
-             simp add: positive_integral_setsum simple_integral_def
-                       positive_integral_cmult
-                       positive_integral_cong[OF eq]
-                       positive_integral_eq_simple_integral[OF f]
-                       emeasure_pair_measure_alt[symmetric])
-qed
-
-lemma (in pair_sigma_finite) positive_integral_fst_measurable:
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
-  shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-      (is "?C f \<in> borel_measurable M1")
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
-proof -
-  from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
-  then have F_borel: "\<And>i. F i \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
-    by (auto intro: borel_measurable_simple_function)
-  note sf = simple_function_cut[OF F(1,5)]
-  then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
-    using F(1) by auto
-  moreover
-  { fix x assume "x \<in> space M1"
-    from F measurable_Pair2[OF F_borel `x \<in> space M1`]
-    have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
-      by (intro positive_integral_monotone_convergence_SUP)
-         (auto simp: incseq_Suc_iff le_fun_def)
-    then have "(SUP i. ?C (F i) x) = ?C f x"
-      unfolding F(4) positive_integral_max_0 by simp }
-  note SUPR_C = this
-  ultimately show "?C f \<in> borel_measurable M1"
-    by (simp cong: measurable_cong)
-  have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>(M1 \<Otimes>\<^isub>M M2)) = (SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i))"
-    using F_borel F
-    by (intro positive_integral_monotone_convergence_SUP) auto
-  also have "(SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
-    unfolding sf(2) by simp
-  also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
-    by (intro positive_integral_monotone_convergence_SUP[symmetric])
-       (auto intro!: positive_integral_mono positive_integral_positive
-             simp: incseq_Suc_iff le_fun_def)
-  also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
-    using F_borel F(2,5)
-    by (auto intro!: positive_integral_cong positive_integral_monotone_convergence_SUP[symmetric] measurable_Pair2
-             simp: incseq_Suc_iff le_fun_def)
-  finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
-    using F by (simp add: positive_integral_max_0)
-qed
-
-lemma (in pair_sigma_finite) positive_integral_snd_measurable:
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
-  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  note measurable_pair_swap[OF f]
-  from Q.positive_integral_fst_measurable[OF this]
-  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1))"
-    by simp
-  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
-    by (subst distr_pair_swap)
-       (auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong)
-  finally show ?thesis .
-qed
-
-lemma (in pair_sigma_finite) Fubini:
-  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
-  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
-  unfolding positive_integral_snd_measurable[OF assms]
-  unfolding positive_integral_fst_measurable[OF assms] ..
-
 lemma (in pair_sigma_finite) AE_pair:
   assumes "AE x in (M1 \<Otimes>\<^isub>M M2). Q x"
   shows "AE x in M1. (AE y in M2. Q (x, y))"
@@ -568,7 +408,7 @@
   proof (rule AE_I)
     from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
     show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
-      by (auto simp: emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg)
+      by (auto simp: M2.emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg)
     show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
       by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)
     { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
@@ -593,7 +433,7 @@
     by (rule sets_Collect) fact
   then have "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} =
       (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
-    by (simp add: emeasure_pair_measure)
+    by (simp add: M2.emeasure_pair_measure)
   also have "\<dots> = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. 0 \<partial>M2 \<partial>M1)"
     using ae
     apply (safe intro!: positive_integral_cong_AE)
@@ -609,24 +449,6 @@
     (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^isub>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 AE_distr_iff:
-  assumes f: "f \<in> measurable M N" and P: "{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])
-  from P show "{x \<in> space (distr M N f). \<not> P x} \<in> sets (distr M N f)"
-    by (auto intro!: sets_Collect_neg)
-  moreover
-  have "f -` {x \<in> space N. P x} \<inter> space M = {x \<in> space M. P (f x)}"
-    using f by (auto dest: measurable_space)
-  then show "{x \<in> space M. \<not> P (f x)} \<in> sets M"
-    using measurable_sets[OF f P] by (auto intro!: sets_Collect_neg)
-  moreover have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
-    using f by (auto dest: measurable_space)
-  ultimately 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)"
-    using f by (simp add: emeasure_distr)
-qed
-
 lemma (in pair_sigma_finite) AE_commute:
   assumes P: "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>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)"
@@ -648,6 +470,82 @@
     done
 qed
 
+section "Fubinis theorem"
+
+lemma measurable_compose_Pair1:
+  "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^isub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
+  by (rule measurable_compose[OF measurable_Pair]) auto
+
+lemma (in pair_sigma_finite) borel_measurable_positive_integral_fst:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" "\<And>x. 0 \<le> f x"
+  shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<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 M2 \<Longrightarrow> u (w, x) = v (w, x)"
+    by (auto simp: space_pair_measure)
+  show ?case
+    apply (subst measurable_cong)
+    apply (rule positive_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 M2 (Pair x -` Q) = \<integral>\<^isup>+ y. indicator Q (x, y) \<partial>M2"
+    by (simp add: sets_Pair1[OF set])
+  from this M2.measurable_emeasure_Pair[OF set] show ?case
+    by (rule measurable_cong[THEN iffD1])
+qed (simp_all add: positive_integral_add positive_integral_cmult measurable_compose_Pair1
+                   positive_integral_monotone_convergence_SUP incseq_def le_fun_def
+              cong: measurable_cong)
+
+lemma (in pair_sigma_finite) positive_integral_fst:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" "\<And>x. 0 \<le> f x"
+  shows "(\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2 \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" (is "?I f = _")
+using f proof induct
+  case (cong u v)
+  moreover then have "?I u = ?I v"
+    by (intro positive_integral_cong) (auto simp: space_pair_measure)
+  ultimately show ?case
+    by (simp cong: positive_integral_cong)
+qed (simp_all add: M2.emeasure_pair_measure positive_integral_cmult positive_integral_add
+                   positive_integral_monotone_convergence_SUP
+                   measurable_compose_Pair1 positive_integral_positive
+                   borel_measurable_positive_integral_fst positive_integral_mono incseq_def le_fun_def
+              cong: positive_integral_cong)
+
+lemma (in pair_sigma_finite) positive_integral_fst_measurable:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
+      (is "?C f \<in> borel_measurable M1")
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
+  using f
+    borel_measurable_positive_integral_fst[of "\<lambda>x. max 0 (f x)"]
+    positive_integral_fst[of "\<lambda>x. max 0 (f x)"]
+  unfolding positive_integral_max_0 by auto
+
+lemma (in pair_sigma_finite) positive_integral_snd_measurable:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  note measurable_pair_swap[OF f]
+  from Q.positive_integral_fst_measurable[OF this]
+  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1))"
+    by simp
+  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
+    by (subst distr_pair_swap)
+       (auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong)
+  finally show ?thesis .
+qed
+
+lemma (in pair_sigma_finite) Fubini:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
+  unfolding positive_integral_snd_measurable[OF assms]
+  unfolding positive_integral_fst_measurable[OF assms] ..
+
 lemma (in pair_sigma_finite) integrable_product_swap:
   assumes "integrable (M1 \<Otimes>\<^isub>M M2) f"
   shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
@@ -817,7 +715,7 @@
     by (intro finite_subset[OF _ B]) auto
   have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)
   show "emeasure ?P X = emeasure ?C X"
-    apply (subst P.emeasure_pair_measure_alt[OF X])
+    apply (subst B.emeasure_pair_measure_alt[OF X])
     apply (subst emeasure_count_space)
     using X_subset apply auto []
     apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)
@@ -861,7 +759,7 @@
     by simp
 
   show "emeasure ?L A = emeasure ?R A"
-    apply (subst L.emeasure_pair_measure[OF A])
+    apply (subst D2.emeasure_pair_measure[OF A])
     apply (subst emeasure_density)
         using f_fst g_snd apply (simp add: split_beta')
       using A apply simp
@@ -924,7 +822,7 @@
     by (auto simp: measurable_pair_iff)
   fix A assume A: "A \<in> sets ?P"
   then have "emeasure ?P A = (\<integral>\<^isup>+x. emeasure (distr N T g) (Pair x -` A) \<partial>distr M S f)"
-    by (rule ST.emeasure_pair_measure_alt)
+    by (rule T.emeasure_pair_measure_alt)
   also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair x -` A) \<inter> space N) \<partial>distr M S f)"
     using g A by (simp add: sets_Pair1 emeasure_distr)
   also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair (f x) -` A) \<inter> space N) \<partial>M)"
@@ -933,7 +831,7 @@
   also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (Pair x -` ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))) \<partial>M)"
     by (intro positive_integral_cong arg_cong2[where f=emeasure]) (auto simp: space_pair_measure)
   also have "\<dots> = emeasure (M \<Otimes>\<^isub>M N) ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))"
-    using fg by (intro MN.emeasure_pair_measure_alt[symmetric] measurable_sets[OF _ A])
+    using fg by (intro N.emeasure_pair_measure_alt[symmetric] measurable_sets[OF _ A])
                 (auto cong: measurable_cong')
   also have "\<dots> = emeasure ?D A"
     using fg A by (subst emeasure_distr) auto
--- a/src/HOL/Probability/Borel_Space.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Borel_Space.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -632,7 +632,7 @@
   "(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_component:
+lemma borel_measurable_euclidean_component':
   "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel"
 proof (rule borel_measurableI)
   fix S::"real set" assume "open S"
@@ -641,13 +641,18 @@
     by (auto intro: borel_open)
 qed
 
+lemma borel_measurable_euclidean_component:
+  fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
+  assumes f: "f \<in> borel_measurable M"
+  shows "(\<lambda>x. f x $$ i) \<in> borel_measurable M"
+  using measurable_comp[OF f borel_measurable_euclidean_component'] by (simp add: comp_def)
+
 lemma borel_measurable_euclidean_space:
   fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)"
 proof safe
   fix i assume "f \<in> borel_measurable M"
   then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M"
-    using measurable_comp[of f _ _ "\<lambda>x. x $$ i", unfolded comp_def]
     by (auto intro: borel_measurable_euclidean_component)
 next
   assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
@@ -657,6 +662,144 @@
 
 subsection "Borel measurable operators"
 
+lemma borel_measurable_continuous_on1:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
+  assumes "continuous_on UNIV f"
+  shows "f \<in> borel_measurable borel"
+  apply(rule borel_measurableI)
+  using continuous_open_preimage[OF assms] unfolding vimage_def by auto
+
+lemma borel_measurable_continuous_on:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::topological_space"
+  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_open':
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
+  assumes cont: "continuous_on A f" "open A"
+  shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _")
+proof (rule borel_measurableI)
+  fix S :: "'b set" assume "open S"
+  then have "open {x\<in>A. f x \<in> S}"
+    by (intro continuous_open_preimage[OF cont]) auto
+  then have *: "{x\<in>A. f x \<in> S} \<in> sets borel" by auto
+  have "?f -` S \<inter> space borel = 
+    {x\<in>A. f x \<in> S} \<union> (if c \<in> S then space borel - A else {})"
+    by (auto split: split_if_asm)
+  also have "\<dots> \<in> sets borel"
+    using * `open A` by (auto simp del: space_borel intro!: Un)
+  finally show "?f -` S \<inter> space borel \<in> sets borel" .
+qed
+
+lemma borel_measurable_continuous_on_open:
+  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
+  assumes cont: "continuous_on A f" "open A"
+  assumes g: "g \<in> borel_measurable M"
+  shows "(\<lambda>x. if g x \<in> A then f (g x) else c) \<in> borel_measurable M"
+  using measurable_comp[OF g borel_measurable_continuous_on_open'[OF cont], of c]
+  by (simp add: comp_def)
+
+lemma borel_measurable_uminus[simp, intro]:
+  fixes g :: "'a \<Rightarrow> real"
+  assumes g: "g \<in> borel_measurable M"
+  shows "(\<lambda>x. - g x) \<in> borel_measurable M"
+  by (rule borel_measurable_continuous_on[OF _ g]) (auto intro: continuous_on_minus continuous_on_id)
+
+lemma euclidean_component_prod:
+  fixes x :: "'a :: euclidean_space \<times> 'b :: euclidean_space"
+  shows "x $$ i = (if i < DIM('a) then fst x $$ i else snd x $$ (i - DIM('a)))"
+  unfolding euclidean_component_def basis_prod_def inner_prod_def by auto
+
+lemma borel_measurable_Pair[simp, intro]:
+  fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
+  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"
+proof (intro borel_measurable_iff_halfspace_le[THEN iffD2] allI impI)
+  fix i and a :: real assume i: "i < DIM('b \<times> 'c)"
+  have [simp]: "\<And>P A B C. {w. (P \<longrightarrow> A w \<and> B w) \<and> (\<not> P \<longrightarrow> A w \<and> C w)} = 
+    {w. A w \<and> (P \<longrightarrow> B w) \<and> (\<not> P \<longrightarrow> C w)}" by auto
+  from i f g show "{w \<in> space M. (f w, g w) $$ i \<le> a} \<in> sets M"
+    by (auto simp: euclidean_component_prod intro!: sets_Collect borel_measurable_euclidean_component)
+qed
+
+lemma continuous_on_fst: "continuous_on UNIV fst"
+proof -
+  have [simp]: "range fst = UNIV" by (auto simp: image_iff)
+  show ?thesis
+    using closed_vimage_fst
+    by (auto simp: continuous_on_closed closed_closedin vimage_def)
+qed
+
+lemma continuous_on_snd: "continuous_on UNIV snd"
+proof -
+  have [simp]: "range snd = UNIV" by (auto simp: image_iff)
+  show ?thesis
+    using closed_vimage_snd
+    by (auto simp: continuous_on_closed closed_closedin vimage_def)
+qed
+
+lemma borel_measurable_continuous_Pair:
+  fixes f :: "'a \<Rightarrow> 'b::ordered_euclidean_space" and g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
+  assumes [simp]: "f \<in> borel_measurable M"
+  assumes [simp]: "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
+
+lemma borel_measurable_add[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
+  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)
+     (auto intro: continuous_on_fst continuous_on_snd continuous_on_add)
+
+lemma borel_measurable_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
+  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_diff[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  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"
+  unfolding diff_minus using assms by fast
+
+lemma borel_measurable_times[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  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)
+     (auto intro: continuous_on_fst continuous_on_snd continuous_on_mult)
+
+lemma continuous_on_dist:
+  fixes f :: "'a :: t2_space \<Rightarrow> 'b :: metric_space"
+  shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. dist (f x) (g x))"
+  unfolding continuous_on_eq_continuous_within by (auto simp: continuous_dist)
+
+lemma borel_measurable_dist[simp, intro]:
+  fixes g f :: "'a \<Rightarrow> 'b::ordered_euclidean_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_on_dist continuous_on_fst continuous_on_snd)
+  
 lemma affine_borel_measurable_vector:
   fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
   assumes "f \<in> borel_measurable M"
@@ -683,116 +826,6 @@
   shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
   using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)
 
-lemma borel_measurable_add[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  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"
-proof -
-  have 1: "\<And>a. {w\<in>space M. a \<le> f w + g w} = {w \<in> space M. a + g w * -1 \<le> f w}"
-    by auto
-  have "\<And>a. (\<lambda>w. a + (g w) * -1) \<in> borel_measurable M"
-    by (rule affine_borel_measurable [OF g])
-  then have "\<And>a. {w \<in> space M. (\<lambda>w. a + (g w) * -1)(w) \<le> f w} \<in> sets M" using f
-    by auto
-  then have "\<And>a. {w \<in> space M. a \<le> f w + g w} \<in> sets M"
-    by (simp add: 1)
-  then show ?thesis
-    by (simp add: borel_measurable_iff_ge)
-qed
-
-lemma borel_measurable_setsum[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
-  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_square:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable M"
-  shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M"
-proof -
-  {
-    fix a
-    have "{w \<in> space M. (f w)\<twosuperior> \<le> a} \<in> sets M"
-    proof (cases rule: linorder_cases [of a 0])
-      case less
-      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} = {}"
-        by auto (metis less order_le_less_trans power2_less_0)
-      also have "... \<in> sets M"
-        by (rule empty_sets)
-      finally show ?thesis .
-    next
-      case equal
-      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
-             {w \<in> space M. f w \<le> 0} \<inter> {w \<in> space M. 0 \<le> f w}"
-        by auto
-      also have "... \<in> sets M"
-        apply (insert f)
-        apply (rule Int)
-        apply (simp add: borel_measurable_iff_le)
-        apply (simp add: borel_measurable_iff_ge)
-        done
-      finally show ?thesis .
-    next
-      case greater
-      have "\<forall>x. (f x ^ 2 \<le> sqrt a ^ 2) = (- sqrt a  \<le> f x & f x \<le> sqrt a)"
-        by (metis abs_le_interval_iff abs_of_pos greater real_sqrt_abs
-                  real_sqrt_le_iff real_sqrt_power)
-      hence "{w \<in> space M. (f w)\<twosuperior> \<le> a} =
-             {w \<in> space M. -(sqrt a) \<le> f w} \<inter> {w \<in> space M. f w \<le> sqrt a}"
-        using greater by auto
-      also have "... \<in> sets M"
-        apply (insert f)
-        apply (rule Int)
-        apply (simp add: borel_measurable_iff_ge)
-        apply (simp add: borel_measurable_iff_le)
-        done
-      finally show ?thesis .
-    qed
-  }
-  thus ?thesis by (auto simp add: borel_measurable_iff_le)
-qed
-
-lemma times_eq_sum_squares:
-   fixes x::real
-   shows"x*y = ((x+y)^2)/4 - ((x-y)^ 2)/4"
-by (simp add: power2_eq_square ring_distribs diff_divide_distrib [symmetric])
-
-lemma borel_measurable_uminus[simp, intro]:
-  fixes g :: "'a \<Rightarrow> real"
-  assumes g: "g \<in> borel_measurable M"
-  shows "(\<lambda>x. - g x) \<in> borel_measurable M"
-proof -
-  have "(\<lambda>x. - g x) = (\<lambda>x. 0 + (g x) * -1)"
-    by simp
-  also have "... \<in> borel_measurable M"
-    by (fast intro: affine_borel_measurable g)
-  finally show ?thesis .
-qed
-
-lemma borel_measurable_times[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  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"
-proof -
-  have 1: "(\<lambda>x. 0 + (f x + g x)\<twosuperior> * inverse 4) \<in> borel_measurable M"
-    using assms by (fast intro: affine_borel_measurable borel_measurable_square)
-  have "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) =
-        (\<lambda>x. 0 + ((f x + -g x) ^ 2 * inverse -4))"
-    by (simp add: minus_divide_right)
-  also have "... \<in> borel_measurable M"
-    using f g by (fast intro: affine_borel_measurable borel_measurable_square f g)
-  finally have 2: "(\<lambda>x. -((f x + -g x) ^ 2 * inverse 4)) \<in> borel_measurable M" .
-  show ?thesis
-    apply (simp add: times_eq_sum_squares diff_minus)
-    using 1 2 by simp
-qed
-
 lemma borel_measurable_setprod[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
@@ -802,26 +835,17 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma borel_measurable_diff[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  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"
-  unfolding diff_minus using assms by fast
-
 lemma borel_measurable_inverse[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
-  assumes "f \<in> borel_measurable M"
+  assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. inverse (f x)) \<in> borel_measurable M"
-  unfolding borel_measurable_iff_ge unfolding inverse_eq_divide
-proof safe
-  fix a :: real
-  have *: "{w \<in> space M. a \<le> 1 / f w} =
-      ({w \<in> space M. 0 < f w} \<inter> {w \<in> space M. a * f w \<le> 1}) \<union>
-      ({w \<in> space M. f w < 0} \<inter> {w \<in> space M. 1 \<le> a * f w}) \<union>
-      ({w \<in> space M. f w = 0} \<inter> {w \<in> space M. a \<le> 0})" by (auto simp: le_divide_eq)
-  show "{w \<in> space M. a \<le> 1 / f w} \<in> sets M" using assms unfolding *
-    by (auto intro!: Int Un)
+proof -
+  have *: "\<And>x::real. inverse x = (if x \<in> UNIV - {0} then inverse x else 0)" by auto
+  show ?thesis
+    apply (subst *)
+    apply (rule borel_measurable_continuous_on_open)
+    apply (auto intro!: f continuous_on_inverse continuous_on_id)
+    done
 qed
 
 lemma borel_measurable_divide[simp, intro]:
@@ -837,30 +861,14 @@
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
-  unfolding borel_measurable_iff_le
-proof safe
-  fix a
-  have "{x \<in> space M. max (g x) (f x) \<le> a} =
-    {x \<in> space M. g x \<le> a} \<inter> {x \<in> space M. f x \<le> a}" by auto
-  thus "{x \<in> space M. max (g x) (f x) \<le> a} \<in> sets M"
-    using assms unfolding borel_measurable_iff_le
-    by (auto intro!: Int)
-qed
+  unfolding max_def by (auto intro!: assms measurable_If)
 
 lemma borel_measurable_min[intro, simp]:
   fixes f g :: "'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
-  unfolding borel_measurable_iff_ge
-proof safe
-  fix a
-  have "{x \<in> space M. a \<le> min (g x) (f x)} =
-    {x \<in> space M. a \<le> g x} \<inter> {x \<in> space M. a \<le> f x}" by auto
-  thus "{x \<in> space M. a \<le> min (g x) (f x)} \<in> sets M"
-    using assms unfolding borel_measurable_iff_ge
-    by (auto intro!: Int)
-qed
+  unfolding min_def by (auto intro!: assms measurable_If)
 
 lemma borel_measurable_abs[simp, intro]:
   assumes "f \<in> borel_measurable M"
@@ -872,76 +880,50 @@
 
 lemma borel_measurable_nth[simp, intro]:
   "(\<lambda>x::real^'n. x $ i) \<in> borel_measurable borel"
-  using borel_measurable_euclidean_component
+  using borel_measurable_euclidean_component'
   unfolding nth_conv_component by auto
 
-lemma borel_measurable_continuous_on1:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
-  assumes "continuous_on UNIV f"
-  shows "f \<in> borel_measurable borel"
-  apply(rule borel_measurableI)
-  using continuous_open_preimage[OF assms] unfolding vimage_def by auto
-
-lemma borel_measurable_continuous_on:
-  fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
-  assumes cont: "continuous_on A f" "open A"
-  shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _")
-proof (rule borel_measurableI)
-  fix S :: "'b set" assume "open S"
-  then have "open {x\<in>A. f x \<in> S}"
-    by (intro continuous_open_preimage[OF cont]) auto
-  then have *: "{x\<in>A. f x \<in> S} \<in> sets borel" by auto
-  have "?f -` S \<inter> space borel = 
-    {x\<in>A. f x \<in> S} \<union> (if c \<in> S then space borel - A else {})"
-    by (auto split: split_if_asm)
-  also have "\<dots> \<in> sets borel"
-    using * `open A` by (auto simp del: space_borel intro!: Un)
-  finally show "?f -` S \<inter> space borel \<in> sets borel" .
-qed
-
 lemma convex_measurable:
   fixes a b :: real
   assumes X: "X \<in> borel_measurable M" "X ` space M \<subseteq> { a <..< b}"
   assumes q: "convex_on { a <..< b} q"
-  shows "q \<circ> X \<in> borel_measurable M"
+  shows "(\<lambda>x. q (X x)) \<in> borel_measurable M"
 proof -
-  have "(\<lambda>x. if x \<in> {a <..< b} then q x else 0) \<in> borel_measurable borel"
-  proof (rule borel_measurable_continuous_on)
+  have "(\<lambda>x. if X x \<in> {a <..< b} then q (X x) else 0) \<in> borel_measurable M" (is "?qX")
+  proof (rule borel_measurable_continuous_on_open[OF _ _ X(1)])
     show "open {a<..<b}" by auto
     from this q show "continuous_on {a<..<b} q"
       by (rule convex_on_continuous)
   qed
-  then have "(\<lambda>x. if x \<in> {a <..< b} then q x else 0) \<circ> X \<in> borel_measurable M" (is ?qX)
-    using X by (intro measurable_comp) auto
-  moreover have "?qX \<longleftrightarrow> q \<circ> X \<in> borel_measurable M"
+  moreover have "?qX \<longleftrightarrow> (\<lambda>x. q (X x)) \<in> borel_measurable M"
     using X by (intro measurable_cong) auto
   ultimately show ?thesis by simp
 qed
 
-lemma borel_measurable_borel_log: assumes "1 < b" shows "log b \<in> borel_measurable borel"
+lemma borel_measurable_ln[simp,intro]:
+  assumes f: "f \<in> borel_measurable M"
+  shows "(\<lambda>x. ln (f x)) \<in> borel_measurable M"
 proof -
   { fix x :: real assume x: "x \<le> 0"
     { fix x::real assume "x \<le> 0" then have "\<And>u. exp u = x \<longleftrightarrow> False" by auto }
-    from this[of x] x this[of 0] have "log b 0 = log b x"
-      by (auto simp: ln_def log_def) }
-  note log_imp = this
-  have "(\<lambda>x. if x \<in> {0<..} then log b x else log b 0) \<in> borel_measurable borel"
-  proof (rule borel_measurable_continuous_on)
-    show "continuous_on {0<..} (log b)"
-      by (auto intro!: continuous_at_imp_continuous_on DERIV_log DERIV_isCont
+    from this[of x] x this[of 0] have "ln 0 = ln x"
+      by (auto simp: ln_def) }
+  note ln_imp = this
+  have "(\<lambda>x. if f x \<in> {0<..} then ln (f x) else ln 0) \<in> borel_measurable M"
+  proof (rule borel_measurable_continuous_on_open[OF _ _ f])
+    show "continuous_on {0<..} ln"
+      by (auto intro!: continuous_at_imp_continuous_on DERIV_ln DERIV_isCont
                simp: continuous_isCont[symmetric])
     show "open ({0<..}::real set)" by auto
   qed
-  also have "(\<lambda>x. if x \<in> {0<..} then log b x else log b 0) = log b"
-    by (simp add: fun_eq_iff not_less log_imp)
+  also have "(\<lambda>x. if x \<in> {0<..} then ln x else ln 0) = ln"
+    by (simp add: fun_eq_iff not_less ln_imp)
   finally show ?thesis .
 qed
 
 lemma borel_measurable_log[simp,intro]:
-  assumes f: "f \<in> borel_measurable M" and "1 < b"
-  shows "(\<lambda>x. log b (f x)) \<in> borel_measurable M"
-  using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]]
-  by (simp add: comp_def)
+  "f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. log b (f x)) \<in> borel_measurable M"
+  unfolding log_def by auto
 
 lemma borel_measurable_real_floor:
   "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
@@ -967,45 +949,91 @@
 
 subsection "Borel space on the extended reals"
 
-lemma borel_measurable_ereal_borel:
-  "ereal \<in> borel_measurable borel"
-proof (rule borel_measurableI)
-  fix X :: "ereal set" assume "open X"
-  then have "open (ereal -` X \<inter> space borel)"
-    by (simp add: open_ereal_vimage)
-  then show "ereal -` X \<inter> space borel \<in> sets borel" by auto
-qed
-
 lemma borel_measurable_ereal[simp, intro]:
   assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
-  using measurable_comp[OF f borel_measurable_ereal_borel] unfolding comp_def .
+  using continuous_on_ereal f by (rule borel_measurable_continuous_on)
 
-lemma borel_measurable_real_of_ereal_borel:
-  "(real :: ereal \<Rightarrow> real) \<in> borel_measurable borel"
-proof (rule borel_measurableI)
-  fix B :: "real set" assume "open B"
-  have *: "ereal -` real -` (B - {0}) = B - {0}" by auto
-  have open_real: "open (real -` (B - {0}) :: ereal set)"
-    unfolding open_ereal_def * using `open B` by auto
-  show "(real -` B \<inter> space borel :: ereal set) \<in> sets borel"
-  proof cases
-    assume "0 \<in> B"
-    then have *: "real -` B = real -` (B - {0}) \<union> {-\<infinity>, \<infinity>, 0::ereal}"
-      by (auto simp add: real_of_ereal_eq_0)
-    then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
-      using open_real by auto
-  next
-    assume "0 \<notin> B"
-    then have *: "(real -` B :: ereal set) = real -` (B - {0})"
-      by (auto simp add: real_of_ereal_eq_0)
-    then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
-      using open_real by auto
-  qed
+lemma borel_measurable_real_of_ereal[simp, intro]:
+  fixes f :: "'a \<Rightarrow> ereal" 
+  assumes f: "f \<in> borel_measurable M"
+  shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+proof -
+  have "(\<lambda>x. if f x \<in> UNIV - { \<infinity>, - \<infinity> } then real (f x) else 0) \<in> borel_measurable M"
+    using continuous_on_real
+    by (rule borel_measurable_continuous_on_open[OF _ _ f]) auto
+  also have "(\<lambda>x. if f x \<in> UNIV - { \<infinity>, - \<infinity> } then real (f x) else 0) = (\<lambda>x. real (f x))"
+    by auto
+  finally show ?thesis .
+qed
+
+lemma borel_measurable_ereal_cases:
+  fixes f :: "'a \<Rightarrow> ereal" 
+  assumes f: "f \<in> borel_measurable M"
+  assumes H: "(\<lambda>x. H (ereal (real (f x)))) \<in> borel_measurable M"
+  shows "(\<lambda>x. H (f x)) \<in> borel_measurable M"
+proof -
+  let ?F = "\<lambda>x. if x \<in> f -` {\<infinity>} then H \<infinity> else if x \<in> f -` {-\<infinity>} then H (-\<infinity>) else H (ereal (real (f x)))"
+  { fix x have "H (f x) = ?F x" by (cases "f x") auto }
+  moreover 
+  have "?F \<in> borel_measurable M"
+    by (intro measurable_If_set f measurable_sets[OF f] H) auto
+  ultimately
+  show ?thesis by simp
 qed
 
-lemma borel_measurable_real_of_ereal[simp, intro]:
-  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: ereal)) \<in> borel_measurable M"
-  using measurable_comp[OF f borel_measurable_real_of_ereal_borel] unfolding comp_def .
+lemma
+  fixes f :: "'a \<Rightarrow> ereal" assumes f[simp]: "f \<in> borel_measurable M"
+  shows borel_measurable_ereal_abs[intro, simp]: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
+    and borel_measurable_ereal_inverse[simp, intro]: "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
+    and borel_measurable_uminus_ereal[intro]: "(\<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 sets_Collect_If_set:
+  assumes "A \<inter> space M \<in> sets M" "{x \<in> space M. P x} \<in> sets M" "{x \<in> space M. Q x} \<in> sets M"
+  shows "{x \<in> space M. if x \<in> A then P x else Q x} \<in> sets M"
+proof -
+  have *: "{x \<in> space M. if x \<in> A then P x else Q x} = 
+    {x \<in> space M. if x \<in> A \<inter> space M then P x else Q x}" by auto
+  show ?thesis unfolding * unfolding if_bool_eq_conj using assms
+    by (auto intro!: sets_Collect simp: Int_def conj_commute)
+qed
+
+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 (f x))) (ereal (real (g x)))} \<in> sets M"
+    "{x \<in> space M. H (-\<infinity>) (ereal (real (g x)))} \<in> sets M"
+    "{x \<in> space M. H (\<infinity>) (ereal (real (g x)))} \<in> sets M"
+    "{x \<in> space M. H (ereal (real (f x))) (-\<infinity>)} \<in> sets M"
+    "{x \<in> space M. H (ereal (real (f x))) (\<infinity>)} \<in> sets M"
+  shows "{x \<in> space M. H (f x) (g x)} \<in> sets M"
+proof -
+  let ?G = "\<lambda>y x. if x \<in> g -` {\<infinity>} then H y \<infinity> else if x \<in> g -` {-\<infinity>} then H y (-\<infinity>) else H y (ereal (real (g x)))"
+  let ?F = "\<lambda>x. if x \<in> f -` {\<infinity>} then ?G \<infinity> x else if x \<in> f -` {-\<infinity>} then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x"
+  { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
+  moreover 
+  have "{x \<in> space M. ?F x} \<in> sets M"
+    by (intro sets_Collect H measurable_sets[OF f] measurable_sets[OF g] sets_Collect_If_set) auto
+  ultimately
+  show ?thesis by simp
+qed
+
+lemma
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows borel_measurable_ereal_le[intro,simp]: "{x \<in> space M. f x \<le> g x} \<in> sets M"
+    and borel_measurable_ereal_less[intro,simp]: "{x \<in> space M. f x < g x} \<in> sets M"
+    and borel_measurable_ereal_eq[intro,simp]: "{w \<in> space M. f w = g w} \<in> sets M"
+    and borel_measurable_ereal_neq[intro,simp]: "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
+  using f g by (auto simp: f g set_Collect_ereal2[OF f g] intro!: sets_Collect_neg)
 
 lemma borel_measurable_ereal_iff:
   shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
@@ -1029,52 +1057,6 @@
   finally show "f \<in> borel_measurable M" .
 qed (auto intro: measurable_sets borel_measurable_real_of_ereal)
 
-lemma less_eq_ge_measurable:
-  fixes f :: "'a \<Rightarrow> 'c::linorder"
-  shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
-proof
-  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 show "f -` {..a} \<inter> space M \<in> sets M" by auto
-next
-  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 show "f -` {a <..} \<inter> space M \<in> sets M" by auto
-qed
-
-lemma greater_eq_le_measurable:
-  fixes f :: "'a \<Rightarrow> 'c::linorder"
-  shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
-proof
-  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 show "f -` {..< a} \<inter> space M \<in> sets M" by auto
-next
-  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 show "f -` {a ..} \<inter> space M \<in> sets M" by auto
-qed
-
-lemma borel_measurable_uminus_borel_ereal:
-  "(uminus :: ereal \<Rightarrow> ereal) \<in> borel_measurable borel"
-proof (rule borel_measurableI)
-  fix X :: "ereal set" assume "open X"
-  have "uminus -` X = uminus ` X" by (force simp: image_iff)
-  then have "open (uminus -` X)" using `open X` ereal_open_uminus by auto
-  then show "uminus -` X \<inter> space borel \<in> sets borel" by auto
-qed
-
-lemma borel_measurable_uminus_ereal[intro]:
-  assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
-  using measurable_comp[OF assms borel_measurable_uminus_borel_ereal] by (simp add: comp_def)
-
-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 borel_measurable_eq_atMost_ereal:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
@@ -1118,94 +1100,88 @@
   then show "f \<in> borel_measurable M" by simp
 qed (simp add: measurable_sets)
 
+lemma greater_eq_le_measurable:
+  fixes f :: "'a \<Rightarrow> 'c::linorder"
+  shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
+proof
+  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 show "f -` {..< a} \<inter> space M \<in> sets M" by auto
+next
+  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 show "f -` {a ..} \<inter> space M \<in> sets M" by auto
+qed
+
 lemma borel_measurable_ereal_iff_less:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
   unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..
 
+lemma less_eq_ge_measurable:
+  fixes f :: "'a \<Rightarrow> 'c::linorder"
+  shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
+proof
+  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 show "f -` {..a} \<inter> space M \<in> sets M" by auto
+next
+  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 show "f -` {a <..} \<inter> space M \<in> sets M" by auto
+qed
+
 lemma borel_measurable_ereal_iff_ge:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
   unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable ..
 
-lemma borel_measurable_ereal_eq_const:
-  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
-  shows "{x\<in>space M. f x = c} \<in> sets M"
-proof -
-  have "{x\<in>space M. f x = c} = (f -` {c} \<inter> space M)" by auto
-  then show ?thesis using assms by (auto intro!: measurable_sets)
-qed
-
-lemma borel_measurable_ereal_neq_const:
-  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
-  shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
-proof -
-  have "{x\<in>space M. f x \<noteq> c} = space M - (f -` {c} \<inter> space M)" by auto
-  then show ?thesis using assms by (auto intro!: measurable_sets)
-qed
-
-lemma borel_measurable_ereal_le[intro,simp]:
-  fixes f g :: "'a \<Rightarrow> ereal"
+lemma borel_measurable_ereal2:
+  fixes f g :: "'a \<Rightarrow> ereal" 
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
-  shows "{x \<in> space M. f x \<le> g x} \<in> sets M"
+  assumes H: "(\<lambda>x. H (ereal (real (f x))) (ereal (real (g x)))) \<in> borel_measurable M"
+    "(\<lambda>x. H (-\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M"
+    "(\<lambda>x. H (\<infinity>) (ereal (real (g x)))) \<in> borel_measurable M"
+    "(\<lambda>x. H (ereal (real (f x))) (-\<infinity>)) \<in> borel_measurable M"
+    "(\<lambda>x. H (ereal (real (f x))) (\<infinity>)) \<in> borel_measurable M"
+  shows "(\<lambda>x. H (f x) (g x)) \<in> borel_measurable M"
 proof -
-  have "{x \<in> space M. f x \<le> g x} =
-    {x \<in> space M. real (f x) \<le> real (g x)} - (f -` {\<infinity>, -\<infinity>} \<inter> space M \<union> g -` {\<infinity>, -\<infinity>} \<inter> space M) \<union>
-    f -` {-\<infinity>} \<inter> space M \<union> g -` {\<infinity>} \<inter> space M" (is "?l = ?r")
-  proof (intro set_eqI)
-    fix x show "x \<in> ?l \<longleftrightarrow> x \<in> ?r" by (cases rule: ereal2_cases[of "f x" "g x"]) auto
-  qed
-  with f g show ?thesis by (auto intro!: Un simp: measurable_sets)
+  let ?G = "\<lambda>y x. if x \<in> g -` {\<infinity>} then H y \<infinity> else if x \<in> g -` {-\<infinity>} then H y (-\<infinity>) else H y (ereal (real (g x)))"
+  let ?F = "\<lambda>x. if x \<in> f -` {\<infinity>} then ?G \<infinity> x else if x \<in> f -` {-\<infinity>} then ?G (-\<infinity>) x else ?G (ereal (real (f x))) x"
+  { fix x have "H (f x) (g x) = ?F x" by (cases "f x" "g x" rule: ereal2_cases) auto }
+  moreover 
+  have "?F \<in> borel_measurable M"
+    by (intro measurable_If_set measurable_sets[OF f] measurable_sets[OF g] H) auto
+  ultimately
+  show ?thesis by simp
 qed
 
-lemma borel_measurable_ereal_less[intro,simp]:
-  fixes f :: "'a \<Rightarrow> ereal"
-  assumes f: "f \<in> borel_measurable M"
-  assumes g: "g \<in> borel_measurable M"
-  shows "{x \<in> space M. f x < g x} \<in> sets M"
-proof -
-  have "{x \<in> space M. f x < g x} = space M - {x \<in> space M. g x \<le> f x}" by auto
-  then show ?thesis using g f by auto
-qed
-
-lemma borel_measurable_ereal_eq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> ereal"
-  assumes f: "f \<in> borel_measurable M"
-  assumes g: "g \<in> borel_measurable M"
-  shows "{w \<in> space M. f w = g w} \<in> sets M"
-proof -
-  have "{x \<in> space M. f x = g x} = {x \<in> space M. g x \<le> f x} \<inter> {x \<in> space M. f x \<le> g x}" by auto
-  then show ?thesis using g f by auto
-qed
-
-lemma borel_measurable_ereal_neq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> ereal"
-  assumes f: "f \<in> borel_measurable M"
-  assumes g: "g \<in> borel_measurable M"
-  shows "{w \<in> space M. f w \<noteq> g w} \<in> sets M"
-proof -
-  have "{w \<in> space M. f w \<noteq> g w} = space M - {w \<in> space M. f w = g w}" by auto
-  thus ?thesis using f g by auto
-qed
+lemma
+  fixes f :: "'a \<Rightarrow> ereal" assumes f: "f \<in> borel_measurable M"
+  shows borel_measurable_ereal_eq_const: "{x\<in>space M. f x = c} \<in> sets M"
+    and borel_measurable_ereal_neq_const: "{x\<in>space M. f x \<noteq> c} \<in> sets M"
+  using f by auto
 
 lemma split_sets:
   "{x\<in>space M. P x \<or> Q x} = {x\<in>space M. P x} \<union> {x\<in>space M. Q x}"
   "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
   by auto
 
-lemma borel_measurable_ereal_add[intro, simp]:
+lemma
   fixes f :: "'a \<Rightarrow> ereal"
-  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
-proof -
-  { fix x assume "x \<in> space M" then have "f x + g x =
-      (if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
-        else if f x = -\<infinity> \<or> g x = -\<infinity> then -\<infinity>
-        else ereal (real (f x) + real (g x)))"
-      by (cases rule: ereal2_cases[of "f x" "g x"]) auto }
-  with assms show ?thesis
-    by (auto cong: measurable_cong simp: split_sets
-             intro!: Un measurable_If measurable_sets)
-qed
+  assumes [simp]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  shows borel_measurable_ereal_add[intro, simp]: "(\<lambda>x. f x + g x) \<in> borel_measurable M"
+    and borel_measurable_ereal_times[intro, simp]: "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+    and borel_measurable_ereal_min[simp, intro]: "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
+    and borel_measurable_ereal_max[simp, intro]: "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
+  by (auto simp add: borel_measurable_ereal2 measurable_If min_def max_def)
+
+lemma
+  fixes f g :: "'a \<Rightarrow> ereal"
+  assumes "f \<in> borel_measurable M"
+  assumes "g \<in> borel_measurable M"
+  shows borel_measurable_ereal_diff[simp, intro]: "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+    and borel_measurable_ereal_divide[simp, intro]: "(\<lambda>x. f x / g x) \<in> borel_measurable M"
+  unfolding minus_ereal_def divide_ereal_def using assms by auto
 
 lemma borel_measurable_ereal_setsum[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
@@ -1215,39 +1191,7 @@
   assume "finite S"
   thus ?thesis using assms
     by induct auto
-qed (simp add: borel_measurable_const)
-
-lemma borel_measurable_ereal_abs[intro, simp]:
-  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
-proof -
-  { fix x have "\<bar>f x\<bar> = (if 0 \<le> f x then f x else - f x)" by auto }
-  then show ?thesis using assms by (auto intro!: measurable_If)
-qed
-
-lemma borel_measurable_ereal_times[intro, simp]:
-  fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
-proof -
-  { fix f g :: "'a \<Rightarrow> ereal"
-    assume b: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-      and pos: "\<And>x. 0 \<le> f x" "\<And>x. 0 \<le> g x"
-    { fix x have *: "f x * g x = (if f x = 0 \<or> g x = 0 then 0
-        else if f x = \<infinity> \<or> g x = \<infinity> then \<infinity>
-        else ereal (real (f x) * real (g x)))"
-      apply (cases rule: ereal2_cases[of "f x" "g x"])
-      using pos[of x] by auto }
-    with b have "(\<lambda>x. f x * g x) \<in> borel_measurable M"
-      by (auto cong: measurable_cong simp: split_sets
-               intro!: Un measurable_If measurable_sets) }
-  note pos_times = this
-  have *: "(\<lambda>x. f x * g x) =
-    (\<lambda>x. if 0 \<le> f x \<and> 0 \<le> g x \<or> f x < 0 \<and> g x < 0 then \<bar>f x\<bar> * \<bar>g x\<bar> else - (\<bar>f x\<bar> * \<bar>g x\<bar>))"
-    by (auto simp: fun_eq_iff)
-  show ?thesis using assms unfolding *
-    by (intro measurable_If pos_times borel_measurable_uminus_ereal)
-       (auto simp: split_sets intro!: Int)
-qed
+qed simp
 
 lemma borel_measurable_ereal_setprod[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
@@ -1258,20 +1202,6 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma borel_measurable_ereal_min[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> ereal"
-  assumes "f \<in> borel_measurable M"
-  assumes "g \<in> borel_measurable M"
-  shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
-  using assms unfolding min_def by (auto intro!: measurable_If)
-
-lemma borel_measurable_ereal_max[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> ereal"
-  assumes "f \<in> borel_measurable M"
-  and "g \<in> borel_measurable M"
-  shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
-  using assms unfolding max_def by (auto intro!: measurable_If)
-
 lemma borel_measurable_SUP[simp, intro]:
   fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
@@ -1298,38 +1228,24 @@
     using assms by auto
 qed
 
-lemma borel_measurable_liminf[simp, intro]:
-  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
-  assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
-  unfolding liminf_SUPR_INFI using assms by auto
-
-lemma borel_measurable_limsup[simp, intro]:
+lemma
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
-  unfolding limsup_INFI_SUPR using assms by auto
-
-lemma borel_measurable_ereal_diff[simp, intro]:
-  fixes f g :: "'a \<Rightarrow> ereal"
-  assumes "f \<in> borel_measurable M"
-  assumes "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-  unfolding minus_ereal_def using assms by auto
+  shows borel_measurable_liminf[simp, intro]: "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
+    and borel_measurable_limsup[simp, intro]: "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
+  unfolding liminf_SUPR_INFI limsup_INFI_SUPR using assms by auto
 
-lemma borel_measurable_ereal_inverse[simp, intro]:
-  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
+lemma borel_measurable_ereal_LIMSEQ:
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
+  assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
+  and u: "\<And>i. u i \<in> borel_measurable M"
+  shows "u' \<in> borel_measurable M"
 proof -
-  { fix x have "inverse (f x) = (if f x = 0 then \<infinity> else ereal (inverse (real (f x))))"
-      by (cases "f x") auto }
-  with f show ?thesis
-    by (auto intro!: measurable_If)
+  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])
+  then show ?thesis by (simp add: u cong: measurable_cong)
 qed
 
-lemma borel_measurable_ereal_divide[simp, intro]:
-  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. f x / g x :: ereal) \<in> borel_measurable M"
-  unfolding divide_ereal_def by auto
-
 lemma borel_measurable_psuminf[simp, intro]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
@@ -1354,4 +1270,38 @@
   ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_ereal_iff)
 qed
 
-end
+lemma sets_Collect_Cauchy: 
+  fixes f :: "nat \<Rightarrow> 'a => real"
+  assumes f: "\<And>i. f i \<in> borel_measurable M"
+  shows "{x\<in>space M. Cauchy (\<lambda>i. f i x)} \<in> sets M"
+  unfolding Cauchy_iff2 using f by (auto intro!: sets_Collect)
+
+lemma borel_measurable_lim:
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+  assumes f: "\<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 Cauchy (\<lambda>i. f i x) then lim (\<lambda>i. if Cauchy (\<lambda>i. f i x) then f i x else 0) else (THE x. False))"
+    by (auto simp: lim_def convergent_eq_cauchy[symmetric])
+  { 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) }
+  note convergent = this
+  show ?thesis
+    unfolding *
+    apply (intro measurable_If sets_Collect_Cauchy f borel_measurable_const)
+    apply (rule borel_measurable_LIMSEQ)
+    apply (rule convergent_LIMSEQ_iff[THEN iffD1, OF convergent])
+    apply (intro measurable_If sets_Collect_Cauchy f borel_measurable_const)
+    done
+qed
+
+lemma borel_measurable_suminf:
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+  assumes f: "\<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 add: f borel_measurable_lim)
+
+end 
--- a/src/HOL/Probability/Caratheodory.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -9,12 +9,6 @@
   imports Measure_Space
 begin
 
-lemma sums_def2:
-  "f sums x \<longleftrightarrow> (\<lambda>n. (\<Sum>i\<le>n. f i)) ----> x"
-  unfolding sums_def
-  apply (subst LIMSEQ_Suc_iff[symmetric])
-  unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost ..
-
 text {*
   Originally from the Hurd/Coble measure theory development, translated by Lawrence Paulson.
 *}
@@ -684,146 +678,6 @@
 
 subsubsection {*Alternative instances of caratheodory*}
 
-lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
-  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)) ----> 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=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
-    using f(1)[unfolded positive_def] dA
-    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos)
-  from LIMSEQ_Suc[OF this]
-  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
-    unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
-  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)) ----> 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)) ----> 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\<le>n. A i)) = (\<Union>i. A i)" by auto
-  have "(\<lambda>n. f (\<Union>i\<le>n. A i)) ----> f (\<Union>i. A i)"
-  proof (unfold *[symmetric], intro cont[rule_format])
-    show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> M"
-      using A * by auto
-  qed (force intro!: incseq_SucI)
-  moreover have "\<And>n. f (\<Union>i\<le>n. A i) = (\<Sum>i\<le>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_def2 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:
-  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)) ----> 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)) ----> 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)) ----> 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)) ----> 0"
-    using `positive M f`[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)) ----> 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
-  { 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 }
-  note f_fin = this
-  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 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 INFI_ereal_add) (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_ereal_INFI[OF decseq_fA]
-  show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
-qed
-
-lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
-lemma positiveD2: "positive M f \<Longrightarrow> A \<in> M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
-
-lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
-  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)) ----> 0"
-  assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
-  shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
-proof -
-  have "\<forall>A\<in>M. \<exists>x. f A = ereal x"
-  proof
-    fix A assume "A \<in> M" with f show "\<exists>x. f A = ereal x"
-      unfolding positive_def by (cases "f A") auto
-  qed
-  from bchoice[OF this] guess f' .. note f' = this[rule_format]
-  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) ----> 0"
-    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
-  moreover
-  { fix i
-    have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)"
-      using A by (intro f(2)[THEN additiveD, symmetric]) 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) - f' (A i) = f' ((\<Union>i. A i) - A i)"
-      using A by (subst (asm) (1 2 3) f') auto
-    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
-      using A f' by auto }
-  ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
-    by (simp add: zero_ereal_def)
-  then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
-    by (rule LIMSEQ_diff_approach_zero2[OF tendsto_const])
-  then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
-    using A by (subst (1 2) f') auto
-qed
-
-lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
-  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)) ----> 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
-
 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)) ----> 0"
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -44,45 +44,45 @@
   unfolding extensional_def by auto
 
 definition
-  "merge I x J y = (\<lambda>i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
+  "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 x J y i = x i"
-  "I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
-  "J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i"
-  "J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i"
-  "i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I x J y i = undefined"
+  "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 x J y = merge J y I x"
+  "I \<inter> J = {} \<Longrightarrow> merge I J (x, y) = merge J I (y, x)"
   by (auto simp: merge_def intro!: ext)
 
 lemma Pi_cancel_merge_range[simp]:
-  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
-  "I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
-  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A"
-  "J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A"
+  "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 x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
-  "I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
-  "J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B"
+  "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 x J y \<in> extensional (I \<union> J)"
+lemma extensional_merge[simp]: "merge I J (x, y) \<in> extensional (I \<union> J)"
   by (auto simp: extensional_def)
 
 lemma restrict_Pi_cancel: "restrict x I \<in> Pi I A \<longleftrightarrow> x \<in> Pi I A"
   by (auto simp: restrict_def Pi_def)
 
 lemma restrict_merge[simp]:
-  "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
-  "I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
-  "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I"
-  "J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J"
+  "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 extensional_insert_undefined[intro, simp]:
@@ -95,7 +95,7 @@
   shows "a \<in> extensional (insert i I)"
   using assms unfolding extensional_def by auto
 
-lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I x {i} y = restrict (x(i := y i)) (insert i I)"
+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 Pi_Int: "Pi I E \<inter> Pi I F = (\<Pi> i\<in>I. E i \<inter> F i)"
@@ -118,24 +118,24 @@
 
 lemma restrict_vimage:
   assumes "I \<inter> J = {}"
-  shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I E J F)"
+  shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I J (E, F))"
   using assms by (auto simp: restrict_Pi_cancel)
 
 lemma merge_vimage:
   assumes "I \<inter> J = {}"
-  shows "(\<lambda>(x,y). merge I x J y) -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E"
+  shows "merge I J -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E"
   using assms by (auto simp: restrict_Pi_cancel)
 
 lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I"
   by (auto simp: restrict_def)
 
 lemma merge_restrict[simp]:
-  "merge I (restrict x I) J y = merge I x J y"
-  "merge I x J (restrict y J) = merge I x J y"
+  "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 x J x = restrict x (I \<union> J)"
+  "merge I J (x, x) = restrict x (I \<union> J)"
   unfolding merge_def by auto
 
 lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A"
@@ -504,8 +504,7 @@
 qed (insert `i \<in> I`, auto simp: space_PiM)
 
 lemma measurable_add_dim:
-  assumes "i \<notin> I"
-  shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
+  "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>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)"
@@ -519,19 +518,18 @@
 qed (auto simp: space_pair_measure space_PiM)
 
 lemma measurable_merge:
-  assumes "I \<inter> J = {}"
-  shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
+  "merge I J \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>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. prod_case (\<lambda>x. merge I x J) \<omega> i \<in> A} =
+  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)"
-    using `I \<inter> J = {}` by auto
+    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. prod_case (\<lambda>x. merge I x J) \<omega> i \<in> A} \<in> sets ?P" .
-qed (insert assms, auto simp: space_pair_measure space_PiM)
+  finally show "{\<omega> \<in> space ?P. merge I J \<omega> i \<in> A} \<in> sets ?P" .
+qed (auto simp: space_pair_measure space_PiM Pi_iff merge_def extensional_def)
 
 lemma measurable_restrict:
   assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
@@ -587,44 +585,38 @@
   qed
 qed
 
-lemma (in product_sigma_finite)
-  assumes "finite I"
-  shows sigma_finite: "sigma_finite_measure (PiM I M)"
-  and emeasure_PiM:
-    "\<And>A. (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
-using `finite I` proof induct
-  case empty
+lemma
+  shows space_PiM_empty: "space (Pi\<^isub>M {} M) = {\<lambda>k. undefined}"
+    and sets_PiM_empty: "sets (Pi\<^isub>M {} M) = { {}, {\<lambda>k. undefined} }"
+  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
+
+lemma emeasure_PiM_empty[simp]: "emeasure (PiM {} M) {\<lambda>_. undefined} = 1"
+proof -
   let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ereal)"
-  have "prod_algebra {} M = {{\<lambda>_. undefined}}"
-    by (auto simp: prod_algebra_def prod_emb_def intro!: image_eqI)
-  then have sets_empty: "sets (PiM {} M) = {{}, {\<lambda>_. undefined}}"
-    by (simp add: sets_PiM)
   have "emeasure (Pi\<^isub>M {} M) (prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = 1"
   proof (subst emeasure_extend_measure_Pair[OF PiM_def])
-    have "finite (space (PiM {} M))"
-      by (simp add: space_PiM)
-    moreover show "positive (PiM {} M) ?\<mu>"
+    show "positive (PiM {} M) ?\<mu>"
       by (auto simp: positive_def)
-    ultimately show "countably_additive (PiM {} M) ?\<mu>"
-      by (rule countably_additiveI_finite) (auto simp: additive_def space_PiM sets_empty)
+    show "countably_additive (PiM {} M) ?\<mu>"
+      by (rule 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>\<^isub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
+  also have "(prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
     by (auto simp: prod_emb_def)
-  finally have emeasure_eq: "emeasure (Pi\<^isub>M {} M) {\<lambda>_. undefined} = 1" .
+  finally show ?thesis
+    by simp
+qed
 
-  interpret finite_measure "PiM {} M"
-    by default (simp add: space_PiM emeasure_eq)
-  case 1 show ?case ..
+lemma PiM_empty: "PiM {} M = count_space {\<lambda>_. undefined}"
+  by (rule measure_eqI) (auto simp add: sets_PiM_empty one_ereal_def)
 
-  case 2 show ?case
-    using emeasure_eq * by simp
-next
+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\<^isub>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 default fact
   have "finite (insert i I)" using `finite I` by auto
   interpret I': finite_product_sigma_finite M "insert i I" by default fact
-  interpret I: sigma_finite_measure "PiM I M" by fact
-  interpret P: pair_sigma_finite "PiM I M" "M i" ..
   let ?h = "(\<lambda>(f, y). f(i := y))"
 
   let ?P = "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M) ?h"
@@ -632,67 +624,73 @@
   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))"
 
-  { case 2
-    have "emeasure (Pi\<^isub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^isub>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\<^isub>E J E)"
-      let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^isub>E j\<in>J-{i}. E j)"
-      have "?\<mu> ?p =
-        emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i))"
-        by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
-      also have "?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
-        using J E[rule_format, THEN sets_into_space]
-        by (force simp: space_pair_measure space_PiM Pi_iff prod_emb_iff split: split_if_asm)
-      also have "emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
-        emeasure (Pi\<^isub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
-        using J E by (intro P.emeasure_pair_measure_Times sets_PiM_I) auto
-      also have "?p' = (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
-        using J E[rule_format, THEN sets_into_space]
-        by (auto simp: prod_emb_iff Pi_iff split: split_if_asm) blast+
-      also have "emeasure (Pi\<^isub>M I M) (\<Pi>\<^isub>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_one_right) auto
-      finally show "?\<mu> ?p = \<dots>" .
+  have "emeasure (Pi\<^isub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^isub>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\<^isub>E J E)"
+    let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^isub>E j\<in>J-{i}. E j)"
+    have "?\<mu> ?p =
+      emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i))"
+      by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
+    also have "?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
+      using J E[rule_format, THEN sets_into_space]
+      by (force simp: space_pair_measure space_PiM Pi_iff prod_emb_iff split: split_if_asm)
+    also have "emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
+      emeasure (Pi\<^isub>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>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
+      using J E[rule_format, THEN sets_into_space]
+      by (auto simp: prod_emb_iff Pi_iff split: split_if_asm) blast+
+    also have "emeasure (Pi\<^isub>M I M) (\<Pi>\<^isub>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_one_right) auto
+    finally show "?\<mu> ?p = \<dots>" .
 
-      show "prod_emb (insert i I) M J (Pi\<^isub>E J E) \<in> Pow (\<Pi>\<^isub>E i\<in>insert i I. space (M i))"
-        using J E[rule_format, THEN sets_into_space] by (auto simp: prod_emb_iff)
-    next
-      show "positive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^isub>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(1,2) 2 by auto
-    qed (auto intro!: setprod_cong)
-    with 2[THEN sets_into_space] show ?case by (subst (asm) prod_emb_PiE_same_index) auto }
-  note product = this
+    show "prod_emb (insert i I) M J (Pi\<^isub>E J E) \<in> Pow (\<Pi>\<^isub>E i\<in>insert i I. space (M i))"
+      using J E[rule_format, THEN sets_into_space] by (auto simp: prod_emb_iff)
+  next
+    show "positive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^isub>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_into_space)
+qed (simp add: emeasure_PiM_empty)
 
-  from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-  then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
-    "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)"
-    "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space (Pi\<^isub>M (insert i I) M)"
-    "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> emeasure (M j) (F j k) \<noteq> \<infinity>"
+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 default fact
+
+  from sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
+  then have F: "\<And>j. j \<in> I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
+    "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> I. F j k)"
+    "(\<Union>k. \<Pi>\<^isub>E j \<in> I. F j k) = space (Pi\<^isub>M I M)"
+    "\<And>k. \<And>j. j \<in> I \<Longrightarrow> emeasure (M j) (F j k) \<noteq> \<infinity>"
     by blast+
-  let ?F = "\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k"
+  let ?F = "\<lambda>k. \<Pi>\<^isub>E j \<in> I. F j k"
 
-  case 1 show ?case
+  show ?thesis
   proof (unfold_locales, intro exI[of _ ?F] conjI allI)
-    show "range ?F \<subseteq> sets (Pi\<^isub>M (insert i I) M)" using F(1) insert(1,2) by auto
+    show "range ?F \<subseteq> sets (Pi\<^isub>M I M)" using F(1) `finite I` by auto
   next
-    from F(3) show "(\<Union>i. ?F i) = space (Pi\<^isub>M (insert i I) M)" by simp
+    from F(3) show "(\<Union>i. ?F i) = space (Pi\<^isub>M I M)" by simp
   next
     fix j
-    from F `finite I` setprod_PInf[of "insert i I", OF emeasure_nonneg, of M]
-    show "emeasure (\<Pi>\<^isub>M i\<in>insert i I. M i) (?F j) \<noteq> \<infinity>"
-      by (subst product) auto
+    from F `finite I` setprod_PInf[of I, OF emeasure_nonneg, of M]
+    show "emeasure (\<Pi>\<^isub>M i\<in>I. M i) (?F j) \<noteq> \<infinity>"
+      by (subst emeasure_PiM) auto
   qed
 qed
 
@@ -703,18 +701,6 @@
   "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^isub>M I M) (Pi\<^isub>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) product_measure_empty[simp]:
-  "emeasure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
-proof -
-  interpret finite_product_sigma_finite M "{}" by default auto
-  from measure_times[of "\<lambda>x. {}"] show ?thesis by simp
-qed
-
-lemma
-  shows space_PiM_empty: "space (Pi\<^isub>M {} M) = {\<lambda>k. undefined}"
-    and sets_PiM_empty: "sets (Pi\<^isub>M {} M) = { {}, {\<lambda>k. undefined} }"
-  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
-
 lemma (in product_sigma_finite) positive_integral_empty:
   assumes pos: "0 \<le> f (\<lambda>k. undefined)"
   shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
@@ -734,7 +720,7 @@
 
 lemma (in product_sigma_finite) distr_merge:
   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
-  shows "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M) (\<lambda>(x,y). merge I x J y) = Pi\<^isub>M (I \<union> J) M"
+  shows "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M) (merge I J) = Pi\<^isub>M (I \<union> J) M"
    (is "?D = ?P")
 proof -
   interpret I: finite_product_sigma_finite M I by default fact
@@ -742,7 +728,7 @@
   have "finite (I \<union> J)" using fin by auto
   interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
   interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
-  let ?g = "\<lambda>(x,y). merge I x J y"
+  let ?g = "merge I J"
 
   from IJ.sigma_finite_pairs obtain F where
     F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
@@ -765,7 +751,6 @@
 
     show "range ?F \<subseteq> prod_algebra (I \<union> J) M" using F
       using fin by (auto simp: prod_algebra_eq_finite)
-    show "incseq ?F" by fact
     show "(\<Union>i. \<Pi>\<^isub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i))"
       using F(3) by (simp add: space_PiM)
   next
@@ -786,7 +771,7 @@
       using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM)
     also have "emeasure ?B ?X = (\<Prod>i\<in>I. emeasure (M i) (F i)) * (\<Prod>i\<in>J. emeasure (M i) (F i))"
       using `finite J` `finite I` F X
-      by (simp add: P.emeasure_pair_measure_Times I.measure_times J.measure_times Pi_iff)
+      by (simp add: J.emeasure_pair_measure_Times I.measure_times J.measure_times Pi_iff)
     also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))"
       using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
     also have "\<dots> = emeasure ?P (Pi\<^isub>E (I \<union> J) F)"
@@ -800,16 +785,16 @@
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
   and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
   shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
-    (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
+    (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I J (x, y)) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
 proof -
   interpret I: finite_product_sigma_finite M I by default fact
   interpret J: finite_product_sigma_finite M J by default fact
   interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
-  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
-    using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def)
+  have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>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 positive_integral_distr[OF measurable_merge f, OF IJ(1)])
+    apply (subst positive_integral_distr[OF measurable_merge f])
     apply (subst P.positive_integral_fst_measurable(2)[symmetric, OF P_borel])
     apply simp
     done
@@ -840,7 +825,7 @@
 qed
 
 lemma (in product_sigma_finite) product_positive_integral_insert:
-  assumes [simp]: "finite I" "i \<notin> I"
+  assumes I[simp]: "finite I" "i \<notin> I"
     and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
   shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))"
 proof -
@@ -849,17 +834,17 @@
   have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
     using f by auto
   show ?thesis
-    unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f]
-  proof (rule positive_integral_cong, subst product_positive_integral_singleton)
+    unfolding product_positive_integral_fold[OF IJ, unfolded insert, OF I(1) i.finite_index f]
+  proof (rule positive_integral_cong, subst product_positive_integral_singleton[symmetric])
     fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
-    let ?f = "\<lambda>y. f (restrict (x(i := y)) (insert i I))"
-    have f'_eq: "\<And>y. ?f y = f (x(i := y))"
-      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def space_PiM)
-    show "?f \<in> borel_measurable (M i)" unfolding f'_eq
+    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 `i \<notin> I`]
       unfolding comp_def .
-    show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i"
-      unfolding f'_eq by simp
+    show "(\<integral>\<^isup>+ y. f (merge I {i} (x, y)) \<partial>Pi\<^isub>M {i} M) = (\<integral>\<^isup>+ y. f (x(i := y i)) \<partial>Pi\<^isub>M {i} M)"
+      using x
+      by (auto intro!: positive_integral_cong arg_cong[where f=f]
+               simp add: space_PiM extensional_def)
   qed
 qed
 
@@ -902,29 +887,54 @@
 lemma (in product_sigma_finite) product_integral_fold:
   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
   and f: "integrable (Pi\<^isub>M (I \<union> J) M) f"
-  shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I x J y) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
+  shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I J (x, y)) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)"
 proof -
   interpret I: finite_product_sigma_finite M I by default fact
   interpret J: finite_product_sigma_finite M J by default fact
   have "finite (I \<union> J)" using fin by auto
   interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
   interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
-  let ?M = "\<lambda>(x, y). merge I x J y"
+  let ?M = "merge I J"
   let ?f = "\<lambda>x. f (?M x)"
   from f have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
     by auto
-  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
-    using measurable_comp[OF measurable_merge[OF IJ(1)] f_borel] by (simp add: comp_def)
+  have P_borel: "(\<lambda>x. f (merge I J x)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
+    using measurable_comp[OF measurable_merge f_borel] by (simp add: comp_def)
   have f_int: "integrable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ?f"
-    by (rule integrable_distr[OF measurable_merge[OF IJ]]) (simp add: distr_merge[OF IJ fin] 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[OF IJ] f_borel])
+    apply (subst integral_distr[OF measurable_merge f_borel])
     apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int])
     apply simp
     done
 qed
 
+lemma (in product_sigma_finite)
+  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+  shows emeasure_fold_integral:
+    "emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
+    and emeasure_fold_measurable:
+    "(\<lambda>x. emeasure (Pi\<^isub>M J M) ((\<lambda>y. merge I J (x, y)) -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
+proof -
+  interpret I: finite_product_sigma_finite M I by default fact
+  interpret J: finite_product_sigma_finite M J by default fact
+  interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" ..
+  have merge: "merge I J -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>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!: positive_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_compose[symmetric] comp_def space_pair_measure cong: measurable_cong)
+qed
+
 lemma (in product_sigma_finite) product_integral_insert:
   assumes I: "finite I" "i \<notin> I"
     and f: "integrable (Pi\<^isub>M (insert i I) M) f"
@@ -932,7 +942,7 @@
 proof -
   have "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = integral\<^isup>L (Pi\<^isub>M (I \<union> {i}) M) f"
     by simp
-  also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I x {i} y) \<partial>Pi\<^isub>M {i} M) \<partial>Pi\<^isub>M I M)"
+  also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I {i} (x,y)) \<partial>Pi\<^isub>M {i} M) \<partial>Pi\<^isub>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\<^isub>M I M)"
   proof (rule integral_cong, subst product_integral_singleton[symmetric])
@@ -942,7 +952,7 @@
     show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
       using measurable_comp[OF measurable_component_update f_borel, OF x `i \<notin> I`]
       unfolding comp_def .
-    from x I show "(\<integral> y. f (merge I x {i} y) \<partial>Pi\<^isub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^isub>M {i} M)"
+    from x I show "(\<integral> y. f (merge I {i} (x,y)) \<partial>Pi\<^isub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^isub>M {i} M)"
       by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def)
   qed
   finally show ?thesis .
@@ -980,30 +990,83 @@
 
 lemma (in product_sigma_finite) product_integral_setprod:
   fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real"
-  assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)"
+  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\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))"
-using assms proof (induct rule: finite_ne_induct)
-  case (singleton i)
-  then show ?case by (simp add: product_integral_singleton integrable_def)
+using assms proof induct
+  case empty
+  interpret finite_measure "Pi\<^isub>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\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))"
-    by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset)
+    by (intro product_integrable_setprod insert(4)) (auto intro: finite_subset)
   interpret I: finite_product_sigma_finite M I by default 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 `i \<notin> I` by (auto intro!: setprod_cong)
   show ?case
-    unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]]
+    unfolding product_integral_insert[OF insert(1,2) prod[OF subset_refl]]
     by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI)
 qed
 
+lemma sets_Collect_single:
+  "i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> { x \<in> space (Pi\<^isub>M I M). x i \<in> A } \<in> sets (Pi\<^isub>M I M)"
+  unfolding sets_PiM_single
+  by (auto intro!: sigma_sets.Basic exI[of _ i] exI[of _ A]) (auto simp: space_PiM)
+
+lemma sigma_prod_algebra_sigma_eq_infinite:
+  fixes E :: "'i \<Rightarrow> 'a set set"
+  assumes S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
+    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
+  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
+    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
+  defines "P == {{f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> E i}"
+  shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
+proof
+  let ?P = "sigma (space (Pi\<^isub>M I M)) P"
+  have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
+    using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq)
+  then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+    by (simp add: space_PiM)
+  have "sets (PiM I M) =
+      sigma_sets (space ?P) {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+    using sets_PiM_single[of I M] by (simp add: space_P)
+  also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)"
+  proof (safe intro!: sigma_sets_subset)
+    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
+    then have "(\<lambda>x. x i) \<in> measurable ?P (sigma (space (M i)) (E i))"
+      apply (subst measurable_iff_measure_of)
+      apply (simp_all add: P_closed)
+      using E_closed
+      apply (force simp: subset_eq space_PiM)
+      apply (force simp: subset_eq space_PiM)
+      apply (auto simp: P_def intro!: sigma_sets.Basic exI[of _ i])
+      apply (rule_tac x=Aa in exI)
+      apply (auto simp: space_PiM)
+      done
+    from measurable_sets[OF this, of A] A `i \<in> I` E_closed
+    have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
+      by (simp add: E_generates)
+    also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A}"
+      using P_closed by (auto simp: space_PiM)
+    finally show "\<dots> \<in> sets ?P" .
+  qed
+  finally show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) P"
+    by (simp add: P_closed)
+  show "sigma_sets (space (PiM I M)) P \<subseteq> sets (PiM I M)"
+    unfolding P_def space_PiM[symmetric]
+    by (intro sigma_sets_subset) (auto simp: E_generates sets_Collect_single)
+qed
+
+lemma bchoice_iff: "(\<forall>a\<in>A. \<exists>b. P a b) \<longleftrightarrow> (\<exists>f. \<forall>a\<in>A. P a (f a))"
+  by metis
+
 lemma sigma_prod_algebra_sigma_eq:
-  fixes E :: "'i \<Rightarrow> 'a set set"
+  fixes E :: "'i \<Rightarrow> 'a set set" and S :: "'i \<Rightarrow> nat \<Rightarrow> 'a set"
   assumes "finite I"
-  assumes S_mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
-    and S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
+  assumes S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
     and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
   assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
     and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
@@ -1011,6 +1074,9 @@
   shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
 proof
   let ?P = "sigma (space (Pi\<^isub>M I M)) P"
+  from `finite I`[THEN ex_bij_betw_finite_nat] guess T ..
+  then have T: "\<And>i. i \<in> I \<Longrightarrow> T i < card I" "\<And>i. i\<in>I \<Longrightarrow> the_inv_into I T (T i) = i"
+    by (auto simp add: bij_betw_def set_eq_iff image_iff the_inv_into_f_f)
   have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
     using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq)
   then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
@@ -1033,15 +1099,18 @@
           using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
         also have "\<dots> = (\<Pi>\<^isub>E j\<in>I. \<Union>n. if i = j then A else S j n)"
           by (intro PiE_cong) (simp add: S_union)
-        also have "\<dots> = (\<Union>n. \<Pi>\<^isub>E j\<in>I. if i = j then A else S j n)"
-          using S_mono
-          by (subst Pi_UN[symmetric, OF `finite I`]) (auto simp: incseq_def)
+        also have "\<dots> = (\<Union>xs\<in>{xs. length xs = card I}. \<Pi>\<^isub>E j\<in>I. if i = j then A else S j (xs ! T j))"
+          using T
+          apply (auto simp: Pi_iff bchoice_iff)
+          apply (rule_tac x="map (\<lambda>n. f (the_inv_into I T n)) [0..<card I]" in exI)
+          apply (auto simp: bij_betw_def)
+          done
         also have "\<dots> \<in> sets ?P"
         proof (safe intro!: countable_UN)
-          fix n show "(\<Pi>\<^isub>E j\<in>I. if i = j then A else S j n) \<in> sets ?P"
+          fix xs show "(\<Pi>\<^isub>E j\<in>I. if i = j then A else S j (xs ! T j)) \<in> sets ?P"
             using A S_in_E
             by (simp add: P_closed)
-               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j n"])
+               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j (xs ! T j)"])
         qed
         finally show "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
           using P_closed by simp
--- a/src/HOL/Probability/Independent_Family.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Independent_Family.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -23,13 +23,6 @@
 qed auto
 
 definition (in prob_space)
-  "indep_events A I \<longleftrightarrow> (A`I \<subseteq> events) \<and>
-    (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j)))"
-
-definition (in prob_space)
-  "indep_event A B \<longleftrightarrow> indep_events (bool_case A B) UNIV"
-
-definition (in prob_space)
   "indep_sets F I \<longleftrightarrow> (\<forall>i\<in>I. F i \<subseteq> events) \<and>
     (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> (\<forall>A\<in>Pi J F. prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))))"
 
@@ -37,22 +30,27 @@
   "indep_set A B \<longleftrightarrow> indep_sets (bool_case A B) UNIV"
 
 definition (in prob_space)
-  "indep_vars M' X I \<longleftrightarrow>
-    (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and>
-    indep_sets (\<lambda>i. sigma_sets (space M) { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
+  indep_events_def_alt: "indep_events A I \<longleftrightarrow> indep_sets (\<lambda>i. {A i}) I"
+
+lemma image_subset_iff_funcset: "F ` A \<subseteq> B \<longleftrightarrow> F \<in> A \<rightarrow> B"
+  by auto
+
+lemma (in prob_space) indep_events_def:
+  "indep_events A I \<longleftrightarrow> (A`I \<subseteq> events) \<and>
+    (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j)))"
+  unfolding indep_events_def_alt indep_sets_def
+  apply (simp add: Ball_def Pi_iff image_subset_iff_funcset)
+  apply (intro conj_cong refl arg_cong[where f=All] ext imp_cong)
+  apply auto
+  done
 
 definition (in prob_space)
-  "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (bool_case Ma Mb) (bool_case A B) UNIV"
+  "indep_event A B \<longleftrightarrow> indep_events (bool_case A B) UNIV"
 
 lemma (in prob_space) indep_sets_cong:
   "I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J"
   by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+
 
-lemma (in prob_space) indep_sets_singleton_iff_indep_events:
-  "indep_sets (\<lambda>i. {F i}) I \<longleftrightarrow> indep_events F I"
-  unfolding indep_sets_def indep_events_def
-  by (simp, intro conj_cong ball_cong all_cong imp_cong) (auto simp: Pi_iff)
-
 lemma (in prob_space) indep_events_finite_index_events:
   "indep_events F I \<longleftrightarrow> (\<forall>J\<subseteq>I. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> indep_events F J)"
   by (auto simp: indep_events_def)
@@ -86,6 +84,15 @@
     using indep by (auto simp: indep_sets_def)
 qed
 
+lemma (in prob_space) indep_sets_mono:
+  assumes indep: "indep_sets F I"
+  assumes mono: "J \<subseteq> I" "\<And>i. i\<in>J \<Longrightarrow> G i \<subseteq> F i"
+  shows "indep_sets G J"
+  apply (rule indep_sets_mono_sets)
+  apply (rule indep_sets_mono_index)
+  apply (fact +)
+  done
+
 lemma (in prob_space) indep_setsI:
   assumes "\<And>i. i \<in> I \<Longrightarrow> F i \<subseteq> events"
     and "\<And>A J. J \<noteq> {} \<Longrightarrow> J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<forall>j\<in>J. A j \<in> F j) \<Longrightarrow> prob (\<Inter>j\<in>J. A j) = (\<Prod>j\<in>J. prob (A j))"
@@ -117,16 +124,6 @@
   using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev
   by (simp add: ac_simps UNIV_bool)
 
-lemma (in prob_space) indep_var_eq:
-  "indep_var S X T Y \<longleftrightarrow>
-    (random_variable S X \<and> random_variable T Y) \<and>
-    indep_set
-      (sigma_sets (space M) { X -` A \<inter> space M | A. A \<in> sets S})
-      (sigma_sets (space M) { Y -` A \<inter> space M | A. A \<in> sets T})"
-  unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool
-  by (intro arg_cong2[where f="op \<and>"] arg_cong2[where f=indep_sets] ext)
-     (auto split: bool.split)
-
 lemma (in prob_space)
   assumes indep: "indep_set A B"
   shows indep_setD_ev1: "A \<subseteq> events"
@@ -326,6 +323,36 @@
     by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic)
 qed
 
+definition (in prob_space)
+  indep_vars_def2: "indep_vars M' X I \<longleftrightarrow>
+    (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and>
+    indep_sets (\<lambda>i. { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
+
+definition (in prob_space)
+  "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (bool_case Ma Mb) (bool_case A B) UNIV"
+
+lemma (in prob_space) indep_vars_def:
+  "indep_vars M' X I \<longleftrightarrow>
+    (\<forall>i\<in>I. random_variable (M' i) (X i)) \<and>
+    indep_sets (\<lambda>i. sigma_sets (space M) { X i -` A \<inter> space M | A. A \<in> sets (M' i)}) I"
+  unfolding indep_vars_def2
+  apply (rule conj_cong[OF refl])
+  apply (rule indep_sets_sigma_sets_iff[symmetric])
+  apply (auto simp: Int_stable_def)
+  apply (rule_tac x="A \<inter> Aa" in exI)
+  apply auto
+  done
+
+lemma (in prob_space) indep_var_eq:
+  "indep_var S X T Y \<longleftrightarrow>
+    (random_variable S X \<and> random_variable T Y) \<and>
+    indep_set
+      (sigma_sets (space M) { X -` A \<inter> space M | A. A \<in> sets S})
+      (sigma_sets (space M) { Y -` A \<inter> space M | A. A \<in> sets T})"
+  unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool
+  by (intro arg_cong2[where f="op \<and>"] arg_cong2[where f=indep_sets] ext)
+     (auto split: bool.split)
+
 lemma (in prob_space) indep_sets2_eq:
   "indep_set A B \<longleftrightarrow> A \<subseteq> events \<and> B \<subseteq> events \<and> (\<forall>a\<in>A. \<forall>b\<in>B. prob (a \<inter> b) = prob a * prob b)"
   unfolding indep_set_def
@@ -468,27 +495,25 @@
     by (simp cong: indep_sets_cong)
 qed
 
-definition (in prob_space) terminal_events where
-  "terminal_events A = (\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
+definition (in prob_space) tail_events where
+  "tail_events A = (\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
 
-lemma (in prob_space) terminal_events_sets:
-  assumes A: "\<And>i. A i \<subseteq> events"
-  assumes "\<And>i::nat. sigma_algebra (space M) (A i)"
-  assumes X: "X \<in> terminal_events A"
-  shows "X \<in> events"
-proof -
+lemma (in prob_space) tail_events_sets:
+  assumes A: "\<And>i::nat. A i \<subseteq> events"
+  shows "tail_events A \<subseteq> events"
+proof
+  fix X assume X: "X \<in> tail_events A"
   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
-  interpret A: sigma_algebra "space M" "A i" for i by fact
-  from X have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by (auto simp: terminal_events_def)
+  from X have "\<And>n::nat. X \<in> sigma_sets (space M) (UNION {n..} A)" by (auto simp: tail_events_def)
   from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
   then show "X \<in> events"
     by induct (insert A, auto)
 qed
 
-lemma (in prob_space) sigma_algebra_terminal_events:
+lemma (in prob_space) sigma_algebra_tail_events:
   assumes "\<And>i::nat. sigma_algebra (space M) (A i)"
-  shows "sigma_algebra (space M) (terminal_events A)"
-  unfolding terminal_events_def
+  shows "sigma_algebra (space M) (tail_events A)"
+  unfolding tail_events_def
 proof (simp add: sigma_algebra_iff2, safe)
   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
   interpret A: sigma_algebra "space M" "A i" for i by fact
@@ -505,20 +530,22 @@
 
 lemma (in prob_space) kolmogorov_0_1_law:
   fixes A :: "nat \<Rightarrow> 'a set set"
-  assumes A: "\<And>i. A i \<subseteq> events"
   assumes "\<And>i::nat. sigma_algebra (space M) (A i)"
   assumes indep: "indep_sets A UNIV"
-  and X: "X \<in> terminal_events A"
+  and X: "X \<in> tail_events A"
   shows "prob X = 0 \<or> prob X = 1"
 proof -
+  have A: "\<And>i. A i \<subseteq> events"
+    using indep unfolding indep_sets_def by simp
+
   let ?D = "{D \<in> events. prob (X \<inter> D) = prob X * prob D}"
   interpret A: sigma_algebra "space M" "A i" for i by fact
-  interpret T: sigma_algebra "space M" "terminal_events A"
-    by (rule sigma_algebra_terminal_events) fact
+  interpret T: sigma_algebra "space M" "tail_events A"
+    by (rule sigma_algebra_tail_events) fact
   have "X \<subseteq> space M" using T.space_closed X by auto
 
   have X_in: "X \<in> events"
-    by (rule terminal_events_sets) fact+
+    using tail_events_sets A X by auto
 
   interpret D: dynkin_system "space M" ?D
   proof (rule dynkin_systemI)
@@ -583,7 +610,7 @@
     proof (simp add: subset_eq, rule)
       fix D assume D: "D \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
       have "X \<in> sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m)"
-        using X unfolding terminal_events_def by simp
+        using X unfolding tail_events_def by simp
       from indep_setD[OF indep D this] indep_setD_ev1[OF indep] D
       show "D \<in> events \<and> prob (X \<inter> D) = prob X * prob D"
         by (auto simp add: ac_simps)
@@ -591,12 +618,12 @@
   then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> ?D" (is "?A \<subseteq> _")
     by auto
 
-  note `X \<in> terminal_events A`
+  note `X \<in> tail_events A`
   also {
     have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A"
       by (intro sigma_sets_subseteq UN_mono) auto
-   then have "terminal_events A \<subseteq> sigma_sets (space M) ?A"
-      unfolding terminal_events_def by auto }
+   then have "tail_events A \<subseteq> sigma_sets (space M) ?A"
+      unfolding tail_events_def by auto }
   also have "sigma_sets (space M) ?A = dynkin (space M) ?A"
   proof (rule sigma_eq_dynkin)
     { fix B n assume "B \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
@@ -628,34 +655,33 @@
 
 lemma (in prob_space) borel_0_1_law:
   fixes F :: "nat \<Rightarrow> 'a set"
-  assumes F: "range F \<subseteq> events" "indep_events F UNIV"
+  assumes F2: "indep_events F UNIV"
   shows "prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 0 \<or> prob (\<Inter>n. \<Union>m\<in>{n..}. F m) = 1"
 proof (rule kolmogorov_0_1_law[of "\<lambda>i. sigma_sets (space M) { F i }"])
-  show "\<And>i. sigma_sets (space M) {F i} \<subseteq> events"
-    using F(1) sets_into_space
-    by (subst sigma_sets_singleton) auto
+  have F1: "range F \<subseteq> events"
+    using F2 by (simp add: indep_events_def subset_eq)
   { fix i show "sigma_algebra (space M) (sigma_sets (space M) {F i})"
-      using sigma_algebra_sigma_sets[of "{F i}" "space M"] F sets_into_space
+      using sigma_algebra_sigma_sets[of "{F i}" "space M"] F1 sets_into_space
       by auto }
   show "indep_sets (\<lambda>i. sigma_sets (space M) {F i}) UNIV"
   proof (rule indep_sets_sigma)
     show "indep_sets (\<lambda>i. {F i}) UNIV"
-      unfolding indep_sets_singleton_iff_indep_events by fact
+      unfolding indep_events_def_alt[symmetric] by fact
     fix i show "Int_stable {F i}"
       unfolding Int_stable_def by simp
   qed
   let ?Q = "\<lambda>n. \<Union>i\<in>{n..}. F i"
-  show "(\<Inter>n. \<Union>m\<in>{n..}. F m) \<in> terminal_events (\<lambda>i. sigma_sets (space M) {F i})"
-    unfolding terminal_events_def
+  show "(\<Inter>n. \<Union>m\<in>{n..}. F m) \<in> tail_events (\<lambda>i. sigma_sets (space M) {F i})"
+    unfolding tail_events_def
   proof
     fix j
     interpret S: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
-      using order_trans[OF F(1) space_closed]
+      using order_trans[OF F1 space_closed]
       by (intro sigma_algebra_sigma_sets) (simp add: sigma_sets_singleton subset_eq)
     have "(\<Inter>n. ?Q n) = (\<Inter>n\<in>{j..}. ?Q n)"
       by (intro decseq_SucI INT_decseq_offset UN_mono) auto
     also have "\<dots> \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
-      using order_trans[OF F(1) space_closed]
+      using order_trans[OF F1 space_closed]
       by (safe intro!: S.countable_INT S.countable_UN)
          (auto simp: sigma_sets_singleton intro!: sigma_sets.Basic bexI)
     finally show "(\<Inter>n. ?Q n) \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
@@ -872,7 +898,7 @@
       show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')"
         by (simp add: sets_PiM space_PiM cong: prod_algebra_cong)
       let ?A = "\<lambda>i. \<Pi>\<^isub>E i\<in>I. space (M' i)"
-      show "range ?A \<subseteq> prod_algebra I M'" "incseq ?A" "(\<Union>i. ?A i) = space (Pi\<^isub>M I M')"
+      show "range ?A \<subseteq> prod_algebra I M'" "(\<Union>i. ?A i) = space (Pi\<^isub>M I M')"
         by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong)
       { fix i show "emeasure ?D (\<Pi>\<^isub>E i\<in>I. space (M' i)) \<noteq> \<infinity>" by auto }
     next
@@ -1022,13 +1048,13 @@
     then show "sets M = sigma_sets (space ?P) ?E"
       using sets[symmetric] by simp
   next
-    show "range F \<subseteq> ?E" "incseq F" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
+    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: emeasure_pair_measure_Times)
+       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"
@@ -1136,7 +1162,7 @@
       also have "\<dots> = emeasure (?S \<Otimes>\<^isub>M ?T) (A \<times> B)"
         unfolding `?S \<Otimes>\<^isub>M ?T = ?J` ..
       also have "\<dots> = emeasure ?S A * emeasure ?T B"
-        using ab by (simp add: XY.emeasure_pair_measure_Times)
+        using ab by (simp add: Y.emeasure_pair_measure_Times)
       finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
         prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
         using rvs ab by (simp add: emeasure_eq_measure emeasure_distr)
@@ -1144,4 +1170,12 @@
   qed
 qed
 
+lemma (in prob_space) distributed_joint_indep:
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes X: "distributed M S X Px" and Y: "distributed M T Y Py"
+  assumes indep: "indep_var S X T Y"
+  shows "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
+  using indep_var_distribution_eq[of S X T Y] indep
+  by (intro distributed_joint_indep'[OF S T X Y]) auto
+
 end
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -8,41 +8,16 @@
   imports Probability_Measure Caratheodory
 begin
 
-lemma (in product_sigma_finite)
-  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
-  shows emeasure_fold_integral:
-    "emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
-    and emeasure_fold_measurable:
-    "(\<lambda>x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
-proof -
-  interpret I: finite_product_sigma_finite M I by default fact
-  interpret J: finite_product_sigma_finite M J by default fact
-  interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" ..
-  have merge: "(\<lambda>(x, y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>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[OF IJ(1)] A])
-    apply (subst IJ.emeasure_pair_measure_alt[OF merge])
-    apply (auto intro!: positive_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_compose[symmetric] comp_def space_pair_measure cong: measurable_cong)
-qed
-
 lemma restrict_extensional_sub[intro]: "A \<subseteq> B \<Longrightarrow> restrict f A \<in> extensional B"
   unfolding restrict_def extensional_def by auto
 
 lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \<inter> B)"
   unfolding restrict_def by (simp add: fun_eq_iff)
 
-lemma split_merge: "P (merge I x J 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)"
+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 extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I x J y \<in> extensional K"
+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 injective_vimage_restrict:
@@ -57,7 +32,7 @@
   show "x \<in> A \<longleftrightarrow> x \<in> B"
   proof cases
     assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)"
-    have "x \<in> A \<longleftrightarrow> merge J x (I - J) y \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
+    have "x \<in> A \<longleftrightarrow> merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
       using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge)
     then show "x \<in> A \<longleftrightarrow> x \<in> B"
       using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge)
@@ -112,7 +87,7 @@
   let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"
   show "Int_stable ?J"
     by (rule Int_stable_PiE)
-  show "range ?F \<subseteq> ?J" "incseq ?F" "(\<Union>i. ?F i) = ?\<Omega>"
+  show "range ?F \<subseteq> ?J" "(\<Union>i. ?F i) = ?\<Omega>"
     using `finite J` by (auto intro!: prod_algebraI_finite)
   { fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp }
   show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets_into_space)
@@ -156,7 +131,7 @@
   show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
     using sets[THEN sets_into_space] by (auto simp: space_PiM)
   have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
-    using M.not_empty by auto
+      using M.not_empty by auto
   from bchoice[OF this]
   show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto
   show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))"
@@ -199,15 +174,24 @@
 qed
 
 lemma (in product_prob_space) sets_PiM_generator:
-  assumes "I \<noteq> {}" shows "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
-proof
-  show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
-    unfolding sets_PiM
-  proof (safe intro!: sigma_sets_subseteq)
-    fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
-      by (auto intro!: generatorI' elim!: prod_algebraE)
-  qed
-qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset)
+  "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
+proof cases
+  assume "I = {}" then show ?thesis
+    unfolding generator_def
+    by (auto simp: sets_PiM_empty sigma_sets_empty_eq cong: conj_cong)
+next
+  assume "I \<noteq> {}"
+  show ?thesis
+  proof
+    show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
+      unfolding sets_PiM
+    proof (safe intro!: sigma_sets_subseteq)
+      fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
+        by (auto intro!: generatorI' elim!: prod_algebraE)
+    qed
+  qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset)
+qed
+
 
 lemma (in product_prob_space) generatorI:
   "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
@@ -255,23 +239,19 @@
 qed
 
 lemma (in product_prob_space) merge_sets:
-  assumes "finite J" "finite K" "J \<inter> K = {}" and A: "A \<in> sets (Pi\<^isub>M (J \<union> K) M)" and x: "x \<in> space (Pi\<^isub>M J M)"
-  shows "merge J x K -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
-proof -
-  from sets_Pair1[OF
-    measurable_merge[THEN measurable_sets, OF `J \<inter> K = {}`], OF A, of x] x
-  show ?thesis
-      by (simp add: space_pair_measure comp_def vimage_compose[symmetric])
-qed
+  assumes "J \<inter> K = {}" and A: "A \<in> sets (Pi\<^isub>M (J \<union> K) M)" and x: "x \<in> space (Pi\<^isub>M J M)"
+  shows "(\<lambda>y. merge J K (x,y)) -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
+  by (rule measurable_sets[OF _ A] measurable_compose[OF measurable_Pair measurable_merge]  
+           measurable_const x measurable_ident)+
 
 lemma (in product_prob_space) merge_emb:
   assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)"
-  shows "(merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =
-    emb I (K - J) (merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"
+  shows "((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =
+    emb I (K - J) ((\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"
 proof -
-  have [simp]: "\<And>x J K L. merge J y K (restrict x L) = merge J y (K \<inter> L) x"
+  have [simp]: "\<And>x J K L. merge J K (y, restrict x L) = merge J (K \<inter> L) (y, x)"
     by (auto simp: restrict_def merge_def)
-  have [simp]: "\<And>x J K L. restrict (merge J y K x) L = merge (J \<inter> L) y (K \<inter> L) x"
+  have [simp]: "\<And>x J K L. restrict (merge J K (y, x)) L = merge (J \<inter> L) (K \<inter> L) (y, x)"
     by (auto simp: restrict_def merge_def)
   have [simp]: "(I - J) \<inter> K = K - J" using `K \<subseteq> I` `J \<subseteq> I` by auto
   have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto
@@ -356,16 +336,16 @@
       "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X"
       by (auto simp: subset_insertI)
 
-    let ?M = "\<lambda>y. merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
+    let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
     { fix y assume y: "y \<in> space (Pi\<^isub>M J M)"
       note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
       moreover
       have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)"
         using J K y by (intro merge_sets) auto
       ultimately
-      have ***: "(merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G"
+      have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G"
         using J K by (intro generatorI) auto
-      have "\<mu>G (merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)"
+      have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)"
         unfolding * using K J by (subst \<mu>G_eq[OF _ _ _ **]) auto
       note * ** *** this }
     note merge_in_G = this
@@ -379,7 +359,7 @@
       using K J by simp
     also have "\<dots> = (\<integral>\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"
       using K J by (subst emeasure_fold_integral) auto
-    also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"
+    also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"
       (is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)")
     proof (intro positive_integral_cong)
       fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
@@ -395,7 +375,7 @@
         by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }
     note le_1 = this
 
-    let ?q = "\<lambda>y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M))"
+    let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^isub>M I M))"
     have "?q \<in> borel_measurable (Pi\<^isub>M J M)"
       unfolding `Z = emb I K X` using J K merge_in_G(3)
       by (simp add: merge_in_G  \<mu>G_eq emeasure_fold_measurable cong: measurable_cong)
@@ -441,7 +421,7 @@
         using \<mu>G_spec[of "J 0" "A 0" "X 0"] J A_eq
         by (auto intro!: INF_lower2[of 0] J.measure_le_1)
 
-      let ?M = "\<lambda>K Z y. merge K y (I - K) -` Z \<inter> space (Pi\<^isub>M I M)"
+      let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^isub>M I M)"
 
       { fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
         then have Z_sets: "\<And>n. Z n \<in> ?G" by auto
@@ -549,9 +529,9 @@
             from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
             obtain w' where w': "w' \<in> space (Pi\<^isub>M ?D M)"
               "\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto
-            let ?w = "merge (J k) (w k) ?D w'"
-            have [simp]: "\<And>x. merge (J k) (w k) (I - J k) (merge ?D w' (I - ?D) x) =
-              merge (J (Suc k)) ?w (I - (J (Suc k))) x"
+            let ?w = "merge (J k) ?D (w k, w')"
+            have [simp]: "\<And>x. merge (J k) (I - J k) (w k, merge ?D (I - ?D) (w', x)) =
+              merge (J (Suc k)) (I - (J (Suc k))) (?w, x)"
               using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"]
               by (auto intro!: ext split: split_merge)
             have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w"
@@ -577,7 +557,7 @@
           using positive_\<mu>G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1`
           by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
         then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto
-        then have "merge (J k) (w k) (I - J k) x \<in> A k" by auto
+        then have "merge (J k) (I - J k) (w k, x) \<in> A k" by auto
         then have "\<exists>x\<in>A k. restrict x (J k) = w k"
           using `w k \<in> space (Pi\<^isub>M (J k) M)`
           by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM)
@@ -657,7 +637,7 @@
       using X by (auto simp add: emeasure_PiM) 
   next
     show "positive (sets (Pi\<^isub>M I M)) \<mu>" "countably_additive (sets (Pi\<^isub>M I M)) \<mu>"
-      using \<mu> unfolding sets_PiM_generator[OF `I \<noteq> {}`] by (auto simp: measure_space_def)
+      using \<mu> unfolding sets_PiM_generator by (auto simp: measure_space_def)
   qed
 qed
 
--- a/src/HOL/Probability/Information.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Information.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -22,104 +22,6 @@
   "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)"
   unfolding setsum_cartesian_product by simp
 
-section "Convex theory"
-
-lemma log_setsum:
-  assumes "finite s" "s \<noteq> {}"
-  assumes "b > 1"
-  assumes "(\<Sum> i \<in> s. a i) = 1"
-  assumes "\<And> i. i \<in> s \<Longrightarrow> a i \<ge> 0"
-  assumes "\<And> i. i \<in> s \<Longrightarrow> y i \<in> {0 <..}"
-  shows "log b (\<Sum> i \<in> s. a i * y i) \<ge> (\<Sum> i \<in> s. a i * log b (y i))"
-proof -
-  have "convex_on {0 <..} (\<lambda> x. - log b x)"
-    by (rule minus_log_convex[OF `b > 1`])
-  hence "- log b (\<Sum> i \<in> s. a i * y i) \<le> (\<Sum> i \<in> s. a i * - log b (y i))"
-    using convex_on_setsum[of _ _ "\<lambda> x. - log b x"] assms pos_is_convex by fastforce
-  thus ?thesis by (auto simp add:setsum_negf le_imp_neg_le)
-qed
-
-lemma log_setsum':
-  assumes "finite s" "s \<noteq> {}"
-  assumes "b > 1"
-  assumes "(\<Sum> i \<in> s. a i) = 1"
-  assumes pos: "\<And> i. i \<in> s \<Longrightarrow> 0 \<le> a i"
-          "\<And> i. \<lbrakk> i \<in> s ; 0 < a i \<rbrakk> \<Longrightarrow> 0 < y i"
-  shows "log b (\<Sum> i \<in> s. a i * y i) \<ge> (\<Sum> i \<in> s. a i * log b (y i))"
-proof -
-  have "\<And>y. (\<Sum> i \<in> s - {i. a i = 0}. a i * y i) = (\<Sum> i \<in> s. a i * y i)"
-    using assms by (auto intro!: setsum_mono_zero_cong_left)
-  moreover have "log b (\<Sum> i \<in> s - {i. a i = 0}. a i * y i) \<ge> (\<Sum> i \<in> s - {i. a i = 0}. a i * log b (y i))"
-  proof (rule log_setsum)
-    have "setsum a (s - {i. a i = 0}) = setsum a s"
-      using assms(1) by (rule setsum_mono_zero_cong_left) auto
-    thus sum_1: "setsum a (s - {i. a i = 0}) = 1"
-      "finite (s - {i. a i = 0})" using assms by simp_all
-
-    show "s - {i. a i = 0} \<noteq> {}"
-    proof
-      assume *: "s - {i. a i = 0} = {}"
-      hence "setsum a (s - {i. a i = 0}) = 0" by (simp add: * setsum_empty)
-      with sum_1 show False by simp
-    qed
-
-    fix i assume "i \<in> s - {i. a i = 0}"
-    hence "i \<in> s" "a i \<noteq> 0" by simp_all
-    thus "0 \<le> a i" "y i \<in> {0<..}" using pos[of i] by auto
-  qed fact+
-  ultimately show ?thesis by simp
-qed
-
-lemma log_setsum_divide:
-  assumes "finite S" and "S \<noteq> {}" and "1 < b"
-  assumes "(\<Sum>x\<in>S. g x) = 1"
-  assumes pos: "\<And>x. x \<in> S \<Longrightarrow> g x \<ge> 0" "\<And>x. x \<in> S \<Longrightarrow> f x \<ge> 0"
-  assumes g_pos: "\<And>x. \<lbrakk> x \<in> S ; 0 < g x \<rbrakk> \<Longrightarrow> 0 < f x"
-  shows "- (\<Sum>x\<in>S. g x * log b (g x / f x)) \<le> log b (\<Sum>x\<in>S. f x)"
-proof -
-  have log_mono: "\<And>x y. 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log b x \<le> log b y"
-    using `1 < b` by (subst log_le_cancel_iff) auto
-
-  have "- (\<Sum>x\<in>S. g x * log b (g x / f x)) = (\<Sum>x\<in>S. g x * log b (f x / g x))"
-  proof (unfold setsum_negf[symmetric], rule setsum_cong)
-    fix x assume x: "x \<in> S"
-    show "- (g x * log b (g x / f x)) = g x * log b (f x / g x)"
-    proof (cases "g x = 0")
-      case False
-      with pos[OF x] g_pos[OF x] have "0 < f x" "0 < g x" by simp_all
-      thus ?thesis using `1 < b` by (simp add: log_divide field_simps)
-    qed simp
-  qed rule
-  also have "... \<le> log b (\<Sum>x\<in>S. g x * (f x / g x))"
-  proof (rule log_setsum')
-    fix x assume x: "x \<in> S" "0 < g x"
-    with g_pos[OF x] show "0 < f x / g x" by (safe intro!: divide_pos_pos)
-  qed fact+
-  also have "... = log b (\<Sum>x\<in>S - {x. g x = 0}. f x)" using `finite S`
-    by (auto intro!: setsum_mono_zero_cong_right arg_cong[where f="log b"]
-        split: split_if_asm)
-  also have "... \<le> log b (\<Sum>x\<in>S. f x)"
-  proof (rule log_mono)
-    have "0 = (\<Sum>x\<in>S - {x. g x = 0}. 0)" by simp
-    also have "... < (\<Sum>x\<in>S - {x. g x = 0}. f x)" (is "_ < ?sum")
-    proof (rule setsum_strict_mono)
-      show "finite (S - {x. g x = 0})" using `finite S` by simp
-      show "S - {x. g x = 0} \<noteq> {}"
-      proof
-        assume "S - {x. g x = 0} = {}"
-        hence "(\<Sum>x\<in>S. g x) = 0" by (subst setsum_0') auto
-        with `(\<Sum>x\<in>S. g x) = 1` show False by simp
-      qed
-      fix x assume "x \<in> S - {x. g x = 0}"
-      thus "0 < f x" using g_pos[of x] pos(1)[of x] by auto
-    qed
-    finally show "0 < ?sum" .
-    show "(\<Sum>x\<in>S - {x. g x = 0}. f x) \<le> (\<Sum>x\<in>S. f x)"
-      using `finite S` pos by (auto intro!: setsum_mono2)
-  qed
-  finally show ?thesis .
-qed
-
 lemma split_pairs:
   "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
   "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto
@@ -194,7 +96,7 @@
   unfolding KL_divergence_def
 proof (subst integral_density)
   show "entropy_density b M (density M (\<lambda>x. ereal (f x))) \<in> borel_measurable M"
-    using f `1 < b`
+    using f
     by (auto simp: comp_def entropy_density_def intro!: borel_measurable_log borel_measurable_RN_deriv_density)
   have "density M (RN_deriv M (density M f)) = density M f"
     using f by (intro density_RN_deriv_density) auto
@@ -437,6 +339,115 @@
   finally show ?thesis .
 qed
 
+subsection {* Finite Entropy *}
+
+definition (in information_space) 
+  "finite_entropy S X f \<longleftrightarrow> distributed M S X f \<and> integrable S (\<lambda>x. f x * log b (f x))"
+
+lemma (in information_space) finite_entropy_simple_function:
+  assumes X: "simple_function M X"
+  shows "finite_entropy (count_space (X`space M)) X (\<lambda>a. measure M {x \<in> space M. X x = a})"
+  unfolding finite_entropy_def
+proof
+  have [simp]: "finite (X ` space M)"
+    using X by (auto simp: simple_function_def)
+  then show "integrable (count_space (X ` space M))
+     (\<lambda>x. prob {xa \<in> space M. X xa = x} * log b (prob {xa \<in> space M. X xa = x}))"
+    by (rule integrable_count_space)
+  have d: "distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (if x \<in> X`space M then prob {xa \<in> space M. X xa = x} else 0))"
+    by (rule distributed_simple_function_superset[OF X]) (auto intro!: arg_cong[where f=prob])
+  show "distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (prob {xa \<in> space M. X xa = x}))"
+    by (rule distributed_cong_density[THEN iffD1, OF _ _ _ d]) auto
+qed
+
+lemma distributed_transform_AE:
+  assumes T: "T \<in> measurable P Q" "absolutely_continuous Q (distr P Q T)"
+  assumes g: "distributed M Q Y g"
+  shows "AE x in P. 0 \<le> g (T x)"
+  using g
+  apply (subst AE_distr_iff[symmetric, OF T(1)])
+  apply (simp add: distributed_borel_measurable)
+  apply (rule absolutely_continuous_AE[OF _ T(2)])
+  apply simp
+  apply (simp add: distributed_AE)
+  done
+
+lemma ac_fst:
+  assumes "sigma_finite_measure T"
+  shows "absolutely_continuous S (distr (S \<Otimes>\<^isub>M T) S fst)"
+proof -
+  interpret sigma_finite_measure T by fact
+  { fix A assume "A \<in> sets S" "emeasure S A = 0"
+    moreover then have "fst -` A \<inter> space (S \<Otimes>\<^isub>M T) = A \<times> space T"
+      by (auto simp: space_pair_measure dest!: sets_into_space)
+    ultimately have "emeasure (S \<Otimes>\<^isub>M T) (fst -` A \<inter> space (S \<Otimes>\<^isub>M T)) = 0"
+      by (simp add: emeasure_pair_measure_Times) }
+  then show ?thesis
+    unfolding absolutely_continuous_def
+    apply (auto simp: null_sets_distr_iff)
+    apply (auto simp: null_sets_def intro!: measurable_sets)
+    done
+qed
+
+lemma ac_snd:
+  assumes "sigma_finite_measure T"
+  shows "absolutely_continuous T (distr (S \<Otimes>\<^isub>M T) T snd)"
+proof -
+  interpret sigma_finite_measure T by fact
+  { fix A assume "A \<in> sets T" "emeasure T A = 0"
+    moreover then have "snd -` A \<inter> space (S \<Otimes>\<^isub>M T) = space S \<times> A"
+      by (auto simp: space_pair_measure dest!: sets_into_space)
+    ultimately have "emeasure (S \<Otimes>\<^isub>M T) (snd -` A \<inter> space (S \<Otimes>\<^isub>M T)) = 0"
+      by (simp add: emeasure_pair_measure_Times) }
+  then show ?thesis
+    unfolding absolutely_continuous_def
+    apply (auto simp: null_sets_distr_iff)
+    apply (auto simp: null_sets_def intro!: measurable_sets)
+    done
+qed
+
+lemma distributed_integrable:
+  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow>
+    integrable N (\<lambda>x. f x * g x) \<longleftrightarrow> integrable M (\<lambda>x. g (X x))"
+  by (auto simp: distributed_real_measurable distributed_real_AE distributed_measurable
+                    distributed_distr_eq_density[symmetric] integral_density[symmetric] integrable_distr_eq)
+  
+lemma distributed_transform_integrable:
+  assumes Px: "distributed M N X Px"
+  assumes "distributed M P Y Py"
+  assumes Y: "Y = (\<lambda>x. T (X x))" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P"
+  shows "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))"
+proof -
+  have "integrable P (\<lambda>x. Py x * f x) \<longleftrightarrow> integrable M (\<lambda>x. f (Y x))"
+    by (rule distributed_integrable) fact+
+  also have "\<dots> \<longleftrightarrow> integrable M (\<lambda>x. f (T (X x)))"
+    using Y by simp
+  also have "\<dots> \<longleftrightarrow> integrable N (\<lambda>x. Px x * f (T x))"
+    using measurable_comp[OF T f] Px by (intro distributed_integrable[symmetric]) (auto simp: comp_def)
+  finally show ?thesis .
+qed
+
+lemma integrable_cong_AE_imp: "integrable M g \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (AE x in M. g x = f x) \<Longrightarrow> integrable M f"
+  using integrable_cong_AE by blast
+
+lemma (in information_space) finite_entropy_integrable:
+  "finite_entropy S X Px \<Longrightarrow> integrable S (\<lambda>x. Px x * log b (Px x))"
+  unfolding finite_entropy_def by auto
+
+lemma (in information_space) finite_entropy_distributed:
+  "finite_entropy S X Px \<Longrightarrow> distributed M S X Px"
+  unfolding finite_entropy_def by auto
+
+lemma (in information_space) finite_entropy_integrable_transform:
+  assumes Fx: "finite_entropy S X Px"
+  assumes Fy: "distributed M T Y Py"
+    and "X = (\<lambda>x. f (Y x))"
+    and "f \<in> measurable T S"
+  shows "integrable T (\<lambda>x. Py x * log b (Px (f x)))"
+  using assms unfolding finite_entropy_def
+  using distributed_transform_integrable[of M T Y Py S X Px f "\<lambda>x. log b (Px x)"]
+  by (auto intro: distributed_real_measurable)
+
 subsection {* Mutual Information *}
 
 definition (in prob_space)
@@ -510,6 +521,120 @@
 
 lemma (in information_space)
   fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Fx: "finite_entropy S X Px" and Fy: "finite_entropy T Y Py"
+  assumes Fxy: "finite_entropy (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
+  shows mutual_information_distr': "mutual_information b S T X Y = integral\<^isup>L (S \<Otimes>\<^isub>M T) f" (is "?M = ?R")
+    and mutual_information_nonneg': "0 \<le> mutual_information b S T X Y"
+proof -
+  have Px: "distributed M S X Px"
+    using Fx by (auto simp: finite_entropy_def)
+  have Py: "distributed M T Y Py"
+    using Fy by (auto simp: finite_entropy_def)
+  have Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+    using Fxy by (auto simp: finite_entropy_def)
+
+  have X: "random_variable S X"
+    using Px by (auto simp: distributed_def finite_entropy_def)
+  have Y: "random_variable T Y"
+    using Py by (auto simp: distributed_def finite_entropy_def)
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T ..
+  interpret X: prob_space "distr M S X" using X by (rule prob_space_distr)
+  interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr)
+  interpret XY: pair_prob_space "distr M S X" "distr M T Y" ..
+  let ?P = "S \<Otimes>\<^isub>M T"
+  let ?D = "distr M ?P (\<lambda>x. (X x, Y x))"
+
+  { fix A assume "A \<in> sets S"
+    with X Y have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)"
+      by (auto simp: emeasure_distr measurable_Pair measurable_space
+               intro!: arg_cong[where f="emeasure M"]) }
+  note marginal_eq1 = this
+  { fix A assume "A \<in> sets T"
+    with X Y have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)"
+      by (auto simp: emeasure_distr measurable_Pair measurable_space
+               intro!: arg_cong[where f="emeasure M"]) }
+  note marginal_eq2 = this
+
+  have eq: "(\<lambda>x. ereal (Px (fst x) * Py (snd x))) = (\<lambda>(x, y). ereal (Px x) * ereal (Py y))"
+    by auto
+
+  have distr_eq: "distr M S X \<Otimes>\<^isub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
+    unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq
+  proof (subst pair_measure_density)
+    show "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ereal (Py y)) \<in> borel_measurable T"
+      "AE x in S. 0 \<le> ereal (Px x)" "AE y in T. 0 \<le> ereal (Py y)"
+      using Px Py by (auto simp: distributed_def)
+    show "sigma_finite_measure (density S Px)" unfolding Px(1)[THEN distributed_distr_eq_density, symmetric] ..
+    show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] ..
+  qed (fact | simp)+
+  
+  have M: "?M = KL_divergence b (density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ereal (Pxy x)))"
+    unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] ..
+
+  from Px Py have f: "(\<lambda>x. Px (fst x) * Py (snd x)) \<in> borel_measurable ?P"
+    by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'')
+  have PxPy_nonneg: "AE x in ?P. 0 \<le> Px (fst x) * Py (snd x)"
+  proof (rule ST.AE_pair_measure)
+    show "{x \<in> space ?P. 0 \<le> Px (fst x) * Py (snd x)} \<in> sets ?P"
+      using f by auto
+    show "AE x in S. AE y in T. 0 \<le> Px (fst (x, y)) * Py (snd (x, y))"
+      using Px Py by (auto simp: zero_le_mult_iff dest!: distributed_real_AE)
+  qed
+
+  have "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)"
+    by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
+  moreover
+  have "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
+    by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
+  ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+    by eventually_elim auto
+
+  show "?M = ?R"
+    unfolding M f_def
+    using b_gt_1 f PxPy_nonneg Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] ac
+    by (rule ST.KL_density_density)
+
+  have X: "X = fst \<circ> (\<lambda>x. (X x, Y x))" and Y: "Y = snd \<circ> (\<lambda>x. (X x, Y x))"
+    by auto
+
+  have "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)))"
+    using finite_entropy_integrable[OF Fxy]
+    using finite_entropy_integrable_transform[OF Fx Pxy, of fst]
+    using finite_entropy_integrable_transform[OF Fy Pxy, of snd]
+    by simp
+  moreover have "f \<in> borel_measurable (S \<Otimes>\<^isub>M T)"
+    unfolding f_def using Px Py Pxy
+    by (auto intro: distributed_real_measurable measurable_fst'' measurable_snd''
+      intro!: borel_measurable_times borel_measurable_log borel_measurable_divide)
+  ultimately have int: "integrable (S \<Otimes>\<^isub>M T) f"
+    apply (rule integrable_cong_AE_imp)
+    using
+      distributed_transform_AE[OF measurable_fst ac_fst, of T, OF T Px]
+      distributed_transform_AE[OF measurable_snd ac_snd, of _ _ _ _ S, OF T Py]
+      subdensity_real[OF measurable_fst Pxy Px X]
+      subdensity_real[OF measurable_snd Pxy Py Y]
+      distributed_real_AE[OF Pxy]
+    by eventually_elim
+       (auto simp: f_def log_divide_eq log_mult_eq field_simps zero_less_mult_iff mult_nonneg_nonneg)
+
+  show "0 \<le> ?M" unfolding M
+  proof (rule ST.KL_density_density_nonneg
+    [OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]])
+    show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x))) "
+      unfolding distributed_distr_eq_density[OF Pxy, symmetric]
+      using distributed_measurable[OF Pxy] by (rule prob_space_distr)
+    show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
+      unfolding distr_eq[symmetric] by unfold_locales
+  qed
+qed
+
+
+lemma (in information_space)
+  fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
   assumes "sigma_finite_measure S" "sigma_finite_measure T"
   assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
   assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
@@ -664,84 +789,209 @@
   entropy_Pow ("\<H>'(_')") where
   "\<H>(X) \<equiv> entropy b (count_space (X`space M)) X"
 
-lemma (in information_space) entropy_distr:
+lemma (in prob_space) distributed_RN_deriv:
+  assumes X: "distributed M S X Px"
+  shows "AE x in S. RN_deriv S (density S Px) x = Px x"
+proof -
+  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X]
+  interpret X: prob_space "distr M S X"
+    using D(1) by (rule prob_space_distr)
+
+  have sf: "sigma_finite_measure (distr M S X)" by default
+  show ?thesis
+    using D
+    apply (subst eq_commute)
+    apply (intro RN_deriv_unique_sigma_finite)
+    apply (auto intro: divide_nonneg_nonneg measure_nonneg
+             simp: distributed_distr_eq_density[symmetric, OF X] sf)
+    done
+qed
+
+lemma (in information_space)
   fixes X :: "'a \<Rightarrow> 'b"
-  assumes "sigma_finite_measure MX" and X: "distributed M MX X f"
-  shows "entropy b MX X = - (\<integral>x. f x * log b (f x) \<partial>MX)"
+  assumes X: "distributed M MX X f"
+  shows entropy_distr: "entropy b MX X = - (\<integral>x. f x * log b (f x) \<partial>MX)" (is ?eq)
+proof -
+  note D = distributed_measurable[OF X] distributed_borel_measurable[OF X] distributed_AE[OF X]
+  note ae = distributed_RN_deriv[OF X]
+
+  have ae_eq: "AE x in distr M MX X. log b (real (RN_deriv MX (distr M MX X) x)) =
+    log b (f x)"
+    unfolding distributed_distr_eq_density[OF X]
+    apply (subst AE_density)
+    using D apply simp
+    using ae apply eventually_elim
+    apply auto
+    done
+
+  have int_eq: "- (\<integral> x. log b (f x) \<partial>distr M MX X) = - (\<integral> x. f x * log b (f x) \<partial>MX)"
+    unfolding distributed_distr_eq_density[OF X]
+    using D
+    by (subst integral_density)
+       (auto simp: borel_measurable_ereal_iff)
+
+  show ?eq
+    unfolding entropy_def KL_divergence_def entropy_density_def comp_def
+    apply (subst integral_cong_AE)
+    apply (rule ae_eq)
+    apply (rule int_eq)
+    done
+qed
+
+lemma (in prob_space) distributed_imp_emeasure_nonzero:
+  assumes X: "distributed M MX X Px"
+  shows "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> 0"
+proof
+  note Px = distributed_borel_measurable[OF X] distributed_AE[OF X]
+  interpret X: prob_space "distr M MX X"
+    using distributed_measurable[OF X] by (rule prob_space_distr)
+
+  assume "emeasure MX {x \<in> space MX. Px x \<noteq> 0} = 0"
+  with Px have "AE x in MX. Px x = 0"
+    by (intro AE_I[OF subset_refl]) (auto simp: borel_measurable_ereal_iff)
+  moreover
+  from X.emeasure_space_1 have "(\<integral>\<^isup>+x. Px x \<partial>MX) = 1"
+    unfolding distributed_distr_eq_density[OF X] using Px
+    by (subst (asm) emeasure_density)
+       (auto simp: borel_measurable_ereal_iff intro!: integral_cong cong: positive_integral_cong)
+  ultimately show False
+    by (simp add: positive_integral_cong_AE)
+qed
+
+lemma (in information_space) entropy_le:
+  fixes Px :: "'b \<Rightarrow> real" and MX :: "'b measure"
+  assumes X: "distributed M MX X Px"
+  and fin: "emeasure MX {x \<in> space MX. Px x \<noteq> 0} \<noteq> \<infinity>"
+  and int: "integrable MX (\<lambda>x. - Px x * log b (Px x))"
+  shows "entropy b MX X \<le> log b (measure MX {x \<in> space MX. Px x \<noteq> 0})"
 proof -
-  interpret MX: sigma_finite_measure MX by fact
-  from X show ?thesis
-    unfolding entropy_def X[THEN distributed_distr_eq_density]
-    by (subst MX.KL_density[OF b_gt_1]) (simp_all add: distributed_real_AE distributed_real_measurable)
+  note Px = distributed_borel_measurable[OF X] distributed_AE[OF X]
+  interpret X: prob_space "distr M MX X"
+    using distributed_measurable[OF X] by (rule prob_space_distr)
+
+  have " - log b (measure MX {x \<in> space MX. Px x \<noteq> 0}) = 
+    - log b (\<integral> x. indicator {x \<in> space MX. Px x \<noteq> 0} x \<partial>MX)"
+    using Px fin
+    by (subst integral_indicator) (auto simp: measure_def borel_measurable_ereal_iff)
+  also have "- log b (\<integral> x. indicator {x \<in> space MX. Px x \<noteq> 0} x \<partial>MX) = - log b (\<integral> x. 1 / Px x \<partial>distr M MX X)"
+    unfolding distributed_distr_eq_density[OF X] using Px
+    apply (intro arg_cong[where f="log b"] arg_cong[where f=uminus])
+    by (subst integral_density) (auto simp: borel_measurable_ereal_iff intro!: integral_cong)
+  also have "\<dots> \<le> (\<integral> x. - log b (1 / Px x) \<partial>distr M MX X)"
+  proof (rule X.jensens_inequality[of "\<lambda>x. 1 / Px x" "{0<..}" 0 1 "\<lambda>x. - log b x"])
+    show "AE x in distr M MX X. 1 / Px x \<in> {0<..}"
+      unfolding distributed_distr_eq_density[OF X]
+      using Px by (auto simp: AE_density)
+    have [simp]: "\<And>x. x \<in> space MX \<Longrightarrow> ereal (if Px x = 0 then 0 else 1) = indicator {x \<in> space MX. Px x \<noteq> 0} x"
+      by (auto simp: one_ereal_def)
+    have "(\<integral>\<^isup>+ x. max 0 (ereal (- (if Px x = 0 then 0 else 1))) \<partial>MX) = (\<integral>\<^isup>+ x. 0 \<partial>MX)"
+      by (intro positive_integral_cong) (auto split: split_max)
+    then show "integrable (distr M MX X) (\<lambda>x. 1 / Px x)"
+      unfolding distributed_distr_eq_density[OF X] using Px
+      by (auto simp: positive_integral_density integrable_def borel_measurable_ereal_iff fin positive_integral_max_0
+              cong: positive_integral_cong)
+    have "integrable MX (\<lambda>x. Px x * log b (1 / Px x)) =
+      integrable MX (\<lambda>x. - Px x * log b (Px x))"
+      using Px
+      by (intro integrable_cong_AE)
+         (auto simp: borel_measurable_ereal_iff log_divide_eq
+                  intro!: measurable_If)
+    then show "integrable (distr M MX X) (\<lambda>x. - log b (1 / Px x))"
+      unfolding distributed_distr_eq_density[OF X]
+      using Px int
+      by (subst integral_density) (auto simp: borel_measurable_ereal_iff)
+  qed (auto simp: minus_log_convex[OF b_gt_1])
+  also have "\<dots> = (\<integral> x. log b (Px x) \<partial>distr M MX X)"
+    unfolding distributed_distr_eq_density[OF X] using Px
+    by (intro integral_cong_AE) (auto simp: AE_density log_divide_eq)
+  also have "\<dots> = - entropy b MX X"
+    unfolding distributed_distr_eq_density[OF X] using Px
+    by (subst entropy_distr[OF X]) (auto simp: borel_measurable_ereal_iff integral_density)
+  finally show ?thesis
+    by simp
+qed
+
+lemma (in information_space) entropy_le_space:
+  fixes Px :: "'b \<Rightarrow> real" and MX :: "'b measure"
+  assumes X: "distributed M MX X Px"
+  and fin: "finite_measure MX"
+  and int: "integrable MX (\<lambda>x. - Px x * log b (Px x))"
+  shows "entropy b MX X \<le> log b (measure MX (space MX))"
+proof -
+  note Px = distributed_borel_measurable[OF X] distributed_AE[OF X]
+  interpret finite_measure MX by fact
+  have "entropy b MX X \<le> log b (measure MX {x \<in> space MX. Px x \<noteq> 0})"
+    using int X by (intro entropy_le) auto
+  also have "\<dots> \<le> log b (measure MX (space MX))"
+    using Px distributed_imp_emeasure_nonzero[OF X]
+    by (intro log_le)
+       (auto intro!: borel_measurable_ereal_iff finite_measure_mono b_gt_1
+                     less_le[THEN iffD2] measure_nonneg simp: emeasure_eq_measure)
+  finally show ?thesis .
+qed
+
+lemma (in prob_space) uniform_distributed_params:
+  assumes X: "distributed M MX X (\<lambda>x. indicator A x / measure MX A)"
+  shows "A \<in> sets MX" "measure MX A \<noteq> 0"
+proof -
+  interpret X: prob_space "distr M MX X"
+    using distributed_measurable[OF X] by (rule prob_space_distr)
+
+  show "measure MX A \<noteq> 0"
+  proof
+    assume "measure MX A = 0"
+    with X.emeasure_space_1 X.prob_space distributed_distr_eq_density[OF X]
+    show False
+      by (simp add: emeasure_density zero_ereal_def[symmetric])
+  qed
+  with measure_notin_sets[of A MX] show "A \<in> sets MX"
+    by blast
 qed
 
 lemma (in information_space) entropy_uniform:
-  assumes "sigma_finite_measure MX"
-  assumes A: "A \<in> sets MX" "emeasure MX A \<noteq> 0" "emeasure MX A \<noteq> \<infinity>"
-  assumes X: "distributed M MX X (\<lambda>x. 1 / measure MX A * indicator A x)"
+  assumes X: "distributed M MX X (\<lambda>x. indicator A x / measure MX A)" (is "distributed _ _ _ ?f")
   shows "entropy b MX X = log b (measure MX A)"
-proof (subst entropy_distr[OF _ X])
-  let ?f = "\<lambda>x. 1 / measure MX A * indicator A x"
-  have "- (\<integral>x. ?f x * log b (?f x) \<partial>MX) = 
-    - (\<integral>x. (log b (1 / measure MX A) / measure MX A) * indicator A x \<partial>MX)"
-    by (auto intro!: integral_cong simp: indicator_def)
-  also have "\<dots> = - log b (inverse (measure MX A))"
-    using A by (subst integral_cmult(2))
-               (simp_all add: measure_def real_of_ereal_eq_0 integral_cmult inverse_eq_divide)
-  also have "\<dots> = log b (measure MX A)"
-    using b_gt_1 A by (subst log_inverse) (auto simp add: measure_def less_le real_of_ereal_eq_0
-                                                          emeasure_nonneg real_of_ereal_pos)
-  finally show "- (\<integral>x. ?f x * log b (?f x) \<partial>MX) = log b (measure MX A)" by simp
-qed fact+
+proof (subst entropy_distr[OF X])
+  have [simp]: "emeasure MX A \<noteq> \<infinity>"
+    using uniform_distributed_params[OF X] by (auto simp add: measure_def)
+  have eq: "(\<integral> x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \<partial>MX) =
+    (\<integral> x. (- log b (measure MX A) / measure MX A) * indicator A x \<partial>MX)"
+    using measure_nonneg[of MX A] uniform_distributed_params[OF X]
+    by (auto intro!: integral_cong split: split_indicator simp: log_divide_eq)
+  show "- (\<integral> x. indicator A x / measure MX A * log b (indicator A x / measure MX A) \<partial>MX) =
+    log b (measure MX A)"
+    unfolding eq using uniform_distributed_params[OF X]
+    by (subst lebesgue_integral_cmult) (auto simp: measure_def)
+qed
 
 lemma (in information_space) entropy_simple_distributed:
-  fixes X :: "'a \<Rightarrow> 'b"
-  assumes X: "simple_distributed M X f"
-  shows "\<H>(X) = - (\<Sum>x\<in>X`space M. f x * log b (f x))"
-proof (subst entropy_distr[OF _ simple_distributed[OF X]])
-  show "sigma_finite_measure (count_space (X ` space M))"
-    using X by (simp add: sigma_finite_measure_count_space_finite simple_distributed_def)
-  show "- (\<integral>x. f x * log b (f x) \<partial>(count_space (X`space M))) = - (\<Sum>x\<in>X ` space M. f x * log b (f x))"
-    using X by (auto simp add: lebesgue_integral_count_space_finite)
-qed
+  "simple_distributed M X f \<Longrightarrow> \<H>(X) = - (\<Sum>x\<in>X`space M. f x * log b (f x))"
+  by (subst entropy_distr[OF simple_distributed])
+     (auto simp add: lebesgue_integral_count_space_finite)
 
 lemma (in information_space) entropy_le_card_not_0:
   assumes X: "simple_distributed M X f"
   shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. f x \<noteq> 0}))"
 proof -
-  have "\<H>(X) = (\<Sum>x\<in>X`space M. f x * log b (1 / f x))"
-    unfolding entropy_simple_distributed[OF X] setsum_negf[symmetric]
-    using X by (auto dest: simple_distributed_nonneg intro!: setsum_cong simp: log_simps less_le)
-  also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. f x * (1 / f x))"
-    using not_empty b_gt_1 `simple_distributed M X f`
-    by (intro log_setsum') (auto simp: simple_distributed_nonneg simple_distributed_setsum_space)
-  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if f x \<noteq> 0 then 1 else 0)"
-    by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) auto
-  finally show ?thesis
-    using `simple_distributed M X f` by (auto simp: setsum_cases real_eq_of_nat)
+  let ?X = "count_space (X`space M)"
+  have "\<H>(X) \<le> log b (measure ?X {x \<in> space ?X. f x \<noteq> 0})"
+    by (rule entropy_le[OF simple_distributed[OF X]])
+       (simp_all add: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space)
+  also have "measure ?X {x \<in> space ?X. f x \<noteq> 0} = card (X ` space M \<inter> {x. f x \<noteq> 0})"
+    by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def Int_def)
+  finally show ?thesis .
 qed
 
 lemma (in information_space) entropy_le_card:
-  assumes "simple_distributed M X f"
+  assumes X: "simple_distributed M X f"
   shows "\<H>(X) \<le> log b (real (card (X ` space M)))"
-proof cases
-  assume "X ` space M \<inter> {x. f x \<noteq> 0} = {}"
-  then have "\<And>x. x\<in>X`space M \<Longrightarrow> f x = 0" by auto
-  moreover
-  have "0 < card (X`space M)"
-    using `simple_distributed M X f` not_empty by (auto simp: card_gt_0_iff)
-  then have "log b 1 \<le> log b (real (card (X`space M)))"
-    using b_gt_1 by (intro log_le) auto
-  ultimately show ?thesis using assms by (simp add: entropy_simple_distributed)
-next
-  assume False: "X ` space M \<inter> {x. f x \<noteq> 0} \<noteq> {}"
-  have "card (X ` space M \<inter> {x. f x \<noteq> 0}) \<le> card (X ` space M)"
-    (is "?A \<le> ?B") using assms not_empty
-    by (auto intro!: card_mono simp: simple_function_def simple_distributed_def)
-  note entropy_le_card_not_0[OF assms]
-  also have "log b (real ?A) \<le> log b (real ?B)"
-    using b_gt_1 False not_empty `?A \<le> ?B` assms
-    by (auto intro!: log_le simp: card_gt_0_iff simp: simple_distributed_def)
+proof -
+  let ?X = "count_space (X`space M)"
+  have "\<H>(X) \<le> log b (measure ?X (space ?X))"
+    by (rule entropy_le_space[OF simple_distributed[OF X]])
+       (simp_all add: simple_distributed_finite[OF X] subset_eq integrable_count_space emeasure_count_space finite_measure_count_space)
+  also have "measure ?X (space ?X) = card (X ` space M)"
+    by (simp_all add: simple_distributed_finite[OF X] subset_eq emeasure_count_space measure_def)
   finally show ?thesis .
 qed
 
@@ -757,7 +1007,18 @@
   "\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b
     (count_space (X ` space M)) (count_space (Y ` space M)) (count_space (Z ` space M)) X Y Z"
 
-lemma (in information_space) conditional_mutual_information_generic_eq:
+lemma (in pair_sigma_finite) borel_measurable_positive_integral_fst:
+  "(\<lambda>(x, y). f x y) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>x. \<integral>\<^isup>+ y. f x y \<partial>M2) \<in> borel_measurable M1"
+  using positive_integral_fst_measurable(1)[of "\<lambda>(x, y). f x y"] by simp
+
+lemma (in pair_sigma_finite) borel_measurable_positive_integral_snd:
+  assumes "(\<lambda>(x, y). f x y) \<in> borel_measurable (M2 \<Otimes>\<^isub>M M1)" shows "(\<lambda>x. \<integral>\<^isup>+ y. f x y \<partial>M1) \<in> borel_measurable M2"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 by default
+  from Q.borel_measurable_positive_integral_fst assms show ?thesis by simp
+qed
+
+lemma (in information_space)
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P"
   assumes Px: "distributed M S X Px"
   assumes Pz: "distributed M P Z Pz"
@@ -766,16 +1027,19 @@
   assumes Pxyz: "distributed M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
   assumes I1: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
   assumes I2: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
-  shows "conditional_mutual_information b S T P X Y Z
-    = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+  shows conditional_mutual_information_generic_eq: "conditional_mutual_information b S T P X Y Z
+    = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" (is "?eq")
+    and conditional_mutual_information_generic_nonneg: "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg")
 proof -
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret P: sigma_finite_measure P by fact
   interpret TP: pair_sigma_finite T P ..
   interpret SP: pair_sigma_finite S P ..
+  interpret ST: pair_sigma_finite S T ..
   interpret SPT: pair_sigma_finite "S \<Otimes>\<^isub>M P" T ..
   interpret STP: pair_sigma_finite S "T \<Otimes>\<^isub>M P" ..
+  interpret TPS: pair_sigma_finite "T \<Otimes>\<^isub>M P" S ..
   have TP: "sigma_finite_measure (T \<Otimes>\<^isub>M P)" ..
   have SP: "sigma_finite_measure (S \<Otimes>\<^isub>M P)" ..
   have YZ: "random_variable (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x))"
@@ -811,27 +1075,27 @@
   finally have mi_eq:
     "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" .
   
-  have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
+  have ae1: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
     by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+  moreover have ae2: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
     by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) (auto intro: measurable_snd')
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+  moreover have ae3: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
     by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto intro: measurable_Pair measurable_snd')
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+  moreover have ae4: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
     by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto intro: measurable_Pair)
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Px (fst x)"
+  moreover have ae5: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Px (fst x)"
     using Px by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pyz (snd x)"
+  moreover have ae6: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pyz (snd x)"
     using Pyz by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd (snd x))"
+  moreover have ae7: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd (snd x))"
     using Pz Pz[THEN distributed_real_measurable] by (auto intro!: measurable_snd'' TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
-  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
+  moreover have ae8: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
     using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
     using measurable_comp[OF measurable_Pair[OF measurable_fst measurable_comp[OF measurable_snd measurable_snd]] Pxz[THEN distributed_real_measurable], of T]
     using measurable_comp[OF measurable_snd measurable_Pair2[OF Pxz[THEN distributed_real_measurable]], of _ T]
     by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure simp: comp_def)
   moreover note Pxyz[THEN distributed_real_AE]
-  ultimately have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P.
+  ultimately have ae: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P.
     Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
     Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
     Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
@@ -846,13 +1110,454 @@
         using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps)
     qed simp
   qed
-  with I1 I2 show ?thesis
+  with I1 I2 show ?eq
     unfolding conditional_mutual_information_def
     apply (subst mi_eq)
     apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz])
     apply (subst integral_diff(2)[symmetric])
     apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
     done
+
+  let ?P = "density (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) Pxyz"
+  interpret P: prob_space ?P
+    unfolding distributed_distr_eq_density[OF Pxyz, symmetric]
+    using distributed_measurable[OF Pxyz] by (rule prob_space_distr)
+
+  let ?Q = "density (T \<Otimes>\<^isub>M P) Pyz"
+  interpret Q: prob_space ?Q
+    unfolding distributed_distr_eq_density[OF Pyz, symmetric]
+    using distributed_measurable[OF Pyz] by (rule prob_space_distr)
+
+  let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)"
+
+  from subdensity_real[of snd, OF _ Pyz Pz]
+  have aeX1: "AE x in T \<Otimes>\<^isub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
+  have aeX2: "AE x in T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd x)"
+    using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
+
+  have aeX3: "AE y in T \<Otimes>\<^isub>M P. (\<integral>\<^isup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
+    using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz]
+    apply (intro TP.AE_pair_measure)
+    apply (auto simp: comp_def measurable_split_conv
+                intro!: measurable_snd'' borel_measurable_ereal_eq borel_measurable_ereal
+                        SP.borel_measurable_positive_integral_snd measurable_compose[OF _ Pxz[THEN distributed_real_measurable]]
+                        measurable_Pair
+                dest: distributed_real_AE distributed_real_measurable)
+    done
+
+  note M = borel_measurable_divide borel_measurable_diff borel_measurable_times borel_measurable_ereal
+           measurable_compose[OF _ measurable_snd]
+           measurable_Pair
+           measurable_compose[OF _ Pxyz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Pxz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Pyz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Pz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Px[THEN distributed_real_measurable]]
+           STP.borel_measurable_positive_integral_snd
+  have "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^isup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+    apply (subst positive_integral_density)
+    apply (rule distributed_borel_measurable[OF Pxyz])
+    apply (rule distributed_AE[OF Pxyz])
+    apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) []
+    apply (rule positive_integral_mono_AE)
+    using ae5 ae6 ae7 ae8
+    apply eventually_elim
+    apply (auto intro!: divide_nonneg_nonneg mult_nonneg_nonneg)
+    done
+  also have "\<dots> = (\<integral>\<^isup>+(y, z). \<integral>\<^isup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^isub>M P)"
+    by (subst STP.positive_integral_snd_measurable[symmetric])
+       (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M)
+  also have "\<dots> = (\<integral>\<^isup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^isub>M P)"
+    apply (rule positive_integral_cong_AE)
+    using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space
+    apply eventually_elim
+  proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure)
+    fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
+      "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)" 
+    then show "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
+      apply (subst positive_integral_multc)
+      apply (auto intro!: borel_measurable_ereal divide_nonneg_nonneg
+                          measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] measurable_Pair
+                  split: prod.split)
+      done
+  qed
+  also have "\<dots> = 1"
+    using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz]
+    by (subst positive_integral_density[symmetric]) (auto intro!: M)
+  finally have le1: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> 1" .
+  also have "\<dots> < \<infinity>" by simp
+  finally have fin: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
+
+  have pos: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> 0"
+    apply (subst positive_integral_density)
+    apply (rule distributed_borel_measurable[OF Pxyz])
+    apply (rule distributed_AE[OF Pxyz])
+    apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) []
+    apply (simp add: split_beta')
+  proof
+    let ?g = "\<lambda>x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
+    assume "(\<integral>\<^isup>+ x. ?g x \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)) = 0"
+    then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. ?g x \<le> 0"
+      by (intro positive_integral_0_iff_AE[THEN iffD1]) (auto intro!: M borel_measurable_ereal measurable_If)
+    then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxyz x = 0"
+      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      by eventually_elim (auto split: split_if_asm simp: mult_le_0_iff divide_le_0_iff)
+    then have "(\<integral>\<^isup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) = 0"
+      by (subst positive_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
+    with P.emeasure_space_1 show False
+      by (subst (asm) emeasure_density) (auto intro!: M cong: positive_integral_cong)
+  qed
+
+  have neg: "(\<integral>\<^isup>+ x. - ?f x \<partial>?P) = 0"
+    apply (rule positive_integral_0_iff_AE[THEN iffD2])
+    apply (auto intro!: M simp: split_beta') []
+    apply (subst AE_density)
+    apply (auto intro!: M simp: split_beta') []
+    using ae5 ae6 ae7 ae8
+    apply eventually_elim
+    apply (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg)
+    done
+
+  have I3: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
+    apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]])
+    using ae
+    apply (auto intro!: M simp: split_beta')
+    done
+
+  have "- log b 1 \<le> - log b (integral\<^isup>L ?P ?f)"
+  proof (intro le_imp_neg_le log_le[OF b_gt_1])
+    show "0 < integral\<^isup>L ?P ?f"
+      using neg pos fin positive_integral_positive[of ?P ?f]
+      by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def less_le split_beta')
+    show "integral\<^isup>L ?P ?f \<le> 1"
+      using neg le1 fin positive_integral_positive[of ?P ?f]
+      by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def)
+  qed
+  also have "- log b (integral\<^isup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
+  proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
+    show "AE x in ?P. ?f x \<in> {0<..}"
+      unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
+      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      by eventually_elim (auto simp: divide_pos_pos mult_pos_pos)
+    show "integrable ?P ?f"
+      unfolding integrable_def 
+      using fin neg by (auto intro!: M simp: split_beta')
+    show "integrable ?P (\<lambda>x. - log b (?f x))"
+      apply (subst integral_density)
+      apply (auto intro!: M) []
+      apply (auto intro!: M distributed_real_AE[OF Pxyz]) []
+      apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+      apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3])
+      apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+      apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      apply eventually_elim
+      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps)
+      done
+  qed (auto simp: b_gt_1 minus_log_convex)
+  also have "\<dots> = conditional_mutual_information b S T P X Y Z"
+    unfolding `?eq`
+    apply (subst integral_density)
+    apply (auto intro!: M) []
+    apply (auto intro!: M distributed_real_AE[OF Pxyz]) []
+    apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+    apply (intro integral_cong_AE)
+    using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+    apply eventually_elim
+    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps)
+    done
+  finally show ?nonneg
+    by simp
+qed
+
+lemma (in information_space)
+  fixes Px :: "_ \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P"
+  assumes Fx: "finite_entropy S X Px"
+  assumes Fz: "finite_entropy P Z Pz"
+  assumes Fyz: "finite_entropy (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
+  assumes Fxz: "finite_entropy (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) Pxz"
+  assumes Fxyz: "finite_entropy (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+  shows conditional_mutual_information_generic_eq': "conditional_mutual_information b S T P X Y Z
+    = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" (is "?eq")
+    and conditional_mutual_information_generic_nonneg': "0 \<le> conditional_mutual_information b S T P X Y Z" (is "?nonneg")
+proof -
+  note Px = Fx[THEN finite_entropy_distributed]
+  note Pz = Fz[THEN finite_entropy_distributed]
+  note Pyz = Fyz[THEN finite_entropy_distributed]
+  note Pxz = Fxz[THEN finite_entropy_distributed]
+  note Pxyz = Fxyz[THEN finite_entropy_distributed]
+
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret P: sigma_finite_measure P by fact
+  interpret TP: pair_sigma_finite T P ..
+  interpret SP: pair_sigma_finite S P ..
+  interpret ST: pair_sigma_finite S T ..
+  interpret SPT: pair_sigma_finite "S \<Otimes>\<^isub>M P" T ..
+  interpret STP: pair_sigma_finite S "T \<Otimes>\<^isub>M P" ..
+  interpret TPS: pair_sigma_finite "T \<Otimes>\<^isub>M P" S ..
+  have TP: "sigma_finite_measure (T \<Otimes>\<^isub>M P)" ..
+  have SP: "sigma_finite_measure (S \<Otimes>\<^isub>M P)" ..
+  have YZ: "random_variable (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x))"
+    using Pyz by (simp add: distributed_measurable)
+
+  have Pxyz_f: "\<And>M f. f \<in> measurable M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. Pxyz (f x)) \<in> borel_measurable M"
+    using measurable_comp[OF _ Pxyz[THEN distributed_real_measurable]] by (auto simp: comp_def)
+
+  { fix f g h M
+    assume f: "f \<in> measurable M S" and g: "g \<in> measurable M P" and h: "h \<in> measurable M (S \<Otimes>\<^isub>M P)"
+    from measurable_comp[OF h Pxz[THEN distributed_real_measurable]]
+         measurable_comp[OF f Px[THEN distributed_real_measurable]]
+         measurable_comp[OF g Pz[THEN distributed_real_measurable]]
+    have "(\<lambda>x. log b (Pxz (h x) / (Px (f x) * Pz (g x)))) \<in> borel_measurable M"
+      by (simp add: comp_def b_gt_1) }
+  note borel_log = this
+
+  have measurable_cut: "(\<lambda>(x, y, z). (x, z)) \<in> measurable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (S \<Otimes>\<^isub>M P)"
+    by (auto simp add: split_beta' comp_def intro!: measurable_Pair measurable_snd')
+  
+  from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) =
+    distr (distr M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). (x, z))"
+    by (subst distr_distr[OF measurable_cut]) (auto dest: distributed_measurable simp: comp_def)
+
+  have "mutual_information b S P X Z =
+    (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^isub>M P))"
+    by (rule mutual_information_distr[OF S P Px Pz Pxz])
+  also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+    using b_gt_1 Pxz Px Pz
+    by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\<lambda>(x, y, z). (x, z)"])
+       (auto simp: split_beta' intro!: measurable_Pair measurable_snd' measurable_snd'' measurable_fst'' borel_measurable_times
+             dest!: distributed_real_measurable)
+  finally have mi_eq:
+    "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" .
+  
+  have ae1: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
+  moreover have ae2: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) (auto intro: measurable_snd')
+  moreover have ae3: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto intro: measurable_Pair measurable_snd')
+  moreover have ae4: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto intro: measurable_Pair)
+  moreover have ae5: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Px (fst x)"
+    using Px by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
+  moreover have ae6: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pyz (snd x)"
+    using Pyz by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
+  moreover have ae7: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd (snd x))"
+    using Pz Pz[THEN distributed_real_measurable] by (auto intro!: measurable_snd'' TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
+  moreover have ae8: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
+    using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
+    using measurable_comp[OF measurable_Pair[OF measurable_fst measurable_comp[OF measurable_snd measurable_snd]] Pxz[THEN distributed_real_measurable], of T]
+    using measurable_comp[OF measurable_snd measurable_Pair2[OF Pxz[THEN distributed_real_measurable]], of _ T]
+    by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure simp: comp_def)
+  moreover note ae9 = Pxyz[THEN distributed_real_AE]
+  ultimately have ae: "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P.
+    Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
+    Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
+    Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
+  proof eventually_elim
+    case (goal1 x)
+    show ?case
+    proof cases
+      assume "Pxyz x \<noteq> 0"
+      with goal1 have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" "0 < Pyz (snd x)" "0 < Pxyz x"
+        by auto
+      then show ?thesis
+        using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps)
+    qed simp
+  qed
+
+  have "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)
+    (\<lambda>x. Pxyz x * log b (Pxyz x) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pyz (snd x)))"
+    using finite_entropy_integrable[OF Fxyz]
+    using finite_entropy_integrable_transform[OF Fx Pxyz, of fst]
+    using finite_entropy_integrable_transform[OF Fyz Pxyz, of snd]
+    by simp
+  moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z)))) \<in> borel_measurable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)"
+    using Pxyz Px Pyz
+    by (auto intro!: borel_measurable_times measurable_fst'' measurable_snd'' dest!: distributed_real_measurable simp: split_beta')
+  ultimately have I1: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
+    apply (rule integrable_cong_AE_imp)
+    using ae1 ae4 ae5 ae6 ae9
+    by eventually_elim
+       (auto simp: log_divide_eq log_mult_eq mult_nonneg_nonneg field_simps zero_less_mult_iff)
+
+  have "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)
+    (\<lambda>x. Pxyz x * log b (Pxz (fst x, snd (snd x))) - Pxyz x * log b (Px (fst x)) - Pxyz x * log b (Pz (snd (snd x))))"
+    using finite_entropy_integrable_transform[OF Fxz Pxyz, of "\<lambda>x. (fst x, snd (snd x))"]
+    using finite_entropy_integrable_transform[OF Fx Pxyz, of fst]
+    using finite_entropy_integrable_transform[OF Fz Pxyz, of "snd \<circ> snd"]
+    by (simp add: measurable_Pair measurable_snd'' comp_def)
+  moreover have "(\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z))) \<in> borel_measurable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)"
+    using Pxyz Px Pz
+    by (auto intro!: measurable_compose[OF _ distributed_real_measurable[OF Pxz]]
+                     measurable_Pair borel_measurable_times measurable_fst'' measurable_snd''
+             dest!: distributed_real_measurable simp: split_beta')
+  ultimately have I2: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
+    apply (rule integrable_cong_AE_imp)
+    using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 ae9
+    by eventually_elim
+       (auto simp: log_divide_eq log_mult_eq mult_nonneg_nonneg field_simps zero_less_mult_iff)
+
+  from ae I1 I2 show ?eq
+    unfolding conditional_mutual_information_def
+    apply (subst mi_eq)
+    apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz])
+    apply (subst integral_diff(2)[symmetric])
+    apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
+    done
+
+  let ?P = "density (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) Pxyz"
+  interpret P: prob_space ?P
+    unfolding distributed_distr_eq_density[OF Pxyz, symmetric]
+    using distributed_measurable[OF Pxyz] by (rule prob_space_distr)
+
+  let ?Q = "density (T \<Otimes>\<^isub>M P) Pyz"
+  interpret Q: prob_space ?Q
+    unfolding distributed_distr_eq_density[OF Pyz, symmetric]
+    using distributed_measurable[OF Pyz] by (rule prob_space_distr)
+
+  let ?f = "\<lambda>(x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) / Pxyz (x, y, z)"
+
+  from subdensity_real[of snd, OF _ Pyz Pz]
+  have aeX1: "AE x in T \<Otimes>\<^isub>M P. Pz (snd x) = 0 \<longrightarrow> Pyz x = 0" by (auto simp: comp_def)
+  have aeX2: "AE x in T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd x)"
+    using Pz by (intro TP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
+
+  have aeX3: "AE y in T \<Otimes>\<^isub>M P. (\<integral>\<^isup>+ x. ereal (Pxz (x, snd y)) \<partial>S) = ereal (Pz (snd y))"
+    using Pz distributed_marginal_eq_joint2[OF P S Pz Pxz]
+    apply (intro TP.AE_pair_measure)
+    apply (auto simp: comp_def measurable_split_conv
+                intro!: measurable_snd'' borel_measurable_ereal_eq borel_measurable_ereal
+                        SP.borel_measurable_positive_integral_snd measurable_compose[OF _ Pxz[THEN distributed_real_measurable]]
+                        measurable_Pair
+                dest: distributed_real_AE distributed_real_measurable)
+    done
+
+  note M = borel_measurable_divide borel_measurable_diff borel_measurable_times borel_measurable_ereal
+           measurable_compose[OF _ measurable_snd]
+           measurable_Pair
+           measurable_compose[OF _ Pxyz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Pxz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Pyz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Pz[THEN distributed_real_measurable]]
+           measurable_compose[OF _ Px[THEN distributed_real_measurable]]
+           STP.borel_measurable_positive_integral_snd
+  have "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> (\<integral>\<^isup>+ (x, y, z). Pxz (x, z) * (Pyz (y, z) / Pz z) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+    apply (subst positive_integral_density)
+    apply (rule distributed_borel_measurable[OF Pxyz])
+    apply (rule distributed_AE[OF Pxyz])
+    apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) []
+    apply (rule positive_integral_mono_AE)
+    using ae5 ae6 ae7 ae8
+    apply eventually_elim
+    apply (auto intro!: divide_nonneg_nonneg mult_nonneg_nonneg)
+    done
+  also have "\<dots> = (\<integral>\<^isup>+(y, z). \<integral>\<^isup>+ x. ereal (Pxz (x, z)) * ereal (Pyz (y, z) / Pz z) \<partial>S \<partial>T \<Otimes>\<^isub>M P)"
+    by (subst STP.positive_integral_snd_measurable[symmetric])
+       (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M)
+  also have "\<dots> = (\<integral>\<^isup>+x. ereal (Pyz x) * 1 \<partial>T \<Otimes>\<^isub>M P)"
+    apply (rule positive_integral_cong_AE)
+    using aeX1 aeX2 aeX3 distributed_AE[OF Pyz] AE_space
+    apply eventually_elim
+  proof (case_tac x, simp del: times_ereal.simps add: space_pair_measure)
+    fix a b assume "Pz b = 0 \<longrightarrow> Pyz (a, b) = 0" "0 \<le> Pz b" "a \<in> space T \<and> b \<in> space P"
+      "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) \<partial>S) = ereal (Pz b)" "0 \<le> Pyz (a, b)" 
+    then show "(\<integral>\<^isup>+ x. ereal (Pxz (x, b)) * ereal (Pyz (a, b) / Pz b) \<partial>S) = ereal (Pyz (a, b))"
+      apply (subst positive_integral_multc)
+      apply (auto intro!: borel_measurable_ereal divide_nonneg_nonneg
+                          measurable_compose[OF _ Pxz[THEN distributed_real_measurable]] measurable_Pair
+                  split: prod.split)
+      done
+  qed
+  also have "\<dots> = 1"
+    using Q.emeasure_space_1 distributed_AE[OF Pyz] distributed_distr_eq_density[OF Pyz]
+    by (subst positive_integral_density[symmetric]) (auto intro!: M)
+  finally have le1: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<le> 1" .
+  also have "\<dots> < \<infinity>" by simp
+  finally have fin: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> \<infinity>" by simp
+
+  have pos: "(\<integral>\<^isup>+ x. ?f x \<partial>?P) \<noteq> 0"
+    apply (subst positive_integral_density)
+    apply (rule distributed_borel_measurable[OF Pxyz])
+    apply (rule distributed_AE[OF Pxyz])
+    apply (auto simp add: borel_measurable_ereal_iff split_beta' intro!: M) []
+    apply (simp add: split_beta')
+  proof
+    let ?g = "\<lambda>x. ereal (if Pxyz x = 0 then 0 else Pxz (fst x, snd (snd x)) * Pyz (snd x) / Pz (snd (snd x)))"
+    assume "(\<integral>\<^isup>+ x. ?g x \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P)) = 0"
+    then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. ?g x \<le> 0"
+      by (intro positive_integral_0_iff_AE[THEN iffD1]) (auto intro!: M borel_measurable_ereal measurable_If)
+    then have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxyz x = 0"
+      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      by eventually_elim (auto split: split_if_asm simp: mult_le_0_iff divide_le_0_iff)
+    then have "(\<integral>\<^isup>+ x. ereal (Pxyz x) \<partial>S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) = 0"
+      by (subst positive_integral_cong_AE[of _ "\<lambda>x. 0"]) auto
+    with P.emeasure_space_1 show False
+      by (subst (asm) emeasure_density) (auto intro!: M cong: positive_integral_cong)
+  qed
+
+  have neg: "(\<integral>\<^isup>+ x. - ?f x \<partial>?P) = 0"
+    apply (rule positive_integral_0_iff_AE[THEN iffD2])
+    apply (auto intro!: M simp: split_beta') []
+    apply (subst AE_density)
+    apply (auto intro!: M simp: split_beta') []
+    using ae5 ae6 ae7 ae8
+    apply eventually_elim
+    apply (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg)
+    done
+
+  have I3: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
+    apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ integral_diff(1)[OF I1 I2]])
+    using ae
+    apply (auto intro!: M simp: split_beta')
+    done
+
+  have "- log b 1 \<le> - log b (integral\<^isup>L ?P ?f)"
+  proof (intro le_imp_neg_le log_le[OF b_gt_1])
+    show "0 < integral\<^isup>L ?P ?f"
+      using neg pos fin positive_integral_positive[of ?P ?f]
+      by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def less_le split_beta')
+    show "integral\<^isup>L ?P ?f \<le> 1"
+      using neg le1 fin positive_integral_positive[of ?P ?f]
+      by (cases "(\<integral>\<^isup>+ x. ?f x \<partial>?P)") (auto simp add: lebesgue_integral_def split_beta' one_ereal_def)
+  qed
+  also have "- log b (integral\<^isup>L ?P ?f) \<le> (\<integral> x. - log b (?f x) \<partial>?P)"
+  proof (rule P.jensens_inequality[where a=0 and b=1 and I="{0<..}"])
+    show "AE x in ?P. ?f x \<in> {0<..}"
+      unfolding AE_density[OF distributed_borel_measurable[OF Pxyz]]
+      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      by eventually_elim (auto simp: divide_pos_pos mult_pos_pos)
+    show "integrable ?P ?f"
+      unfolding integrable_def 
+      using fin neg by (auto intro!: M simp: split_beta')
+    show "integrable ?P (\<lambda>x. - log b (?f x))"
+      apply (subst integral_density)
+      apply (auto intro!: M) []
+      apply (auto intro!: M distributed_real_AE[OF Pxyz]) []
+      apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+      apply (rule integrable_cong_AE[THEN iffD1, OF _ _ _ I3])
+      apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+      apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+      using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+      apply eventually_elim
+      apply (auto simp: log_divide_eq log_mult_eq zero_le_mult_iff zero_less_mult_iff zero_less_divide_iff field_simps)
+      done
+  qed (auto simp: b_gt_1 minus_log_convex)
+  also have "\<dots> = conditional_mutual_information b S T P X Y Z"
+    unfolding `?eq`
+    apply (subst integral_density)
+    apply (auto intro!: M) []
+    apply (auto intro!: M distributed_real_AE[OF Pxyz]) []
+    apply (auto intro!: M borel_measurable_uminus borel_measurable_log simp: split_beta') []
+    apply (intro integral_cong_AE)
+    using ae1 ae2 ae3 ae4 ae5 ae6 ae7 ae8 Pxyz[THEN distributed_real_AE]
+    apply eventually_elim
+    apply (auto simp: log_divide_eq zero_less_mult_iff zero_less_divide_iff field_simps)
+    done
+  finally show ?nonneg
+    by simp
 qed
 
 lemma (in information_space) conditional_mutual_information_eq:
@@ -909,92 +1614,28 @@
   assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z"
   shows "0 \<le> \<I>(X ; Y | Z)"
 proof -
-  def Pz \<equiv> "\<lambda>x. if x \<in> Z`space M then measure M (Z -` {x} \<inter> space M) else 0"
-  def Pxz \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, Z x))`space M then measure M ((\<lambda>x. (X x, Z x)) -` {x} \<inter> space M) else 0"
-  def Pyz \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (Y x, Z x))`space M then measure M ((\<lambda>x. (Y x, Z x)) -` {x} \<inter> space M) else 0"
-  def Pxyz \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x))`space M then measure M ((\<lambda>x. (X x, Y x, Z x)) -` {x} \<inter> space M) else 0"
-  let ?M = "X`space M \<times> Y`space M \<times> Z`space M"
-
-  note XZ = simple_function_Pair[OF X Z]
-  note YZ = simple_function_Pair[OF Y Z]
-  note XYZ = simple_function_Pair[OF X simple_function_Pair[OF Y Z]]
-  have Pz: "simple_distributed M Z Pz"
-    using Z by (rule simple_distributedI) (auto simp: Pz_def)
-  have Pxz: "simple_distributed M (\<lambda>x. (X x, Z x)) Pxz"
-    using XZ by (rule simple_distributedI) (auto simp: Pxz_def)
-  have Pyz: "simple_distributed M (\<lambda>x. (Y x, Z x)) Pyz"
-    using YZ by (rule simple_distributedI) (auto simp: Pyz_def)
-  have Pxyz: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Pxyz"
-    using XYZ by (rule simple_distributedI) (auto simp: Pxyz_def)
-
-  { fix z assume z: "z \<in> Z ` space M" then have "(\<Sum>x\<in>X ` space M. Pxz (x, z)) = Pz z"
-      using distributed_marginal_eq_joint_simple[OF X Pz Pxz z]
-      by (auto intro!: setsum_cong simp: Pxz_def) }
-  note marginal1 = this
-
-  { fix z assume z: "z \<in> Z ` space M" then have "(\<Sum>y\<in>Y ` space M. Pyz (y, z)) = Pz z"
-      using distributed_marginal_eq_joint_simple[OF Y Pz Pyz z]
-      by (auto intro!: setsum_cong simp: Pyz_def) }
-  note marginal2 = this
-
-  have "- \<I>(X ; Y | Z) = - (\<Sum>(x, y, z) \<in> ?M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
-    unfolding conditional_mutual_information_eq[OF Pz Pyz Pxz Pxyz]
-    using X Y Z by (auto intro!: setsum_mono_zero_left simp: Pxyz_def simple_functionD)
-  also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. Pxz (x, z) * (Pyz (y,z) / Pz z))"
-    unfolding split_beta'
-  proof (rule log_setsum_divide)
-    show "?M \<noteq> {}" using not_empty by simp
-    show "1 < b" using b_gt_1 .
-
-    show "finite ?M" using X Y Z by (auto simp: simple_functionD)
-
-    then show "(\<Sum>x\<in>?M. Pxyz (fst x, fst (snd x), snd (snd x))) = 1"
-      apply (subst Pxyz[THEN simple_distributed_setsum_space, symmetric])
-      apply simp
-      apply (intro setsum_mono_zero_right)
-      apply (auto simp: Pxyz_def)
-      done
-    let ?N = "(\<lambda>x. (X x, Y x, Z x)) ` space M"
-    fix x assume x: "x \<in> ?M"
-    let ?Q = "Pxyz (fst x, fst (snd x), snd (snd x))"
-    let ?P = "Pxz (fst x, snd (snd x)) * (Pyz (fst (snd x), snd (snd x)) / Pz (snd (snd x)))"
-    from x show "0 \<le> ?Q" "0 \<le> ?P"
-      using Pxyz[THEN simple_distributed, THEN distributed_real_AE]
-      using Pxz[THEN simple_distributed, THEN distributed_real_AE]
-      using Pyz[THEN simple_distributed, THEN distributed_real_AE]
-      using Pz[THEN simple_distributed, THEN distributed_real_AE]
-      by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg simp: AE_count_space Pxyz_def Pxz_def Pyz_def Pz_def)
-    moreover assume "0 < ?Q"
-    moreover have "AE x in count_space ?N. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-      by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz[THEN simple_distributed] Pz[THEN simple_distributed]]) (auto intro: measurable_snd')
-    then have "\<And>x. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-      by (auto simp: Pz_def Pxyz_def AE_count_space)
-    moreover have "AE x in count_space ?N. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-      by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz[THEN simple_distributed] Pxz[THEN simple_distributed]]) (auto intro: measurable_Pair measurable_snd')
-    then have "\<And>x. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
-      by (auto simp: Pz_def Pxyz_def AE_count_space)
-    moreover have "AE x in count_space ?N. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
-      by (intro subdensity_real[of snd, OF _ Pxyz[THEN simple_distributed] Pyz[THEN simple_distributed]]) (auto intro: measurable_Pair)
-    then have "\<And>x. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
-      by (auto simp: Pz_def Pxyz_def AE_count_space)
-    ultimately show "0 < ?P" using x by (auto intro!: divide_pos_pos mult_pos_pos simp: less_le)
-  qed
-  also have "(\<Sum>(x, y, z) \<in> ?M. Pxz (x, z) * (Pyz (y,z) / Pz z)) = (\<Sum>z\<in>Z`space M. Pz z)"
-    apply (simp add: setsum_cartesian_product')
-    apply (subst setsum_commute)
-    apply (subst (2) setsum_commute)
-    apply (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] marginal1 marginal2
-          intro!: setsum_cong)
-    done
-  also have "log b (\<Sum>z\<in>Z`space M. Pz z) = 0"
-    using Pz[THEN simple_distributed_setsum_space] by simp
-  finally show ?thesis by simp
+  have [simp]: "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) \<Otimes>\<^isub>M count_space (Z ` space M) =
+      count_space (X`space M \<times> Y`space M \<times> Z`space M)"
+    by (simp add: pair_measure_count_space X Y Z simple_functionD)
+  note sf = sigma_finite_measure_count_space_finite[OF simple_functionD(1)]
+  note sd = simple_distributedI[OF _ refl]
+  note sp = simple_function_Pair
+  show ?thesis
+   apply (rule conditional_mutual_information_generic_nonneg[OF sf[OF X] sf[OF Y] sf[OF Z]])
+   apply (rule simple_distributed[OF sd[OF X]])
+   apply (rule simple_distributed[OF sd[OF Z]])
+   apply (rule simple_distributed_joint[OF sd[OF sp[OF Y Z]]])
+   apply (rule simple_distributed_joint[OF sd[OF sp[OF X Z]]])
+   apply (rule simple_distributed_joint2[OF sd[OF sp[OF X sp[OF Y Z]]]])
+   apply (auto intro!: integrable_count_space simp: X Y Z simple_functionD)
+   done
 qed
 
 subsection {* Conditional Entropy *}
 
 definition (in prob_space)
-  "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y"
+  "conditional_entropy b S T X Y = - (\<integral>(x, y). log b (real (RN_deriv (S \<Otimes>\<^isub>M T) (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) (x, y)) / 
+    real (RN_deriv T (distr M T Y) y)) \<partial>distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)))"
 
 abbreviation (in information_space)
   conditional_entropy_Pow ("\<H>'(_ | _')") where
@@ -1003,66 +1644,112 @@
 lemma (in information_space) conditional_entropy_generic_eq:
   fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
   assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
-  assumes Px: "distributed M S X Px"
   assumes Py: "distributed M T Y Py"
   assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
-  assumes I1: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
-  assumes I2: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
   shows "conditional_entropy b S T X Y = - (\<integral>(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \<partial>(S \<Otimes>\<^isub>M T))"
 proof -
   interpret S: sigma_finite_measure S by fact
   interpret T: sigma_finite_measure T by fact
   interpret ST: pair_sigma_finite S T ..
-  have ST: "sigma_finite_measure (S \<Otimes>\<^isub>M T)" ..
 
-  interpret Pxy: prob_space "density (S \<Otimes>\<^isub>M T) Pxy"
-    unfolding Pxy[THEN distributed_distr_eq_density, symmetric]
-    using Pxy[THEN distributed_measurable] by (rule prob_space_distr)
+  have "AE x in density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x)). Pxy x = real (RN_deriv (S \<Otimes>\<^isub>M T) (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) x)"
+    unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
+    unfolding distributed_distr_eq_density[OF Pxy]
+    using distributed_RN_deriv[OF Pxy]
+    by auto
+  moreover
+  have "AE x in density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x)). Py (snd x) = real (RN_deriv T (distr M T Y) (snd x))"
+    unfolding AE_density[OF distributed_borel_measurable, OF Pxy]
+    unfolding distributed_distr_eq_density[OF Py]
+    apply (rule ST.AE_pair_measure)
+    apply (auto intro!: sets_Collect borel_measurable_eq measurable_compose[OF _ distributed_real_measurable[OF Py]]
+                        distributed_real_measurable[OF Pxy] distributed_real_AE[OF Py]
+                        borel_measurable_real_of_ereal measurable_compose[OF _ borel_measurable_RN_deriv_density])
+    using distributed_RN_deriv[OF Py]
+    apply auto
+    done    
+  ultimately
+  have "conditional_entropy b S T X Y = - (\<integral>x. Pxy x * log b (Pxy x / Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+    unfolding conditional_entropy_def neg_equal_iff_equal
+    apply (subst integral_density(1)[symmetric])
+    apply (auto simp: distributed_real_measurable[OF Pxy] distributed_real_AE[OF Pxy]
+                      measurable_compose[OF _ distributed_real_measurable[OF Py]]
+                      distributed_distr_eq_density[OF Pxy]
+                intro!: integral_cong_AE)
+    done
+  then show ?thesis by (simp add: split_beta')
+qed
 
-  from Py Pxy have distr_eq: "distr M T Y =
-    distr (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) T snd"
-    by (subst distr_distr[OF measurable_snd]) (auto dest: distributed_measurable simp: comp_def)
+lemma (in information_space) conditional_entropy_eq_entropy:
+  fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes I1: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+  assumes I2: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+  shows "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y"
+proof -
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T ..
 
   have "entropy b T Y = - (\<integral>y. Py y * log b (Py y) \<partial>T)"
-    by (rule entropy_distr[OF T Py])
+    by (rule entropy_distr[OF Py])
   also have "\<dots> = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^isub>M T))"
     using b_gt_1 Py[THEN distributed_real_measurable]
     by (subst distributed_transform_integral[OF Pxy Py, where T=snd]) (auto intro!: integral_cong)
   finally have e_eq: "entropy b T Y = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^isub>M T))" .
-  
-  have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
-    by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair)
-  moreover have "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+
+  have ae2: "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
     by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
-  moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Px (fst x)"
-    using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
-  moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Py (snd x)"
+  moreover have ae4: "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Py (snd x)"
     using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
-  moreover note Pxy[THEN distributed_real_AE]
-  ultimately have pos: "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Pxy x \<and> 0 \<le> Px (fst x) \<and> 0 \<le> Py (snd x) \<and>
-    (Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Px (fst x) \<and> 0 < Py (snd x)))"
+  moreover note ae5 = Pxy[THEN distributed_real_AE]
+  ultimately have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Pxy x \<and> 0 \<le> Py (snd x) \<and>
+    (Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Py (snd x)))"
     by eventually_elim auto
-
-  from pos have "AE x in S \<Otimes>\<^isub>M T.
+  then have ae: "AE x in S \<Otimes>\<^isub>M T.
      Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))"
     by eventually_elim (auto simp: log_simps mult_pos_pos field_simps b_gt_1)
-  with I1 I2 show ?thesis
-    unfolding conditional_entropy_def
-    apply (subst e_eq)
-    apply (subst entropy_distr[OF ST Pxy])
-    unfolding minus_diff_minus
-    apply (subst integral_diff(2)[symmetric])
-    apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
+  have "conditional_entropy b S T X Y = 
+    - (\<integral>x. Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+    unfolding conditional_entropy_generic_eq[OF S T Py Pxy] neg_equal_iff_equal
+    apply (intro integral_cong_AE)
+    using ae
+    apply eventually_elim
+    apply auto
     done
+  also have "\<dots> = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^isub>M T)) - - (\<integral>x.  Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+    by (simp add: integral_diff[OF I1 I2])
+  finally show ?thesis 
+    unfolding conditional_entropy_generic_eq[OF S T Py Pxy] entropy_distr[OF Pxy] e_eq
+    by (simp add: split_beta')
+qed
+
+lemma (in information_space) conditional_entropy_eq_entropy_simple:
+  assumes X: "simple_function M X" and Y: "simple_function M Y"
+  shows "\<H>(X | Y) = entropy b (count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x)) - \<H>(Y)"
+proof -
+  have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
+    (is "?P = ?C") using X Y by (simp add: simple_functionD pair_measure_count_space)
+  show ?thesis
+    by (rule conditional_entropy_eq_entropy sigma_finite_measure_count_space_finite
+                 simple_functionD  X Y simple_distributed simple_distributedI[OF _ refl]
+                 simple_distributed_joint simple_function_Pair integrable_count_space)+
+       (auto simp: `?P = ?C` intro!: integrable_count_space simple_functionD  X Y)
 qed
 
 lemma (in information_space) conditional_entropy_eq:
-  assumes Y: "simple_distributed M Y Py" and X: "simple_function M X"
+  assumes Y: "simple_distributed M Y Py"
   assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
     shows "\<H>(X | Y) = - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
 proof (subst conditional_entropy_generic_eq[OF _ _
-  simple_distributed[OF simple_distributedI[OF X refl]] simple_distributed[OF Y] simple_distributed_joint[OF XY]])
-  have [simp]: "finite (X`space M)" using X by (simp add: simple_function_def)
+  simple_distributed[OF Y] simple_distributed_joint[OF XY]])
+  have "finite ((\<lambda>x. (X x, Y x))`space M)"
+    using XY unfolding simple_distributed_def by auto
+  from finite_imageI[OF this, of fst]
+  have [simp]: "finite (X`space M)"
+    by (simp add: image_compose[symmetric] comp_def)
   note Y[THEN simple_distributed_finite, simp]
   show "sigma_finite_measure (count_space (X ` space M))"
     by (simp add: sigma_finite_measure_count_space_finite)
@@ -1071,15 +1758,11 @@
   let ?f = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x else 0)"
   have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
     (is "?P = ?C")
-    using X Y by (simp add: simple_distributed_finite pair_measure_count_space)
-  with X Y show
-      "integrable ?P (\<lambda>x. ?f x * log b (?f x))"
-      "integrable ?P (\<lambda>x. ?f x * log b (Py (snd x)))"
-    by (auto intro!: integrable_count_space simp: simple_distributed_finite)
+    using Y by (simp add: simple_distributed_finite pair_measure_count_space)
   have eq: "(\<lambda>(x, y). ?f (x, y) * log b (?f (x, y) / Py y)) =
     (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / Py (snd x)) else 0)"
     by auto
-  from X Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
+  from Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
     - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
     by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite eq setsum_cases split_beta')
 qed
@@ -1111,11 +1794,11 @@
     by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]]) (auto intro: measurable_Pair)
   then show ?thesis
     apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy])
-    apply (subst conditional_entropy_eq[OF Py X Pxy])
+    apply (subst conditional_entropy_eq[OF Py Pxy])
     apply (auto intro!: setsum_cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum_reindex[OF inj]
                 log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space)
     using Py[THEN simple_distributed, THEN distributed_real_AE] Pxy[THEN simple_distributed, THEN distributed_real_AE]
-    apply (auto simp add: not_le[symmetric] AE_count_space)
+  apply (auto simp add: not_le[symmetric] AE_count_space)
     done
 qed
 
@@ -1138,14 +1821,14 @@
 proof -
   have X: "entropy b S X = - (\<integral>x. Pxy x * log b (Px (fst x)) \<partial>(S \<Otimes>\<^isub>M T))"
     using b_gt_1 Px[THEN distributed_real_measurable]
-    apply (subst entropy_distr[OF S Px])
+    apply (subst entropy_distr[OF Px])
     apply (subst distributed_transform_integral[OF Pxy Px, where T=fst])
     apply (auto intro!: integral_cong)
     done
 
   have Y: "entropy b T Y = - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
     using b_gt_1 Py[THEN distributed_real_measurable]
-    apply (subst entropy_distr[OF T Py])
+    apply (subst entropy_distr[OF Py])
     apply (subst distributed_transform_integral[OF Pxy Py, where T=snd])
     apply (auto intro!: integral_cong)
     done
@@ -1156,7 +1839,7 @@
   have ST: "sigma_finite_measure (S \<Otimes>\<^isub>M T)" ..
 
   have XY: "entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^isub>M T))"
-    by (subst entropy_distr[OF ST Pxy]) (auto intro!: integral_cong)
+    by (subst entropy_distr[OF Pxy]) (auto intro!: integral_cong)
   
   have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
     by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair)
@@ -1197,6 +1880,20 @@
   finally show ?thesis ..
 qed
 
+lemma (in information_space) mutual_information_eq_entropy_conditional_entropy':
+  fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes Ix: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
+  assumes Iy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+  assumes Ixy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+  shows  "mutual_information b S T X Y = entropy b S X - conditional_entropy b S T X Y"
+  using
+    mutual_information_eq_entropy_conditional_entropy_distr[OF S T Px Py Pxy Ix Iy Ixy]
+    conditional_entropy_eq_entropy[OF S T Py Pxy Ixy Iy]
+  by simp
+
 lemma (in information_space) mutual_information_eq_entropy_conditional_entropy:
   assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y"
   shows  "\<I>(X ; Y) = \<H>(X) - \<H>(X | Y)"
@@ -1217,7 +1914,7 @@
     using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]
     by (rule mutual_information_eq_entropy_conditional_entropy_distr) (auto simp: eq integrable_count_space)
   then show ?thesis
-    unfolding conditional_entropy_def by simp
+    unfolding conditional_entropy_eq_entropy_simple[OF sf_X sf_Y] by simp
 qed
 
 lemma (in information_space) mutual_information_nonneg_simple:
@@ -1252,6 +1949,25 @@
   finally show ?thesis by auto
 qed
 
+lemma (in information_space) 
+  fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Px: "finite_entropy S X Px" and Py: "finite_entropy T Y Py"
+  assumes Pxy: "finite_entropy (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  shows "conditional_entropy b S T X Y \<le> entropy b S X"
+proof -
+
+  have "0 \<le> mutual_information b S T X Y" 
+    by (rule mutual_information_nonneg') fact+
+  also have "\<dots> = entropy b S X - conditional_entropy b S T X Y"
+    apply (rule mutual_information_eq_entropy_conditional_entropy')
+    using assms
+    by (auto intro!: finite_entropy_integrable finite_entropy_distributed
+      finite_entropy_integrable_transform[OF Px]
+      finite_entropy_integrable_transform[OF Py])
+  finally show ?thesis by auto
+qed
+
 lemma (in information_space) entropy_chain_rule:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
   shows  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
@@ -1269,12 +1985,12 @@
   also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` space M. ?h x * log b (?h x))"
     by (auto intro!: setsum_cong)
   also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^isub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
-    by (subst entropy_distr[OF _ simple_distributed_joint[OF YX]])
+    by (subst entropy_distr[OF simple_distributed_joint[OF YX]])
        (auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite
              cong del: setsum_cong  intro!: setsum_mono_zero_left)
   finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^isub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
   then show ?thesis
-    unfolding conditional_entropy_def by simp
+    unfolding conditional_entropy_eq_entropy_simple[OF Y X] by simp
 qed
 
 lemma (in information_space) entropy_partition:
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -355,6 +355,98 @@
     "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
   using borel_measurable_implies_simple_function_sequence[OF u] by auto
 
+lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]:
+  fixes u :: "'a \<Rightarrow> ereal"
+  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 2, case_names cong set mult add]:
+  fixes u :: "'a \<Rightarrow> ereal"
+  assumes u: "simple_function M u" and nn: "\<And>x. 0 \<le> u x"
+  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. 0 \<le> c \<Longrightarrow> simple_function M u \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)"
+  assumes add: "\<And>u v. simple_function M u \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<Longrightarrow> (\<And>x. 0 \<le> v x) \<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 nn have "finite (u ` space M)" "\<And>x. x \<in> u ` space M \<Longrightarrow> 0 \<le> x"
+      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 setsum_nonneg
+       simple_function_setsum)
+  qed fact
+qed
+
+lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]:
+  fixes u :: "'a \<Rightarrow> ereal"
+  assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
+  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. 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> 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> 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. \<infinity> \<notin> range (U i)" and
+    sup: "\<And>x. (SUP i. U i x) = max 0 (u x)" and nn: "\<And>i x. 0 \<le> U i x"
+  have u_eq: "u = (SUP i. U i)"
+    using nn u sup by (auto simp: max_def)
+  
+  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 `simple_function M (U i)` nn
+      by (induct rule: simple_function_induct_nn)
+         (auto intro: set mult add cong dest!: borel_measurable_simple_function)
+  qed fact+
+qed
+
 lemma simple_function_If_set:
   assumes sf: "simple_function M f" "simple_function M g" and A: "A \<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")
@@ -595,7 +687,7 @@
 
 lemma simple_integral_indicator:
   assumes "A \<in> sets M"
-  assumes "simple_function M f"
+  assumes f: "simple_function M f"
   shows "(\<integral>\<^isup>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 cases
@@ -670,32 +762,6 @@
   shows "integral\<^isup>S M f = (\<integral>\<^isup>Sx. f x * indicator S x \<partial>M)"
   using assms by (intro simple_integral_cong_AE) auto
 
-lemma simple_integral_distr:
-  assumes T: "T \<in> measurable M M'"
-    and f: "simple_function M' f"
-  shows "integral\<^isup>S (distr M M' T) f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
-  unfolding simple_integral_def
-proof (intro setsum_mono_zero_cong_right ballI)
-  show "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space (distr M M' T)"
-    using T unfolding measurable_def by auto
-  show "finite (f ` space (distr M M' T))"
-    using f unfolding simple_function_def by auto
-next
-  fix i assume "i \<in> f ` space (distr M M' T) - (\<lambda>x. f (T x)) ` space M"
-  then have "T -` (f -` {i} \<inter> space (distr M M' T)) \<inter> space M = {}" by (auto simp: image_iff)
-  with f[THEN simple_functionD(2), of "{i}"]
-  show "i * emeasure (distr M M' T) (f -` {i} \<inter> space (distr M M' T)) = 0"
-    using T by (simp add: emeasure_distr)
-next
-  fix i assume "i \<in> (\<lambda>x. f (T x)) ` space M"
-  then 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
-  with f[THEN simple_functionD(2), of "{i}"] T
-  show "i * emeasure (distr M M' T) (f -` {i} \<inter> space (distr M M' T)) =
-      i * (emeasure M) ((\<lambda>x. f (T x)) -` {i} \<inter> space M)"
-    by (auto simp add: emeasure_distr)
-qed
-
 lemma simple_integral_cmult_indicator:
   assumes A: "A \<in> sets M"
   shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * (emeasure M) A"
@@ -1088,7 +1154,7 @@
 qed
 
 lemma positive_integral_cmult:
-  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> c"
+  assumes f: "f \<in> borel_measurable M" "0 \<le> c"
   shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
 proof -
   have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
@@ -1101,14 +1167,14 @@
 qed
 
 lemma positive_integral_multc:
-  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> c"
+  assumes "f \<in> borel_measurable M" "0 \<le> c"
   shows "(\<integral>\<^isup>+ x. f x * c \<partial>M) = integral\<^isup>P M f * c"
   unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
 
 lemma positive_integral_indicator[simp]:
   "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x\<partial>M) = (emeasure M) A"
   by (subst positive_integral_eq_simple_integral)
-     (auto simp: simple_function_indicator simple_integral_indicator)
+     (auto simp: simple_integral_indicator)
 
 lemma positive_integral_cmult_indicator:
   "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A"
@@ -1151,7 +1217,7 @@
 qed simp
 
 lemma positive_integral_Markov_inequality:
-  assumes u: "u \<in> borel_measurable M" "AE x in M. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" "c \<noteq> \<infinity>"
+  assumes u: "u \<in> borel_measurable M" "AE x in M. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c"
   shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
     (is "(emeasure M) ?A \<le> _ * ?PI")
 proof -
@@ -1342,23 +1408,21 @@
   by (auto intro!: positive_integral_0_iff_AE[THEN iffD2])
 
 lemma positive_integral_subalgebra:
-  assumes f: "f \<in> borel_measurable N" "AE x in N. 0 \<le> f x"
+  assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x"
   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\<^isup>P N f = integral\<^isup>P M f"
 proof -
-  from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this
-  note sf = simple_function_subalgebra[OF fs(1) N(1,2)]
-  from positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric]
-  have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * emeasure N (fs i -` {x} \<inter> space M))"
-    unfolding fs(4) positive_integral_max_0
-    unfolding simple_integral_def `space N = space M` by simp
-  also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * (emeasure M) (fs i -` {x} \<inter> space M))"
-    using N simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto
-  also have "\<dots> = integral\<^isup>P M f"
-    using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric]
-    unfolding fs(4) positive_integral_max_0
-    unfolding simple_integral_def `space N = space M` by simp
-  finally show ?thesis .
+  have [simp]: "\<And>f :: 'a \<Rightarrow> ereal. 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)
+  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: positive_integral_add positive_integral_cmult positive_integral_monotone_convergence_SUP N)
+    apply (auto intro!: positive_integral_cong cong: positive_integral_cong simp: N(2)[symmetric])
+    done
 qed
 
 section "Lebesgue Integral"
@@ -1423,6 +1487,25 @@
   "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable M g"
   by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
 
+lemma integral_mono_AE:
+  assumes fg: "integrable M f" "integrable M g"
+  and mono: "AE t in M. f t \<le> g t"
+  shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
+proof -
+  have "AE x in M. ereal (f x) \<le> ereal (g x)"
+    using mono by auto
+  moreover have "AE x in M. ereal (- g x) \<le> ereal (- f x)"
+    using mono by auto
+  ultimately show ?thesis using fg
+    by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
+             simp: positive_integral_positive lebesgue_integral_def diff_minus)
+qed
+
+lemma integral_mono:
+  assumes "integrable M f" "integrable M g" "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
+  shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
+  using assms by (auto intro: integral_mono_AE)
+
 lemma positive_integral_eq_integral:
   assumes f: "integrable M f"
   assumes nonneg: "AE x in M. 0 \<le> f x" 
@@ -1571,20 +1654,16 @@
   assumes f: "f \<in> borel_measurable M" and "0 \<le> c"
   shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^isup>L M f"
 proof -
-  { have "c * real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (f x)))) =
-      real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (f x))))"
-      by simp
-    also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (f x))))"
+  { have "real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (f x)))) =
+      real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (f x))))"
       using f `0 \<le> c` by (subst positive_integral_cmult) auto
     also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (c * f x))))"
       using `0 \<le> c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def zero_le_mult_iff)
     finally have "real (integral\<^isup>P M (\<lambda>x. ereal (c * f x))) = c * real (integral\<^isup>P M (\<lambda>x. ereal (f x)))"
       by (simp add: positive_integral_max_0) }
   moreover
-  { have "c * real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (- f x)))) =
-      real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (- f x))))"
-      by simp
-    also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (- f x))))"
+  { have "real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (- f x)))) =
+      real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (- f x))))"
       using f `0 \<le> c` by (subst positive_integral_cmult) auto
     also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (- c * f x))))"
       using `0 \<le> c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def mult_le_0_iff)
@@ -1615,25 +1694,6 @@
   shows "(\<integral> x. f x * c \<partial>M) = integral\<^isup>L M f * c"
   unfolding mult_commute[of _ c] integral_cmult[OF assms] ..
 
-lemma integral_mono_AE:
-  assumes fg: "integrable M f" "integrable M g"
-  and mono: "AE t in M. f t \<le> g t"
-  shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
-proof -
-  have "AE x in M. ereal (f x) \<le> ereal (g x)"
-    using mono by auto
-  moreover have "AE x in M. ereal (- g x) \<le> ereal (- f x)"
-    using mono by auto
-  ultimately show ?thesis using fg
-    by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
-             simp: positive_integral_positive lebesgue_integral_def diff_minus)
-qed
-
-lemma integral_mono:
-  assumes "integrable M f" "integrable M g" "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
-  shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
-  using assms by (auto intro: integral_mono_AE)
-
 lemma integral_diff[simp, intro]:
   assumes f: "integrable M f" and g: "integrable M g"
   shows "integrable M (\<lambda>t. f t - g t)"
@@ -1685,8 +1745,33 @@
   thus "?int S" and "?I S" by auto
 qed
 
+lemma integrable_bound:
+  assumes "integrable M f" and f: "AE x in M. \<bar>g x\<bar> \<le> f x"
+  assumes borel: "g \<in> borel_measurable M"
+  shows "integrable M g"
+proof -
+  have "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
+    by (auto intro!: positive_integral_mono)
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
+    using f by (auto intro!: positive_integral_mono_AE)
+  also have "\<dots> < \<infinity>"
+    using `integrable M f` unfolding integrable_def by auto
+  finally have pos: "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
+
+  have "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
+    by (auto intro!: positive_integral_mono)
+  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
+    using f by (auto intro!: positive_integral_mono_AE)
+  also have "\<dots> < \<infinity>"
+    using `integrable M f` unfolding integrable_def by auto
+  finally have neg: "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
+
+  from neg pos borel show ?thesis
+    unfolding integrable_def by auto
+qed
+
 lemma integrable_abs:
-  assumes "integrable M f"
+  assumes f: "integrable M f"
   shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
 proof -
   from assms have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>)\<partial>M) = 0"
@@ -1711,33 +1796,6 @@
     by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0)
 qed
 
-lemma integrable_bound:
-  assumes "integrable M f"
-  and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
-    "\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
-  assumes borel: "g \<in> borel_measurable M"
-  shows "integrable M g"
-proof -
-  have "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal \<bar>g x\<bar> \<partial>M)"
-    by (auto intro!: positive_integral_mono)
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
-    using f by (auto intro!: positive_integral_mono)
-  also have "\<dots> < \<infinity>"
-    using `integrable M f` unfolding integrable_def by auto
-  finally have pos: "(\<integral>\<^isup>+ x. ereal (g x) \<partial>M) < \<infinity>" .
-
-  have "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) \<le> (\<integral>\<^isup>+ x. ereal (\<bar>g x\<bar>) \<partial>M)"
-    by (auto intro!: positive_integral_mono)
-  also have "\<dots> \<le> (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
-    using f by (auto intro!: positive_integral_mono)
-  also have "\<dots> < \<infinity>"
-    using `integrable M f` unfolding integrable_def by auto
-  finally have neg: "(\<integral>\<^isup>+ x. ereal (- g x) \<partial>M) < \<infinity>" .
-
-  from neg pos borel show ?thesis
-    unfolding integrable_def by auto
-qed
-
 lemma lebesgue_integral_nonneg:
   assumes ae: "(AE x in M. 0 \<le> f x)" shows "0 \<le> integral\<^isup>L M f"
 proof -
@@ -1760,11 +1818,7 @@
     using int by (simp add: integrable_abs)
   show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M"
     using int unfolding integrable_def by auto
-next
-  fix x assume "x \<in> space M"
-  show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>max (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>"
-    by auto
-qed
+qed auto
 
 lemma integrable_min:
   assumes int: "integrable M f" "integrable M g"
@@ -1774,11 +1828,7 @@
     using int by (simp add: integrable_abs)
   show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M"
     using int unfolding integrable_def by auto
-next
-  fix x assume "x \<in> space M"
-  show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>min (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>"
-    by auto
-qed
+qed auto
 
 lemma integral_triangle_inequality:
   assumes "integrable M f"
@@ -1802,74 +1852,71 @@
 qed
 
 lemma integral_monotone_convergence_pos:
-  assumes i: "\<And>i. integrable M (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
-  and pos: "\<And>x i. 0 \<le> f i x"
-  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
-  and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+  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) ----> u x"
+    and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+    and u: "u \<in> borel_measurable M"
   shows "integrable M u"
   and "integral\<^isup>L M u = x"
 proof -
-  { fix x have "0 \<le> u x"
-      using mono pos[of 0 x] incseq_le[OF _ lim, of x 0]
-      by (simp add: mono_def incseq_def) }
-  note pos_u = this
-
-  have SUP_F: "\<And>x. (SUP n. ereal (f n x)) = ereal (u x)"
-    unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim)
-
-  have borel_f: "\<And>i. (\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
-    using i unfolding integrable_def by auto
-  hence "(\<lambda>x. SUP i. ereal (f i x)) \<in> borel_measurable M"
-    by auto
-  hence borel_u: "u \<in> borel_measurable M"
-    by (auto simp: borel_measurable_ereal_iff SUP_F)
-
-  hence [simp]: "\<And>i. (\<integral>\<^isup>+x. ereal (- f i x) \<partial>M) = 0" "(\<integral>\<^isup>+x. ereal (- u x) \<partial>M) = 0"
-    using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
-
-  have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M) = ereal (integral\<^isup>L M (f n))"
-    using i positive_integral_positive[of M] by (auto simp: ereal_real lebesgue_integral_def integrable_def)
-
-  have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
-    using pos i by (auto simp: integral_positive)
-  hence "0 \<le> x"
-    using LIMSEQ_le_const[OF ilim, of 0] by auto
-
-  from mono pos i have pI: "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M))"
-    by (auto intro!: positive_integral_monotone_convergence_SUP
-      simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric])
-  also have "\<dots> = ereal x" unfolding integral_eq
-  proof (rule SUP_eq_LIMSEQ[THEN iffD2])
-    show "mono (\<lambda>n. integral\<^isup>L M (f n))"
-      using mono i by (auto simp: mono_def intro!: integral_mono)
-    show "(\<lambda>n. integral\<^isup>L M (f n)) ----> x" using ilim .
+  have "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = (SUP n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M))"
+  proof (subst positive_integral_monotone_convergence_SUP_AE[symmetric])
+    fix i
+    from mono pos show "AE x in M. ereal (f i x) \<le> ereal (f (Suc i) x) \<and> 0 \<le> ereal (f i x)"
+      by eventually_elim (auto simp: mono_def)
+    show "(\<lambda>x. ereal (f i x)) \<in> borel_measurable M"
+      using i by (auto simp: integrable_def)
+  next
+    show "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = \<integral>\<^isup>+ x. (SUP i. ereal (f i x)) \<partial>M"
+      apply (rule positive_integral_cong_AE)
+      using lim mono
+      by eventually_elim (simp add: SUP_eq_LIMSEQ[THEN iffD2])
   qed
-  finally show  "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \<le> x`
-    unfolding integrable_def lebesgue_integral_def by auto
+  also have "\<dots> = ereal x"
+    using mono i unfolding positive_integral_eq_integral[OF i pos]
+    by (subst SUP_eq_LIMSEQ) (auto simp: mono_def intro!: integral_mono_AE ilim)
+  finally have "(\<integral>\<^isup>+ x. ereal (u x) \<partial>M) = ereal x" .
+  moreover have "(\<integral>\<^isup>+ x. ereal (- u x) \<partial>M) = 0"
+  proof (subst positive_integral_0_iff_AE)
+    show "(\<lambda>x. ereal (- u x)) \<in> borel_measurable M"
+      using u by auto
+    from mono pos[of 0] lim show "AE x in M. ereal (- u x) \<le> 0"
+    proof eventually_elim
+      fix x assume "mono (\<lambda>n. f n x)" "0 \<le> f 0 x" "(\<lambda>i. f i x) ----> u x"
+      then show "ereal (- u x) \<le> 0"
+        using incseq_le[of "\<lambda>n. f n x" "u x" 0] by (simp add: mono_def incseq_def)
+    qed
+  qed
+  ultimately show "integrable M u" "integral\<^isup>L M u = x"
+    by (auto simp: integrable_def lebesgue_integral_def u)
 qed
 
 lemma integral_monotone_convergence:
-  assumes f: "\<And>i. integrable M (f i)" and "mono f"
-  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
+  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) ----> u x"
   and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
+  and u: "u \<in> borel_measurable M"
   shows "integrable M u"
   and "integral\<^isup>L M u = x"
 proof -
   have 1: "\<And>i. integrable M (\<lambda>x. f i x - f 0 x)"
-      using f by (auto intro!: integral_diff)
-  have 2: "\<And>x. mono (\<lambda>n. f n x - f 0 x)" using `mono f`
-      unfolding mono_def le_fun_def by auto
-  have 3: "\<And>x n. 0 \<le> f n x - f 0 x" using `mono f`
-      unfolding mono_def le_fun_def by (auto simp: field_simps)
-  have 4: "\<And>x. (\<lambda>i. f i x - f 0 x) ----> u 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) ----> 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)) ----> x - integral\<^isup>L M (f 0)"
-    using f ilim by (auto intro!: tendsto_diff simp: integral_diff)
-  note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5]
+    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_pos[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 integral_add(1))
   with diff(2) f show "integrable M u" "integral\<^isup>L M u = x"
-    by (auto simp: integral_diff)
+    by auto
 qed
 
 lemma integral_0_iff:
@@ -1993,58 +2040,47 @@
   using gt by (intro integral_less_AE[OF int, where A="space M"]) auto
 
 lemma integral_dominated_convergence:
-  assumes u: "\<And>i. integrable M (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
+  assumes u: "\<And>i. integrable M (u i)" and bound: "\<And>j. AE x in M. \<bar>u j x\<bar> \<le> w x"
   and w: "integrable M w"
-  and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
+  and u': "AE x in M. (\<lambda>i. u i x) ----> u' x"
+  and borel: "u' \<in> borel_measurable M"
   shows "integrable M u'"
   and "(\<lambda>i. (\<integral>x. \<bar>u i x - u' x\<bar> \<partial>M)) ----> 0" (is "?lim_diff")
   and "(\<lambda>i. integral\<^isup>L M (u i)) ----> integral\<^isup>L M u'" (is ?lim)
 proof -
-  { fix x j assume x: "x \<in> space M"
-    from u'[OF x] have "(\<lambda>i. \<bar>u i x\<bar>) ----> \<bar>u' x\<bar>" by (rule tendsto_rabs)
-    from LIMSEQ_le_const2[OF this]
-    have "\<bar>u' x\<bar> \<le> w x" using bound[OF x] by auto }
-  note u'_bound = this
+  have all_bound: "AE x in M. \<forall>j. \<bar>u j x\<bar> \<le> w x"
+    using bound by (auto simp: AE_all_countable)
+  with u' have u'_bound: "AE x in M. \<bar>u' x\<bar> \<le> w x"
+    by eventually_elim (auto intro: LIMSEQ_le_const2 tendsto_rabs)
 
-  from u[unfolded integrable_def]
-  have u'_borel: "u' \<in> borel_measurable M"
-    using u' by (blast intro: borel_measurable_LIMSEQ[of M u])
-
-  { fix x assume x: "x \<in> space M"
-    then have "0 \<le> \<bar>u 0 x\<bar>" by auto
-    also have "\<dots> \<le> w x" using bound[OF x] by auto
-    finally have "0 \<le> w x" . }
-  note w_pos = this
+  from bound[of 0] have w_pos: "AE x in M. 0 \<le> w x"
+    by eventually_elim auto
 
   show "integrable M u'"
-  proof (rule integrable_bound)
-    show "integrable M w" by fact
-    show "u' \<in> borel_measurable M" by fact
-  next
-    fix x assume x: "x \<in> space M" then show "0 \<le> w x" by fact
-    show "\<bar>u' x\<bar> \<le> w x" using u'_bound[OF x] .
-  qed
+    by (rule integrable_bound) fact+
 
   let ?diff = "\<lambda>n x. 2 * w x - \<bar>u n x - u' x\<bar>"
   have diff: "\<And>n. integrable M (\<lambda>x. \<bar>u n x - u' x\<bar>)"
-    using w u `integrable M u'`
-    by (auto intro!: integral_add integral_diff integral_cmult integrable_abs)
+    using w u `integrable M u'` by (auto intro!: integrable_abs)
 
-  { fix j x assume x: "x \<in> space M"
-    have "\<bar>u j x - u' x\<bar> \<le> \<bar>u j x\<bar> + \<bar>u' x\<bar>" by auto
+  from u'_bound all_bound
+  have diff_less_2w: "AE x in M. \<forall>j. \<bar>u j x - u' x\<bar> \<le> 2 * w x"
+  proof (eventually_elim, intro allI)
+    fix x j assume *: "\<bar>u' x\<bar> \<le> w x" "\<forall>j. \<bar>u j x\<bar> \<le> w x"
+    then have "\<bar>u j x - u' x\<bar> \<le> \<bar>u j x\<bar> + \<bar>u' x\<bar>" by auto
     also have "\<dots> \<le> w x + w x"
-      by (rule add_mono[OF bound[OF x] u'_bound[OF x]])
-    finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp }
-  note diff_less_2w = this
+      using * by (intro add_mono) auto
+    finally show "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp
+  qed
 
   have PI_diff: "\<And>n. (\<integral>\<^isup>+ x. ereal (?diff n x) \<partial>M) =
     (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) - (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
     using diff w diff_less_2w w_pos
     by (subst positive_integral_diff[symmetric])
-       (auto simp: integrable_def intro!: positive_integral_cong)
+       (auto simp: integrable_def intro!: positive_integral_cong_AE)
 
   have "integrable M (\<lambda>x. 2 * w x)"
-    using w by (auto intro: integral_cmult)
+    using w by auto
   hence I2w_fin: "(\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) \<noteq> \<infinity>" and
     borel_2w: "(\<lambda>x. ereal (2 * w x)) \<in> borel_measurable M"
     unfolding integrable_def by auto
@@ -2054,8 +2090,8 @@
     assume eq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) = 0" (is "?wx = 0")
     { fix n
       have "?f n \<le> ?wx" (is "integral\<^isup>P M ?f' \<le> _")
-        using diff_less_2w[of _ n] unfolding positive_integral_max_0
-        by (intro positive_integral_mono) auto
+        using diff_less_2w unfolding positive_integral_max_0
+        by (intro positive_integral_mono_AE) auto
       then have "?f n = 0"
         using positive_integral_positive[of M ?f'] eq_0 by auto }
     then show ?thesis by (simp add: Limsup_const)
@@ -2066,19 +2102,20 @@
       by (intro limsup_mono positive_integral_positive)
     finally have pos: "0 \<le> limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)" .
     have "?wx = (\<integral>\<^isup>+ x. liminf (\<lambda>n. max 0 (ereal (?diff n x))) \<partial>M)"
-    proof (rule positive_integral_cong)
-      fix x assume x: "x \<in> space M"
+      using u'
+    proof (intro positive_integral_cong_AE, eventually_elim)
+      fix x assume u': "(\<lambda>i. u i x) ----> u' x"
       show "max 0 (ereal (2 * w x)) = liminf (\<lambda>n. max 0 (ereal (?diff n x)))"
         unfolding ereal_max_0
       proof (rule lim_imp_Liminf[symmetric], unfold lim_ereal)
         have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>"
-          using u'[OF x] by (safe intro!: tendsto_intros)
+          using u' by (safe intro!: tendsto_intros)
         then show "(\<lambda>i. max 0 (?diff i x)) ----> max 0 (2 * w x)"
-          by (auto intro!: tendsto_real_max simp add: lim_ereal)
+          by (auto intro!: tendsto_real_max)
       qed (rule trivial_limit_sequentially)
     qed
     also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^isup>+ x. max 0 (ereal (?diff n x)) \<partial>M)"
-      using u'_borel w u unfolding integrable_def
+      using borel w u unfolding integrable_def
       by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF)
     also have "\<dots> = (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) -
         limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
@@ -2106,7 +2143,7 @@
     by (subst integral_eq_positive_integral[of _ M]) (auto simp: ereal_real integrable_def)
   then show ?lim_diff
     using ereal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
-    by (simp add: lim_ereal)
+    by simp
 
   show ?lim
   proof (rule LIMSEQ_I)
@@ -2119,9 +2156,9 @@
     proof (safe intro!: exI[of _ N])
       fix n assume "N \<le> n"
       have "\<bar>integral\<^isup>L M (u n) - integral\<^isup>L M u'\<bar> = \<bar>(\<integral>x. u n x - u' x \<partial>M)\<bar>"
-        using u `integrable M u'` by (auto simp: integral_diff)
+        using u `integrable M u'` by auto
       also have "\<dots> \<le> (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)" using u `integrable M u'`
-        by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff)
+        by (rule_tac integral_triangle_inequality) auto
       also note N[OF `N \<le> n`]
       finally show "norm (integral\<^isup>L M (u n) - integral\<^isup>L M u') < r" by simp
     qed
@@ -2139,6 +2176,8 @@
     using summable unfolding summable_def by auto
   from bchoice[OF this]
   obtain w where w: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. \<bar>f i x\<bar>) sums w x" by auto
+  then have w_borel: "w \<in> borel_measurable M" unfolding sums_def
+    by (rule borel_measurable_LIMSEQ) (auto simp: borel[THEN integrableD(1)])
 
   let ?w = "\<lambda>y. if y \<in> space M then w y else 0"
 
@@ -2146,13 +2185,16 @@
     using sums unfolding summable_def ..
 
   have 1: "\<And>n. integrable M (\<lambda>x. \<Sum>i = 0..<n. f i x)"
-    using borel by (auto intro!: integral_setsum)
+    using borel by auto
 
-  { fix j x assume [simp]: "x \<in> space M"
+  have 2: "\<And>j. AE x in M. \<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x"
+    using AE_space
+  proof eventually_elim
+    fix j x assume [simp]: "x \<in> space M"
     have "\<bar>\<Sum>i = 0..< j. f i x\<bar> \<le> (\<Sum>i = 0..< j. \<bar>f i x\<bar>)" by (rule setsum_abs)
     also have "\<dots> \<le> w x" using w[of x] series_pos_le[of "\<lambda>i. \<bar>f i x\<bar>"] unfolding sums_iff by auto
-    finally have "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp }
-  note 2 = this
+    finally show "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp
+  qed
 
   have 3: "integrable M ?w"
   proof (rule integral_monotone_convergence(1))
@@ -2161,21 +2203,22 @@
     have "\<And>n. integrable M (?F n)"
       using borel by (auto intro!: integral_setsum integrable_abs)
     thus "\<And>n. integrable M (?w' n)" by (simp cong: integrable_cong)
-    show "mono ?w'"
+    show "AE x in M. mono (\<lambda>n. ?w' n x)"
       by (auto simp: mono_def le_fun_def intro!: setsum_mono2)
-    { fix x show "(\<lambda>n. ?w' n x) ----> ?w x"
-        using w by (cases "x \<in> space M") (simp_all add: tendsto_const sums_def) }
+    show "AE x in M. (\<lambda>n. ?w' n x) ----> ?w x"
+        using w by (simp_all add: tendsto_const sums_def)
     have *: "\<And>n. integral\<^isup>L M (?w' n) = (\<Sum>i = 0..< n. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
-      using borel by (simp add: integral_setsum integrable_abs cong: integral_cong)
+      using borel by (simp add: integrable_abs cong: integral_cong)
     from abs_sum
     show "(\<lambda>i. integral\<^isup>L M (?w' i)) ----> x" unfolding * sums_def .
-  qed
+  qed (simp add: w_borel measurable_If_set)
 
   from summable[THEN summable_rabs_cancel]
-  have 4: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)"
+  have 4: "AE x in M. (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)"
     by (auto intro: summable_sumr_LIMSEQ_suminf)
 
-  note int = integral_dominated_convergence(1,3)[OF 1 2 3 4]
+  note int = integral_dominated_convergence(1,3)[OF 1 2 3 4
+    borel_measurable_suminf[OF integrableD(1)[OF borel]]]
 
   from int show "integrable M ?S" by simp
 
@@ -2244,59 +2287,54 @@
 
 section {* Distributions *}
 
-lemma simple_function_distr[simp]:
-  "simple_function (distr M M' T) f \<longleftrightarrow> simple_function M' (\<lambda>x. f x)"
- unfolding simple_function_def by simp
+lemma positive_integral_distr':
+  assumes T: "T \<in> measurable M M'"
+  and f: "f \<in> borel_measurable (distr M M' T)" "\<And>x. 0 \<le> f x"
+  shows "integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
+  using f 
+proof induct
+  case (cong f g)
+  with T show ?case
+    apply (subst positive_integral_cong[of _ f g])
+    apply simp
+    apply (subst positive_integral_cong[of _ "\<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 positive_integral_cong[OF eq])
+       (auto simp add: emeasure_distr intro!: positive_integral_indicator[symmetric] measurable_sets)
+qed (simp_all add: measurable_compose[OF T] T positive_integral_cmult positive_integral_add
+                   positive_integral_monotone_convergence_SUP le_fun_def incseq_def)
 
 lemma positive_integral_distr:
   assumes T: "T \<in> measurable M M'"
   and f: "f \<in> borel_measurable M'"
   shows "integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
-proof -
-  from borel_measurable_implies_simple_function_sequence'[OF f]
-  guess f' . note f' = this
-  then have f_distr: "\<And>i. simple_function (distr M M' T) (f' i)"
-    by simp
-  let ?f = "\<lambda>i x. f' i (T x)"
-  have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def)
-  have sup: "\<And>x. (SUP i. ?f i x) = max 0 (f (T x))"
-    using f'(4) .
-  have sf: "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
-    using simple_function_comp[OF T(1) f'(1)] .
-  show "integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
-    using
-      positive_integral_monotone_convergence_simple[OF f'(2,5) f_distr]
-      positive_integral_monotone_convergence_simple[OF inc f'(5) sf]
-    by (simp add: positive_integral_max_0 simple_integral_distr[OF T f'(1)] f')
-qed
+  apply (subst (1 2) positive_integral_max_0[symmetric])
+  apply (rule positive_integral_distr')
+  apply (auto simp: f T)
+  done
 
 lemma integral_distr:
-  assumes T: "T \<in> measurable M M'"
-  assumes f: "f \<in> borel_measurable M'"
-  shows "integral\<^isup>L (distr M M' T) f = (\<integral>x. f (T x) \<partial>M)"
-proof -
-  from measurable_comp[OF T, of f borel]
-  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
-    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
-    using f by (auto simp: comp_def)
-  then show ?thesis
-    using f unfolding lebesgue_integral_def integrable_def
-    by (auto simp: borel[THEN positive_integral_distr[OF T]])
-qed
+  "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^isup>L (distr M M' T) f = (\<integral> x. f (T x) \<partial>M)"
+  unfolding lebesgue_integral_def
+  by (subst (1 2) positive_integral_distr) auto
+
+lemma integrable_distr_eq:
+  assumes T: "T \<in> measurable M M'" "f \<in> borel_measurable M'"
+  shows "integrable (distr M M' T) f \<longleftrightarrow> integrable M (\<lambda>x. f (T x))"
+  using T measurable_comp[OF T]
+  unfolding integrable_def 
+  by (subst (1 2) positive_integral_distr) (auto simp: comp_def)
 
 lemma integrable_distr:
-  assumes T: "T \<in> measurable M M'" and f: "integrable (distr M M' T) f"
-  shows "integrable M (\<lambda>x. f (T x))"
-proof -
-  from measurable_comp[OF T, of f borel]
-  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
-    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
-    using f by (auto simp: comp_def)
-  then show ?thesis
-    using f unfolding lebesgue_integral_def integrable_def
-    using borel[THEN positive_integral_distr[OF T]]
-    by (auto simp: borel[THEN positive_integral_distr[OF T]])
-qed
+  assumes T: "T \<in> measurable M M'" shows "integrable (distr M M' T) f \<Longrightarrow> integrable M (\<lambda>x. f (T x))"
+  by (subst integrable_distr_eq[symmetric, OF T]) auto
 
 section {* Lebesgue integration on @{const count_space} *}
 
@@ -2329,6 +2367,26 @@
   by (subst positive_integral_max_0[symmetric])
      (auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le)
 
+lemma lebesgue_integral_count_space_finite_support:
+  assumes f: "finite {a\<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 *: "\<And>r::real. 0 < max 0 r \<longleftrightarrow> 0 < r" "\<And>x. max 0 (ereal x) = ereal (max 0 x)"
+    "\<And>a. a \<in> A \<and> 0 < f a \<Longrightarrow> max 0 (f a) = f a"
+    "\<And>a. a \<in> A \<and> f a < 0 \<Longrightarrow> max 0 (- f a) = - f a"
+    "{a \<in> A. f a \<noteq> 0} = {a \<in> A. 0 < f a} \<union> {a \<in> A. f a < 0}"
+    "({a \<in> A. 0 < f a} \<inter> {a \<in> A. f a < 0}) = {}"
+    by (auto split: split_max)
+  have "finite {a \<in> A. 0 < f a}" "finite {a \<in> A. f a < 0}"
+    by (auto intro: finite_subset[OF _ f])
+  then show ?thesis
+    unfolding lebesgue_integral_def
+    apply (subst (1 2) positive_integral_max_0[symmetric])
+    apply (subst (1 2) positive_integral_count_space)
+    apply (auto simp add: * setsum_negf setsum_Un
+                simp del: ereal_max)
+    done
+qed
+
 lemma lebesgue_integral_count_space_finite:
     "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
   apply (auto intro!: setsum_mono_zero_left
@@ -2337,6 +2395,10 @@
   apply (auto simp: max_def setsum_subtractf[symmetric] intro!: setsum_cong)
   done
 
+lemma borel_measurable_count_space[simp, intro!]:
+  "f \<in> borel_measurable (count_space A)"
+  by simp
+
 section {* Measure spaces with an associated density *}
 
 definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
@@ -2449,65 +2511,50 @@
        (auto elim: eventually_elim2)
 qed
 
-lemma positive_integral_density:
+lemma positive_integral_density':
   assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
-  assumes g': "g' \<in> borel_measurable M"
-  shows "integral\<^isup>P (density M f) g' = (\<integral>\<^isup>+ x. f x * g' x \<partial>M)"
-proof -
-  def g \<equiv> "\<lambda>x. max 0 (g' x)"
-  then have g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
-    using g' by auto
-  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this
-  note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)]
-  note G'(2)[simp]
-  { fix P have "AE x in M. P x \<Longrightarrow> AE x in M. P x"
-      using positive_integral_null_set[of _ _ f]
-      by (auto simp: eventually_ae_filter ) }
-  note ac = this
-  with G(4) f g have G_M': "AE x in density M f. (SUP i. G i x) = g x"
-    by (auto simp add: AE_density[OF f(1)] max_def)
-  { fix i
-    let ?I = "\<lambda>y x. indicator (G i -` {y} \<inter> space M) x"
-    { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
-      then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
-      from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
-        by (subst setsum_ereal_right_distrib) (auto simp: ac_simps)
-      also have "\<dots> = f x * G i x"
-        by (simp add: indicator_def if_distrib setsum_cases)
-      finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
-    note to_singleton = this
-    have "integral\<^isup>P (density M f) (G i) = integral\<^isup>S (density M f) (G i)"
-      using G by (intro positive_integral_eq_simple_integral) simp_all
-    also have "\<dots> = (\<Sum>y\<in>G i`space M. y * (\<integral>\<^isup>+x. f x * ?I y x \<partial>M))"
-      using f G(1)
-      by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_density
-               simp: simple_function_def simple_integral_def)
-    also have "\<dots> = (\<Sum>y\<in>G i`space M. (\<integral>\<^isup>+x. y * (f x * ?I y x) \<partial>M))"
-      using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric])
-    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) \<partial>M)"
-      using f G' G by (auto intro!: positive_integral_setsum[symmetric])
-    finally have "integral\<^isup>P (density M f) (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
-      using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) }
-  note [simp] = this
-  have "integral\<^isup>P (density M f) g = (SUP i. integral\<^isup>P (density M f) (G i))" using G'(1) G_M'(1) G
-    using positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`, of "density M f"]
-    by (simp cong: positive_integral_cong_AE)
-  also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f x * G i x \<partial>M))" by simp
-  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
-    using f G' G(2)[THEN incseq_SucD] G
-    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
-       (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff)
-  also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
-    by (intro positive_integral_cong_AE)
-       (auto simp add: SUPR_ereal_cmult split: split_max)
-  also have "\<dots> = (\<integral>\<^isup>+x. f x * g' x \<partial>M)"
-    using f(2)
-    by (subst (2) positive_integral_max_0[symmetric])
-       (auto simp: g_def max_def ereal_zero_le_0_iff intro!: positive_integral_cong_AE)
-  finally show "integral\<^isup>P (density M f) g' = (\<integral>\<^isup>+x. f x * g' x \<partial>M)"
-    unfolding g_def positive_integral_max_0 .
+  assumes g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
+  shows "integral\<^isup>P (density M f) g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
+using g proof induct
+  case (cong u v)
+  then show ?case
+    apply (subst positive_integral_cong[OF cong(3)])
+    apply (simp_all cong: positive_integral_cong)
+    done
+next
+  case (set A) then show ?case
+    by (simp add: emeasure_density f)
+next
+  case (mult u c)
+  moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps)
+  ultimately show ?case
+    by (simp add: f positive_integral_cmult)
+next
+  case (add u v)
+  moreover then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x"
+    by (simp add: ereal_right_distrib)
+  moreover note f
+  ultimately show ?case
+    by (auto simp add: positive_integral_add ereal_zero_le_0_iff intro!: positive_integral_add[symmetric])
+next
+  case (seq U)
+  from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)"
+    by eventually_elim (simp add: SUPR_ereal_cmult seq)
+  from seq f show ?case
+    apply (simp add: positive_integral_monotone_convergence_SUP)
+    apply (subst positive_integral_cong_AE[OF eq])
+    apply (subst positive_integral_monotone_convergence_SUP_AE)
+    apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono)
+    done
 qed
 
+lemma positive_integral_density:
+  "f \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> g' \<in> borel_measurable M \<Longrightarrow> 
+    integral\<^isup>P (density M f) g' = (\<integral>\<^isup>+ x. f x * g' x \<partial>M)"
+  by (subst (1 2) positive_integral_max_0[symmetric])
+     (auto intro!: positive_integral_cong_AE
+           simp: measurable_If max_def ereal_zero_le_0_iff positive_integral_density')
+
 lemma integral_density:
   assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
     and g: "g \<in> borel_measurable M"
@@ -2615,9 +2662,14 @@
 qed
 
 lemma emeasure_point_measure_finite:
-  "finite A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a|a\<in>X. f a)"
+  "finite A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<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_zero_left simp: less_le)
 
+lemma emeasure_point_measure_finite2:
+  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> (\<And>i. i \<in> X \<Longrightarrow> 0 \<le> f i) \<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_zero_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 \<le> 0)"
  by (auto simp: AE_count_space null_sets_density_iff point_measure_def)
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -9,20 +9,23 @@
   imports Finite_Product_Measure
 begin
 
+lemma borel_measurable_indicator':
+  "A \<in> sets borel \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M"
+  using measurable_comp[OF _ borel_measurable_indicator, of f M borel A] by (auto simp add: comp_def)
+
 lemma borel_measurable_sets:
   assumes "f \<in> measurable borel M" "A \<in> sets M"
   shows "f -` A \<in> sets borel"
   using measurable_sets[OF assms] by simp
 
-lemma measurable_identity[intro,simp]:
-  "(\<lambda>x. x) \<in> measurable M M"
-  unfolding measurable_def by auto
-
 subsection {* Standard Cubes *}
 
 definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
   "cube n \<equiv> {\<chi>\<chi> i. - real n .. \<chi>\<chi> i. real n}"
 
+lemma borel_cube[intro]: "cube n \<in> sets borel"
+  unfolding cube_def by auto
+
 lemma cube_closed[intro]: "closed (cube n)"
   unfolding cube_def by auto
 
@@ -154,7 +157,7 @@
     then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n"
       by (auto simp: sets_lebesgue)
     show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (\<Union>i. A i)"
-    proof (subst suminf_SUP_eq, safe intro!: incseq_SucI)
+    proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) 
       fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)"
         using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
     next
@@ -193,7 +196,6 @@
       qed
     qed
   qed
-next
 qed (auto, fact)
 
 lemma has_integral_interval_cube:
@@ -279,14 +281,16 @@
 
 lemma lmeasure_finite_has_integral:
   fixes s :: "'a::ordered_euclidean_space set"
-  assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m" "0 \<le> m"
+  assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m"
   shows "(indicator s has_integral m) UNIV"
 proof -
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
+  have "0 \<le> m"
+    using emeasure_nonneg[of lebesgue s] `emeasure lebesgue s = ereal m` by simp
   have **: "(?I s) integrable_on UNIV \<and> (\<lambda>k. integral UNIV (?I (s \<inter> cube k))) ----> integral UNIV (?I s)"
   proof (intro monotone_convergence_increasing allI ballI)
     have LIMSEQ: "(\<lambda>n. integral (cube n) (?I s)) ----> m"
-      using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1, 3)] .
+      using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] .
     { fix n have "integral (cube n) (?I s) \<le> m"
         using cube_subset assms
         by (intro incseq_le[where L=m] LIMSEQ incseq_def[THEN iffD2] integral_subset_le allI impI)
@@ -316,7 +320,7 @@
   note ** = conjunctD2[OF this]
   have m: "m = integral UNIV (?I s)"
     apply (intro LIMSEQ_unique[OF _ **(2)])
-    using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1,3)] integral_indicator_UNIV .
+    using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] integral_indicator_UNIV .
   show ?thesis
     unfolding m by (intro integrable_integral **)
 qed
@@ -366,14 +370,14 @@
 qed
 
 lemma has_integral_iff_lmeasure:
-  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m)"
+  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m)"
 proof
   assume "(indicator A has_integral m) UNIV"
   with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
-  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m"
+  show "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m"
     by (auto intro: has_integral_nonneg)
 next
-  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m"
+  assume "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m"
   then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
 qed
 
@@ -450,6 +454,9 @@
     by (auto simp: cube_def content_closed_interval_cases setprod_constant)
 qed simp
 
+lemma lmeasure_complete: "A \<subseteq> B \<Longrightarrow> B \<in> null_sets lebesgue \<Longrightarrow> A \<in> null_sets lebesgue"
+  unfolding negligible_iff_lebesgue_null_sets[symmetric] by (auto simp: negligible_subset)
+
 lemma
   fixes a b ::"'a::ordered_euclidean_space"
   shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})"
@@ -475,43 +482,44 @@
   fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0"
   using lmeasure_atLeastAtMost[of a a] by simp
 
+lemma AE_lebesgue_singleton:
+  fixes a :: "'a::ordered_euclidean_space" shows "AE x in lebesgue. x \<noteq> a"
+  by (rule AE_I[where N="{a}"]) auto
+
 declare content_real[simp]
 
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanAtMost[simp]:
     "emeasure lebesgue {a <.. b} = ereal (if a \<le> b then b - a else 0)"
-proof cases
-  assume "a < b"
-  then have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b} - emeasure lebesgue {a}"
-    by (subst emeasure_Diff[symmetric])
-       (auto intro!: arg_cong[where f="emeasure lebesgue"])
+proof -
+  have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b}"
+    using AE_lebesgue_singleton[of a]
+    by (intro emeasure_eq_AE) auto
   then show ?thesis by auto
-qed auto
+qed
 
 lemma
   fixes a b :: real
   shows lmeasure_real_atLeastLessThan[simp]:
     "emeasure lebesgue {a ..< b} = ereal (if a \<le> b then b - a else 0)"
-proof cases
-  assume "a < b"
-  then have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b} - emeasure lebesgue {b}"
-    by (subst emeasure_Diff[symmetric])
-       (auto intro!: arg_cong[where f="emeasure lebesgue"])
+proof -
+  have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b}"
+    using AE_lebesgue_singleton[of b]
+    by (intro emeasure_eq_AE) auto
   then show ?thesis by auto
-qed auto
+qed
 
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanLessThan[simp]:
     "emeasure lebesgue {a <..< b} = ereal (if a \<le> b then b - a else 0)"
-proof cases
-  assume "a < b"
-  then have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a <.. b} - emeasure lebesgue {b}"
-    by (subst emeasure_Diff[symmetric])
-       (auto intro!: arg_cong[where f="emeasure lebesgue"])
+proof -
+  have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a .. b}"
+    using AE_lebesgue_singleton[of a] AE_lebesgue_singleton[of b]
+    by (intro emeasure_eq_AE) auto
   then show ?thesis by auto
-qed auto
+qed
 
 subsection {* Lebesgue-Borel measure *}
 
@@ -544,6 +552,61 @@
     by (intro exI[of _ A]) (auto simp: subset_eq)
 qed
 
+lemma Int_stable_atLeastAtMost:
+  fixes x::"'a::ordered_euclidean_space"
+  shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
+  by (auto simp: inter_interval Int_stable_def)
+
+lemma lborel_eqI:
+  fixes M :: "'a::ordered_euclidean_space measure"
+  assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
+  assumes sets_eq: "sets M = sets borel"
+  shows "lborel = M"
+proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])
+  let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
+  let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
+  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
+    by (simp_all add: borel_eq_atLeastAtMost sets_eq)
+
+  show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto
+  { fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
+  then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto
+
+  { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
+  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
+      by (auto simp: emeasure_eq) }
+qed
+
+lemma lebesgue_real_affine:
+  fixes c :: real assumes "c \<noteq> 0"
+  shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D")
+proof (rule lborel_eqI)
+  fix a b show "emeasure ?D {a..b} = content {a .. b}"
+  proof cases
+    assume "0 < c"
+    then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
+      by (auto simp: field_simps)
+    with `0 < c` show ?thesis
+      by (cases "a \<le> b")
+         (auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult
+                     borel_measurable_indicator' emeasure_distr)
+  next
+    assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
+    then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
+      by (auto simp: field_simps)
+    with `c < 0` show ?thesis
+      by (cases "a \<le> b")
+         (auto simp: field_simps emeasure_density positive_integral_distr
+                     positive_integral_cmult borel_measurable_indicator' emeasure_distr)
+  qed
+qed simp
+
+lemma lebesgue_integral_real_affine:
+  fixes c :: real assumes c: "c \<noteq> 0" and f: "f \<in> borel_measurable borel"
+  shows "(\<integral> x. f x \<partial> lborel) = \<bar>c\<bar> * (\<integral> x. f (t + c * x) \<partial>lborel)"
+  by (subst lebesgue_real_affine[OF c, of t])
+     (simp add: f integral_density integral_distr lebesgue_integral_cmult)
+
 subsection {* Lebesgue integrable implies Gauge integrable *}
 
 lemma has_integral_cmult_real:
@@ -632,6 +695,69 @@
   qed
 qed
 
+lemma borel_measurable_real_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]:
+  fixes u :: "'a \<Rightarrow> real"
+  assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x"
+  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. 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> P v \<Longrightarrow> P (\<lambda>x. v x + u x)"
+  assumes seq: "\<And>U u. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>x. (\<lambda>i. U i x) ----> u x) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> P u"
+  shows "P u"
+proof -
+  have "(\<lambda>x. ereal (u x)) \<in> borel_measurable M"
+    using u by auto
+  then obtain U where U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i. \<infinity> \<notin> range (U i)" and
+    "\<And>x. (SUP i. U i x) = max 0 (ereal (u x))" and nn: "\<And>i x. 0 \<le> U i x"
+    using borel_measurable_implies_simple_function_sequence'[of u M] by auto
+  then have u_eq: "\<And>x. ereal (u x) = (SUP i. U i x)"
+    using u by (auto simp: max_def)
+
+  have [simp]: "\<And>i x. U i x \<noteq> \<infinity>" using U by (auto simp: image_def eq_commute)
+
+  { fix i x have [simp]: "U i x \<noteq> -\<infinity>" using nn[of i x] by auto }
+  note this[simp]
+
+  show "P u"
+  proof (rule seq)
+    show "\<And>i. (\<lambda>x. real (U i x)) \<in> borel_measurable M"
+      using U by (auto intro: borel_measurable_simple_function)
+    show "\<And>i x. 0 \<le> real (U i x)"
+      using nn by (auto simp: real_of_ereal_pos)
+    show "incseq (\<lambda>i x. real (U i x))"
+      using U(2) by (auto simp: incseq_def image_iff le_fun_def intro!: real_of_ereal_positive_mono nn)
+    then show "\<And>x. (\<lambda>i. real (U i x)) ----> u x"
+      by (intro SUP_eq_LIMSEQ[THEN iffD1])
+         (auto simp: incseq_mono incseq_def le_fun_def u_eq ereal_real
+               intro!: arg_cong2[where f=SUPR] ext)
+    show "\<And>i. P (\<lambda>x. real (U i x))"
+    proof (rule cong)
+      fix x i assume x: "x \<in> space M"
+      have [simp]: "\<And>A x. real (indicator A x :: ereal) = indicator A x"
+        by (auto simp: indicator_def one_ereal_def)
+      { fix y assume "y \<in> U i ` space M"
+        then have "0 \<le> y" "y \<noteq> \<infinity>" using nn by auto
+        then have "\<bar>y * indicator (U i -` {y} \<inter> space M) x\<bar> \<noteq> \<infinity>"
+          by (auto simp: indicator_def) }
+      then show "real (U i x) = (\<Sum>y \<in> U i ` space M. real y * indicator (U i -` {y} \<inter> space M) x)"
+        unfolding simple_function_indicator_representation[OF U(1) x]
+        by (subst setsum_real_of_ereal[symmetric]) auto
+    next
+      fix i
+      have "finite (U i ` space M)" "\<And>x. x \<in> U i ` space M \<Longrightarrow> 0 \<le> x""\<And>x. x \<in> U i ` space M \<Longrightarrow> x \<noteq> \<infinity>"
+        using U(1) nn by (auto simp: simple_function_def)
+      then show "P (\<lambda>x. \<Sum>y \<in> U i ` space M. real y * indicator (U i -` {y} \<inter> space M) x)"
+      proof induct
+        case empty then show ?case
+          using set[of "{}"] by (simp add: indicator_def[abs_def])
+      qed (auto intro!: add mult set simple_functionD U setsum_nonneg borel_measurable_setsum mult_nonneg_nonneg real_of_ereal_pos)
+    qed (auto intro: borel_measurable_simple_function U simple_functionD intro!: borel_measurable_setsum borel_measurable_times)
+  qed
+qed
+
+lemma ereal_indicator: "ereal (indicator A x) = indicator A x"
+  by (auto simp: indicator_def one_ereal_def)
+
 lemma positive_integral_has_integral:
   fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
@@ -777,201 +903,22 @@
     unfolding lebesgue_integral_eq_borel[OF borel] by simp
 qed
 
-subsection {* Equivalence between product spaces and euclidean spaces *}
-
-definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where
-  "e2p x = (\<lambda>i\<in>{..<DIM('a)}. x$$i)"
-
-definition p2e :: "(nat \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
-  "p2e x = (\<chi>\<chi> i. x i)"
-
-lemma e2p_p2e[simp]:
-  "x \<in> extensional {..<DIM('a)} \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
-  by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def)
-
-lemma p2e_e2p[simp]:
-  "p2e (e2p x) = (x::'a::ordered_euclidean_space)"
-  by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def)
-
-interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"
-  by default
-
-interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "{..<n}" for n :: nat
-  by default auto
-
-lemma bchoice_iff: "(\<forall>x\<in>A. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>A. P x (f x))"
-  by metis
-
-lemma sets_product_borel:
-  assumes I: "finite I"
-  shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
-proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
-  show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
-    by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)
-qed (auto simp: borel_eq_lessThan incseq_def reals_Archimedean2 image_iff intro: real_natceiling_ge)
-
-lemma measurable_e2p:
-  "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))"
-proof (rule measurable_sigma_sets[OF sets_product_borel])
-  fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i} |x. True} "
-  then obtain x where  "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i})" by auto
-  then have "e2p -` A = {..< (\<chi>\<chi> i. x i) :: 'a}"
-    using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def
-      euclidean_eq[where 'a='a] eucl_less[where 'a='a])
-  then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp
-qed (auto simp: e2p_def)
-
-lemma measurable_p2e:
-  "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))
-    (borel :: 'a::ordered_euclidean_space measure)"
-  (is "p2e \<in> measurable ?P _")
-proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
-  fix x i
-  let ?A = "{w \<in> space ?P. (p2e w :: 'a) $$ i \<le> x}"
-  assume "i < DIM('a)"
-  then have "?A = (\<Pi>\<^isub>E j\<in>{..<DIM('a)}. if i = j then {.. x} else UNIV)"
-    using DIM_positive by (auto simp: space_PiM p2e_def split: split_if_asm)
-  then show "?A \<in> sets ?P"
-    by auto
-qed
-
-lemma Int_stable_atLeastAtMost:
-  fixes x::"'a::ordered_euclidean_space"
-  shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
-  by (auto simp: inter_interval Int_stable_def)
-
-lemma lborel_eqI:
-  fixes M :: "'a::ordered_euclidean_space measure"
-  assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
-  assumes sets_eq: "sets M = sets borel"
-  shows "lborel = M"
-proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])
-  let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
-  let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
-  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
-    by (simp_all add: borel_eq_atLeastAtMost sets_eq)
-
-  show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto
-  show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI)
-  { fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
-  then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto
-
-  { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
-  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
-      by (auto simp: emeasure_eq) }
-qed
-
-lemma lborel_eq_lborel_space:
-  "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel) lborel p2e"
-  (is "?B = ?D")
-proof (rule lborel_eqI)
-  show "sets ?D = sets borel" by simp
-  let ?P = "(\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel)"
-  fix a b :: 'a
-  have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
-    by (auto simp: Pi_iff eucl_le[where 'a='a] p2e_def space_PiM)
-  have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
-  proof cases
-    assume "{a..b} \<noteq> {}"
-    then have "a \<le> b"
-      by (simp add: interval_ne_empty eucl_le[where 'a='a])
-    then have "emeasure lborel {a..b} = (\<Prod>x<DIM('a). emeasure lborel {a $$ x .. b $$ x})"
-      by (auto simp: content_closed_interval eucl_le[where 'a='a]
-               intro!: setprod_ereal[symmetric])
-    also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
-      unfolding * by (subst lborel_space.measure_times) auto
-    finally show ?thesis by simp
-  qed simp
-  then show "emeasure ?D {a .. b} = content {a .. b}"
-    by (simp add: emeasure_distr measurable_p2e)
-qed
-
-lemma borel_fubini_positiv_integral:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
-  assumes f: "f \<in> borel_measurable borel"
-  shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel)"
-  by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)
-
-lemma borel_fubini_integrable:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
-  shows "integrable lborel f \<longleftrightarrow>
-    integrable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) (\<lambda>x. f (p2e x))"
-    (is "_ \<longleftrightarrow> integrable ?B ?f")
-proof
-  assume "integrable lborel f"
-  moreover then have f: "f \<in> borel_measurable borel"
-    by auto
-  moreover with measurable_p2e
-  have "f \<circ> p2e \<in> borel_measurable ?B"
-    by (rule measurable_comp)
-  ultimately show "integrable ?B ?f"
-    by (simp add: comp_def borel_fubini_positiv_integral integrable_def)
-next
-  assume "integrable ?B ?f"
-  moreover
-  then have "?f \<circ> e2p \<in> borel_measurable (borel::'a measure)"
-    by (auto intro!: measurable_e2p)
-  then have "f \<in> borel_measurable borel"
-    by (simp cong: measurable_cong)
-  ultimately show "integrable lborel f"
-    by (simp add: borel_fubini_positiv_integral integrable_def)
-qed
-
-lemma borel_fubini:
-  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable borel"
-  shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
-  using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
-
-lemma borel_measurable_indicator':
-  "A \<in> sets borel \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M"
-  using measurable_comp[OF _ borel_measurable_indicator, of f M borel A] by (auto simp add: comp_def)
-
-lemma lebesgue_real_affine:
-  fixes c :: real assumes "c \<noteq> 0"
-  shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D")
-proof (rule lborel_eqI)
-  fix a b show "emeasure ?D {a..b} = content {a .. b}"
-  proof cases
-    assume "0 < c"
-    then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
-      by (auto simp: field_simps)
-    with `0 < c` show ?thesis
-      by (cases "a \<le> b")
-         (auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult
-                     borel_measurable_indicator' emeasure_distr)
-  next
-    assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
-    then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
-      by (auto simp: field_simps)
-    with `c < 0` show ?thesis
-      by (cases "a \<le> b")
-         (auto simp: field_simps emeasure_density positive_integral_distr
-                     positive_integral_cmult borel_measurable_indicator' emeasure_distr)
-  qed
-qed simp
-
-lemma borel_cube[intro]: "cube n \<in> sets borel"
-  unfolding cube_def by auto
-
 lemma integrable_on_cmult_iff:
   fixes c :: real assumes "c \<noteq> 0"
   shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s"
   using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] `c \<noteq> 0`
   by auto
 
-lemma positive_integral_borel_has_integral:
+lemma positive_integral_lebesgue_has_integral:
   fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
-  assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"
+  assumes f_borel: "f \<in> borel_measurable lebesgue" and nonneg: "\<And>x. 0 \<le> f x"
   assumes I: "(f has_integral I) UNIV"
-  shows "(\<integral>\<^isup>+x. f x \<partial>lborel) = I"
+  shows "(\<integral>\<^isup>+x. f x \<partial>lebesgue) = I"
 proof -
-  from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable borel" by auto
+  from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue" by auto
   from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this
 
-  have lebesgue_eq: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (f x) \<partial>lborel)"
-    using f_borel by (intro lebesgue_positive_integral_eq_borel) auto
-  also have "\<dots> = (SUP i. integral\<^isup>S lborel (F i))"
+  have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^isup>S lebesgue (F i))"
     using F
     by (subst positive_integral_monotone_convergence_simple)
        (simp_all add: positive_integral_max_0 simple_function_def)
@@ -1043,11 +990,8 @@
       unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast
     moreover have "0 \<le> integral\<^isup>S lebesgue (F i)"
       using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def)
-    moreover have "integral\<^isup>S lebesgue (F i) = integral\<^isup>S lborel (F i)"
-      using F(1)[of i, THEN borel_measurable_simple_function]
-      by (rule lebesgue_simple_integral_eq_borel)
-    ultimately show "integral\<^isup>S lborel (F i) \<le> ereal I"
-      by (cases "integral\<^isup>S lborel (F i)") auto
+    ultimately show "integral\<^isup>S lebesgue (F i) \<le> ereal I"
+      by (cases "integral\<^isup>S lebesgue (F i)") auto
   qed
   also have "\<dots> < \<infinity>" by simp
   finally have finite: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) \<noteq> \<infinity>" by simp
@@ -1059,14 +1003,142 @@
   with I have "I = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)"
     by (rule has_integral_unique)
   with finite positive_integral_positive[of _ "\<lambda>x. ereal (f x)"] show ?thesis
-    by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>lborel") (auto simp: lebesgue_eq)
+    by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue") auto
 qed
 
-lemma has_integral_iff_positive_integral:
+lemma has_integral_iff_positive_integral_lebesgue:
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable lebesgue" "\<And>x. 0 \<le> f x"
+  shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lebesgue f = I"
+  using f positive_integral_lebesgue_has_integral[of f I] positive_integral_has_integral[of f]
+  by (auto simp: subset_eq)
+
+lemma has_integral_iff_positive_integral_lborel:
   fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x"
   shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lborel f = I"
-  using f positive_integral_borel_has_integral[of f I] positive_integral_has_integral[of f]
-  by (auto simp: subset_eq borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel)
+  using assms
+  by (subst has_integral_iff_positive_integral_lebesgue)
+     (auto simp: borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel)
+
+subsection {* Equivalence between product spaces and euclidean spaces *}
+
+definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where
+  "e2p x = (\<lambda>i\<in>{..<DIM('a)}. x$$i)"
+
+definition p2e :: "(nat \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where
+  "p2e x = (\<chi>\<chi> i. x i)"
+
+lemma e2p_p2e[simp]:
+  "x \<in> extensional {..<DIM('a)} \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
+  by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def)
+
+lemma p2e_e2p[simp]:
+  "p2e (e2p x) = (x::'a::ordered_euclidean_space)"
+  by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def)
+
+interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"
+  by default
+
+interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "{..<n}" for n :: nat
+  by default auto
+
+lemma bchoice_iff: "(\<forall>x\<in>A. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>A. P x (f x))"
+  by metis
+
+lemma sets_product_borel:
+  assumes I: "finite I"
+  shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
+proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
+  show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
+    by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)
+qed (auto simp: borel_eq_lessThan reals_Archimedean2)
+
+lemma measurable_e2p:
+  "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))"
+proof (rule measurable_sigma_sets[OF sets_product_borel])
+  fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i} |x. True} "
+  then obtain x where  "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i})" by auto
+  then have "e2p -` A = {..< (\<chi>\<chi> i. x i) :: 'a}"
+    using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def
+      euclidean_eq[where 'a='a] eucl_less[where 'a='a])
+  then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp
+qed (auto simp: e2p_def)
+
+lemma measurable_p2e:
+  "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))
+    (borel :: 'a::ordered_euclidean_space measure)"
+  (is "p2e \<in> measurable ?P _")
+proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
+  fix x i
+  let ?A = "{w \<in> space ?P. (p2e w :: 'a) $$ i \<le> x}"
+  assume "i < DIM('a)"
+  then have "?A = (\<Pi>\<^isub>E j\<in>{..<DIM('a)}. if i = j then {.. x} else UNIV)"
+    using DIM_positive by (auto simp: space_PiM p2e_def split: split_if_asm)
+  then show "?A \<in> sets ?P"
+    by auto
+qed
+
+lemma lborel_eq_lborel_space:
+  "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel) borel p2e"
+  (is "?B = ?D")
+proof (rule lborel_eqI)
+  show "sets ?D = sets borel" by simp
+  let ?P = "(\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel)"
+  fix a b :: 'a
+  have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
+    by (auto simp: Pi_iff eucl_le[where 'a='a] p2e_def space_PiM)
+  have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
+  proof cases
+    assume "{a..b} \<noteq> {}"
+    then have "a \<le> b"
+      by (simp add: interval_ne_empty eucl_le[where 'a='a])
+    then have "emeasure lborel {a..b} = (\<Prod>x<DIM('a). emeasure lborel {a $$ x .. b $$ x})"
+      by (auto simp: content_closed_interval eucl_le[where 'a='a]
+               intro!: setprod_ereal[symmetric])
+    also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
+      unfolding * by (subst lborel_space.measure_times) auto
+    finally show ?thesis by simp
+  qed simp
+  then show "emeasure ?D {a .. b} = content {a .. b}"
+    by (simp add: emeasure_distr measurable_p2e)
+qed
+
+lemma borel_fubini_positiv_integral:
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
+  assumes f: "f \<in> borel_measurable borel"
+  shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel)"
+  by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)
+
+lemma borel_fubini_integrable:
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
+  shows "integrable lborel f \<longleftrightarrow>
+    integrable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) (\<lambda>x. f (p2e x))"
+    (is "_ \<longleftrightarrow> integrable ?B ?f")
+proof
+  assume "integrable lborel f"
+  moreover then have f: "f \<in> borel_measurable borel"
+    by auto
+  moreover with measurable_p2e
+  have "f \<circ> p2e \<in> borel_measurable ?B"
+    by (rule measurable_comp)
+  ultimately show "integrable ?B ?f"
+    by (simp add: comp_def borel_fubini_positiv_integral integrable_def)
+next
+  assume "integrable ?B ?f"
+  moreover
+  then have "?f \<circ> e2p \<in> borel_measurable (borel::'a measure)"
+    by (auto intro!: measurable_e2p)
+  then have "f \<in> borel_measurable borel"
+    by (simp cong: measurable_cong)
+  ultimately show "integrable lborel f"
+    by (simp add: borel_fubini_positiv_integral integrable_def)
+qed
+
+lemma borel_fubini:
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable borel"
+  shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
+  using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
 
 end
--- a/src/HOL/Probability/Measure_Space.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Measure_Space.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -12,6 +12,12 @@
   "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
 begin
 
+lemma sums_def2:
+  "f sums x \<longleftrightarrow> (\<lambda>n. (\<Sum>i\<le>n. f i)) ----> x"
+  unfolding sums_def
+  apply (subst LIMSEQ_Suc_iff[symmetric])
+  unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost ..
+
 lemma suminf_cmult_indicator:
   fixes f :: "nat \<Rightarrow> ereal"
   assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
@@ -90,6 +96,9 @@
 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 positiveD2: "positive M f \<Longrightarrow> A \<in> M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
+
 lemma positiveD_empty:
   "positive M f \<Longrightarrow> f {} = 0"
   by (auto simp add: positive_def)
@@ -240,6 +249,143 @@
   finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
 qed
 
+lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
+  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)) ----> 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=0..<n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
+    using f(1)[unfolded positive_def] dA
+    by (auto intro!: summable_sumr_LIMSEQ_suminf summable_ereal_pos)
+  from LIMSEQ_Suc[OF this]
+  have "(\<lambda>n. (\<Sum>i\<le>n. f (disjointed A i))) ----> (\<Sum>i. f (disjointed A i))"
+    unfolding atLeastLessThanSuc_atLeastAtMost atLeast0AtMost .
+  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)) ----> 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)) ----> 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\<le>n. A i)) = (\<Union>i. A i)" by auto
+  have "(\<lambda>n. f (\<Union>i\<le>n. A i)) ----> f (\<Union>i. A i)"
+  proof (unfold *[symmetric], intro cont[rule_format])
+    show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> M"
+      using A * by auto
+  qed (force intro!: incseq_SucI)
+  moreover have "\<And>n. f (\<Union>i\<le>n. A i) = (\<Sum>i\<le>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_def2 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:
+  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)) ----> 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)) ----> 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)) ----> 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)) ----> 0"
+    using `positive M f`[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)) ----> 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
+  { 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 }
+  note f_fin = this
+  have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 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 INFI_ereal_add) (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_ereal_INFI[OF decseq_fA]
+  show "(\<lambda>i. f (A i)) ----> f (\<Inter>i. A i)" by simp
+qed
+
+lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
+  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)) ----> 0"
+  assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
+  shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
+proof -
+  have "\<forall>A\<in>M. \<exists>x. f A = ereal x"
+  proof
+    fix A assume "A \<in> M" with f show "\<exists>x. f A = ereal x"
+      unfolding positive_def by (cases "f A") auto
+  qed
+  from bchoice[OF this] guess f' .. note f' = this[rule_format]
+  from A have "(\<lambda>i. f ((\<Union>i. A i) - A i)) ----> 0"
+    by (intro cont[rule_format]) (auto simp: decseq_def incseq_def)
+  moreover
+  { fix i
+    have "f ((\<Union>i. A i) - A i) + f (A i) = f ((\<Union>i. A i) - A i \<union> A i)"
+      using A by (intro f(2)[THEN additiveD, symmetric]) 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) - f' (A i) = f' ((\<Union>i. A i) - A i)"
+      using A by (subst (asm) (1 2 3) f') auto
+    then have "f ((\<Union>i. A i) - A i) = ereal (f' (\<Union>i. A i) - f' (A i))"
+      using A f' by auto }
+  ultimately have "(\<lambda>i. f' (\<Union>i. A i) - f' (A i)) ----> 0"
+    by (simp add: zero_ereal_def)
+  then have "(\<lambda>i. f' (A i)) ----> f' (\<Union>i. A i)"
+    by (rule LIMSEQ_diff_approach_zero2[OF tendsto_const])
+  then show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
+    using A by (subst (1 2) f') auto
+qed
+
+lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
+  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)) ----> 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
+
 section {* Properties of @{const emeasure} *}
 
 lemma emeasure_positive: "positive (sets M) (emeasure M)"
@@ -314,84 +460,21 @@
     using finite `0 \<le> emeasure M B` by auto
 qed
 
-lemma emeasure_countable_increasing:
-  assumes A: "range A \<subseteq> sets M"
-      and A0: "A 0 = {}"
-      and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
-  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
-proof -
-  { fix n
-    have "emeasure M (A n) = (\<Sum>i<n. emeasure M (A (Suc i) - A i))"
-      proof (induct n)
-        case 0 thus ?case by (auto simp add: A0)
-      next
-        case (Suc m)
-        have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
-          by (metis ASuc Un_Diff_cancel Un_absorb1)
-        hence "emeasure M (A (Suc m)) =
-               emeasure M (A m) + emeasure M (A (Suc m) - A m)"
-          by (subst plus_emeasure)
-             (auto simp add: emeasure_additive range_subsetD [OF A])
-        with Suc show ?case
-          by simp
-      qed }
-  note Meq = this
-  have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
-    proof (rule UN_finite2_eq [where k=1], simp)
-      fix i
-      show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
-        proof (induct i)
-          case 0 thus ?case by (simp add: A0)
-        next
-          case (Suc i)
-          thus ?case
-            by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
-        qed
-    qed
-  have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
-    by (metis A Diff range_subsetD)
-  have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
-    by (blast intro: range_subsetD [OF A])
-  have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = (\<Sum>i. emeasure M (A (Suc i) - A i))"
-    using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
-  also have "\<dots> = emeasure M (\<Union>i. A (Suc i) - A i)"
-    by (rule suminf_emeasure)
-       (auto simp add: disjoint_family_Suc ASuc A1 A2)
-  also have "... =  emeasure M (\<Union>i. A i)"
-    by (simp add: Aeq)
-  finally have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = emeasure M (\<Union>i. A i)" .
-  then show ?thesis by (auto simp add: Meq)
-qed
-
-lemma SUP_emeasure_incseq:
-  assumes A: "range A \<subseteq> sets M" and "incseq A"
-  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
-proof -
-  have *: "(SUP n. emeasure M (nat_case {} A (Suc n))) = (SUP n. emeasure M (nat_case {} A n))"
-    using A by (auto intro!: SUPR_eq exI split: nat.split)
-  have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
-    by (auto simp add: split: nat.splits)
-  have meq: "\<And>n. emeasure M (A n) = (emeasure M \<circ> nat_case {} A) (Suc n)"
-    by simp
-  have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. nat_case {} A i)"
-    using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
-    by (force split: nat.splits intro!: emeasure_countable_increasing)
-  also have "emeasure M (\<Union>i. nat_case {} A i) = emeasure M (\<Union>i. A i)"
-    by (simp add: ueq)
-  finally have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. A i)" .
-  thus ?thesis unfolding meq * comp_def .
-qed
+lemma Lim_emeasure_incseq:
+  "range A \<subseteq> sets M \<Longrightarrow> incseq A \<Longrightarrow> (\<lambda>i. (emeasure M (A i))) ----> emeasure M (\<Union>i. A i)"
+  using emeasure_countably_additive
+  by (auto simp add: 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 Lim_emeasure_incseq:
+lemma SUP_emeasure_incseq:
   assumes A: "range A \<subseteq> sets M" "incseq A"
-  shows "(\<lambda>i. (emeasure M (A i))) ----> emeasure M (\<Union>i. A i)"
-  using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A]
-    SUP_emeasure_incseq[OF A] by simp
+  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
+  using LIMSEQ_ereal_SUPR[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"
@@ -545,70 +628,78 @@
   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" "incseq A" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  and A: "range A \<subseteq> E" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
   shows "M = N"
 proof -
-  let ?D = "\<lambda>F. {D. D \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)}"
+  let ?\<mu>  = "emeasure M" and ?\<nu> = "emeasure N"
   interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
-  { fix F assume "F \<in> E" and "emeasure M F \<noteq> \<infinity>"
+  have "space M = \<Omega>"
+    using top[of M] space_closed[of M] S.top S.space_closed `sets M = sigma_sets \<Omega> E` 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 "emeasure N F \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` `F \<in> E` eq by simp
-    interpret D: dynkin_system \<Omega> "?D F"
-    proof (rule dynkin_systemI, simp_all)
-      fix A assume "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
-      then show "A \<subseteq> \<Omega>" using S.sets_into_space by auto
+    have "?\<nu> F \<noteq> \<infinity>" using `?\<mu> F \<noteq> \<infinity>` `F \<in> E` eq by simp
+    assume "D \<in> sets M"
+    with `Int_stable E` `E \<subseteq> Pow \<Omega>` 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 `Int_stable E` `F \<in> E` by (auto simp: Int_stable_def)
+      then show ?case using eq by auto
     next
-      have "F \<inter> \<Omega> = F" using `F \<in> E` `E \<subseteq> Pow \<Omega>` by auto
-      then show "emeasure M (F \<inter> \<Omega>) = emeasure N (F \<inter> \<Omega>)"
-        using `F \<in> E` eq by (auto intro: sigma_sets_top)
+      case empty then show ?case by simp
     next
-      fix A assume *: "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
+      case (compl A)
       then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
         and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
-        using `F \<in> E` S.sets_into_space by auto
-      have "emeasure N (F \<inter> A) \<le> emeasure N F" by (auto intro!: emeasure_mono simp: M N)
-      then have "emeasure N (F \<inter> A) \<noteq> \<infinity>" using `emeasure N F \<noteq> \<infinity>` by auto
-      have "emeasure M (F \<inter> A) \<le> emeasure M F" by (auto intro!: emeasure_mono simp: M N)
-      then have "emeasure M (F \<inter> A) \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` by auto
-      then have "emeasure M (F \<inter> (\<Omega> - A)) = emeasure M F - emeasure M (F \<inter> A)" unfolding **
+        using `F \<in> E` 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 `?\<nu> F \<noteq> \<infinity>` by auto
+      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 `?\<mu> F \<noteq> \<infinity>` by auto
+      then have "?\<mu> (F \<inter> (\<Omega> - A)) = ?\<mu> F - ?\<mu> (F \<inter> A)" unfolding **
         using `F \<inter> A \<in> sigma_sets \<Omega> E` by (auto intro!: emeasure_Diff simp: M N)
-      also have "\<dots> = emeasure N F - emeasure N (F \<inter> A)" using eq `F \<in> E` * by simp
-      also have "\<dots> = emeasure N (F \<inter> (\<Omega> - A))" unfolding **
-        using `F \<inter> A \<in> sigma_sets \<Omega> E` `emeasure N (F \<inter> A) \<noteq> \<infinity>`
+      also have "\<dots> = ?\<nu> F - ?\<nu> (F \<inter> A)" using eq `F \<in> E` compl by simp
+      also have "\<dots> = ?\<nu> (F \<inter> (\<Omega> - A))" unfolding **
+        using `F \<inter> A \<in> sigma_sets \<Omega> E` `?\<nu> (F \<inter> A) \<noteq> \<infinity>`
         by (auto intro!: emeasure_Diff[symmetric] simp: M N)
-      finally show "\<Omega> - A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Omega> - A)) = emeasure N (F \<inter> (\<Omega> - A))"
-        using * by auto
+      finally show ?case
+        using `space M = \<Omega>` by auto
     next
-      fix A :: "nat \<Rightarrow> 'a set"
-      assume "disjoint_family A" "range A \<subseteq> {X \<in> sigma_sets \<Omega> E. emeasure M (F \<inter> X) = emeasure N (F \<inter> X)}"
-      then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sigma_sets \<Omega> E" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
-        "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. emeasure M (F \<inter> A i) = emeasure N (F \<inter> A i)" "range A \<subseteq> sigma_sets \<Omega> E"
-        by (auto simp: disjoint_family_on_def subset_eq)
-      then show "(\<Union>x. A x) \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Union>x. A x)) = emeasure N (F \<inter> (\<Union>x. A x))"
-        by (auto simp: M N suminf_emeasure[symmetric] simp del: UN_simps)
-    qed
-    have *: "sigma_sets \<Omega> E = ?D F"
-      using `F \<in> E` `Int_stable E`
-      by (intro D.dynkin_lemma) (auto simp add: Int_stable_def eq)
-    have "\<And>D. D \<in> sigma_sets \<Omega> E \<Longrightarrow> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
-      by (subst (asm) *) auto }
+      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
-    fix X assume "X \<in> sets M"
-    then have "X \<in> sigma_sets \<Omega> E"
-      using M by simp
-    let ?A = "\<lambda>i. A i \<inter> X"
-    have "range ?A \<subseteq> sigma_sets \<Omega> E" "incseq ?A"
-      using A(1,2) `X \<in> sigma_sets \<Omega> E` by (auto simp: incseq_def)
-    moreover
-    { fix i have "emeasure M (?A i) = emeasure N (?A i)"
-        using *[of "A i" X] `X \<in> sigma_sets \<Omega> E` A finite by auto }
-    ultimately show "emeasure M X = emeasure N X"
-      using SUP_emeasure_incseq[of ?A M] SUP_emeasure_incseq[of ?A N] A(3) `X \<in> sigma_sets \<Omega> E`
-      by (auto simp: M N SUP_emeasure_incseq)
+    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 `space M = \<Omega>` have F_eq: "F = (\<Union>i. ?D i)"
+      using `F \<in> sets M`[THEN sets_into_space] A(2)[symmetric] by (auto simp: UN_disjointed_eq)
+    have [simp, intro]: "\<And>i. ?D i \<in> sets M"
+      using range_disjointed_sets[of "\<lambda>i. F \<inter> A i" M] `F \<in> sets M`
+      by (auto simp: subset_eq)
+    have "disjoint_family ?D"
+      by (auto simp: disjoint_family_disjointed)
+     moreover
+    { fix i
+      have "A i \<inter> ?D i = ?D i"
+        by (auto simp: disjointed_def)
+      then have "emeasure M (?D i) = emeasure N (?D i)"
+        using *[of "A i" "?D i", OF _ A(3)] A(1) by auto }
+     ultimately show "emeasure M F = emeasure N F"
+      using N M
+      apply (subst (1 2) F_eq)
+      apply (subst (1 2) suminf_emeasure[symmetric])
+      apply auto
+      done
   qed
 qed
 
@@ -1016,6 +1107,24 @@
        (auto simp: emeasure_distr measurable_def)
 qed
 
+lemma AE_distr_iff:
+  assumes f: "f \<in> measurable M N" and P: "{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])
+  from P show "{x \<in> space (distr M N f). \<not> P x} \<in> sets (distr M N f)"
+    by (auto intro!: sets_Collect_neg)
+  moreover
+  have "f -` {x \<in> space N. P x} \<inter> space M = {x \<in> space M. P (f x)}"
+    using f by (auto dest: measurable_space)
+  then show "{x \<in> space M. \<not> P (f x)} \<in> sets M"
+    using measurable_sets[OF f P] by (auto intro!: sets_Collect_neg)
+  moreover have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
+    using f by (auto dest: measurable_space)
+  ultimately 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)"
+    using f by (simp add: emeasure_distr)
+qed
+
 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 measurable_sets)
@@ -1295,103 +1404,81 @@
     unfolding measurable_def by auto
 qed
 
+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 ereal (card X) else \<infinity>)"
     (is "_ = ?M X")
   unfolding count_space_def
 proof (rule emeasure_measure_of_sigma)
+  show "X \<in> Pow A" using `X \<subseteq> A` by auto
   show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
-
-  show "positive (Pow A) ?M"
+  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)
 
-  show "countably_additive (Pow A) ?M"
-  proof (unfold countably_additive_def, safe)
-      fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F"
-      show "(\<Sum>i. ?M (F i)) = ?M (\<Union>i. F i)"
-      proof cases
-        assume "\<forall>i. finite (F i)"
-        then have finite_F: "\<And>i. finite (F i)" by auto
-        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
+  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)) ----> ?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 `incseq F` 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 "\<And>i. i \<le> f i" "\<And>i. F i \<noteq> F (f i)" by metis
+      moreover then have "\<And>i. F i \<subseteq> F (f i)" using `incseq F` by (auto simp: incseq_def)
+      ultimately have *: "\<And>i. F i \<subset> F (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 fin_eq: "finite (\<Union>i. F i) \<longleftrightarrow> finite {i. F i \<noteq> {}}"
-        proof
-          assume "finite (\<Union>i. F i)"
-          show "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
-        next
-          assume "finite {i. F i \<noteq> {}}"
-          with finite_F have "finite (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
-            by auto
-          also have "(\<Union>i\<in>{i. F i \<noteq> {}}. F i) = (\<Union>i. F i)"
-            by auto
-          finally show "finite (\<Union>i. F i)" .
-        qed
-        
-        show ?thesis
-        proof cases
-          assume *: "finite (\<Union>i. F i)"
-          with finite_F have "finite {i. ?M (F i) \<noteq> 0} "
-            by (simp add: fin_eq)
-          then have "(\<Sum>i. ?M (F i)) = (\<Sum>i | ?M (F i) \<noteq> 0. ?M (F i))"
-            by (rule suminf_finite) auto
-          also have "\<dots> = ereal (\<Sum>i | F i \<noteq> {}. card (F i))"
-            using finite_F by simp
-          also have "\<dots> = ereal (card (\<Union>i \<in> {i. F i \<noteq> {}}. F i))"
-            using * finite_F disj
-            by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def fin_eq)
-          also have "\<dots> = ?M (\<Union>i. F i)"
-            using * by (auto intro!: arg_cong[where f=card])
-          finally show ?thesis .
-        next
-          assume inf: "infinite (\<Union>i. F i)"
-          { fix i
-            have "\<exists>N. i \<le> (\<Sum>i<N. card (F i))"
-            proof (induct i)
-              case (Suc j)
-              from Suc obtain N where N: "j \<le> (\<Sum>i<N. card (F i))" by auto
-              have "infinite ({i. F i \<noteq> {}} - {..< N})"
-                using inf by (auto simp: fin_eq)
-              then have "{i. F i \<noteq> {}} - {..< N} \<noteq> {}"
-                by (metis finite.emptyI)
-              then obtain i where i: "F i \<noteq> {}" "N \<le> i"
-                by (auto simp: not_less[symmetric])
+      have "incseq (\<lambda>i. ?M (F i))"
+        using `incseq F` unfolding incseq_def by (auto simp: card_mono dest: finite_subset)
+      then have "(\<lambda>i. ?M (F i)) ----> (SUP n. ?M (F n))"
+        by (rule LIMSEQ_ereal_SUPR)
 
-              note N
-              also have "(\<Sum>i<N. card (F i)) \<le> (\<Sum>i<i. card (F i))"
-                by (rule setsum_mono2) (auto simp: i)
-              also have "\<dots> < (\<Sum>i<i. card (F i)) + card (F i)"
-                using finite_F `F i \<noteq> {}` by (simp add: card_gt_0_iff)
-              finally have "j < (\<Sum>i<Suc i. card (F i))"
-                by simp
-              then show ?case unfolding Suc_le_eq by blast
-            qed simp }
-          with finite_F inf show ?thesis
-            by (auto simp del: real_of_nat_setsum intro!: SUP_PInfty
-                     simp add: suminf_ereal_eq_SUPR real_of_nat_setsum[symmetric])
-        qed
-      next
-        assume "\<not> (\<forall>i. finite (F i))"
-        then obtain j where j: "infinite (F j)" by auto
-        then have "infinite (\<Union>i. F i)"
-          using finite_subset[of "F j" "\<Union>i. F i"] by auto
-        moreover have "\<And>i. 0 \<le> ?M (F i)" by auto
-        ultimately show ?thesis
-          using suminf_PInfty[of "\<lambda>i. ?M (F i)" j] j by auto
+      moreover have "(SUP n. ?M (F n)) = \<infinity>"
+      proof (rule SUP_PInfty)
+        fix n :: nat show "\<exists>k::nat\<in>UNIV. ereal 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 `F k \<subset> F (f k)` by (simp add: psubset_card_mono)
+          moreover have "finite (F (f k)) \<Longrightarrow> finite (F k)"
+            using `k \<le> f k` `incseq F` by (auto simp: incseq_def dest: finite_subset)
+          ultimately show ?case
+            by (auto intro!: exI[of _ "f k"])
+        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
-  show "X \<in> Pow A" using `X \<subseteq> A` by simp
 qed
 
 lemma emeasure_count_space_finite[simp]:
--- a/src/HOL/Probability/Probability_Measure.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -219,26 +219,26 @@
   with AE_in_set_eq_1 assms show ?thesis by simp
 qed
 
-lemma (in prob_space) prob_space_increasing: "increasing M prob"
+lemma (in finite_measure) prob_space_increasing: "increasing M (measure M)"
   by (auto intro!: finite_measure_mono simp: increasing_def)
 
-lemma (in prob_space) prob_zero_union:
-  assumes "s \<in> events" "t \<in> events" "prob t = 0"
-  shows "prob (s \<union> t) = prob s"
+lemma (in finite_measure) prob_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 "prob (s \<union> t) \<le> prob s"
+  have "measure M (s \<union> t) \<le> measure M s"
     using finite_measure_subadditive[of s t] assms by auto
-  moreover have "prob (s \<union> t) \<ge> prob s"
+  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 prob_space) prob_eq_compl:
-  assumes "s \<in> events" "t \<in> events"
-  assumes "prob (space M - s) = prob (space M - t)"
-  shows "prob s = prob t"
-  using assms prob_compl by auto
+lemma (in finite_measure) prob_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 prob_space) prob_one_inter:
   assumes events:"s \<in> events" "t \<in> events"
@@ -253,26 +253,26 @@
     using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
 qed
 
-lemma (in prob_space) prob_eq_bigunion_image:
-  assumes "range f \<subseteq> events" "range g \<subseteq> events"
+lemma (in finite_measure) prob_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. prob (f n) = prob (g n)"
-  shows "(prob (\<Union> i. f i) = prob (\<Union> i. g i))"
+  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. prob (f i)) sums (prob (\<Union> i. f i))"
+  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. prob (g i)) sums (prob (\<Union> i. g i))"
+  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 prob_space) prob_countably_zero:
-  assumes "range c \<subseteq> events"
-  assumes "\<And> i. prob (c i) = 0"
-  shows "prob (\<Union> i :: nat. c i) = 0"
+lemma (in finite_measure) prob_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 "prob (\<Union> i :: nat. c i) \<le> 0"
+  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 add: measure_nonneg)
 
@@ -316,7 +316,7 @@
 
 lemma (in prob_space) expectation_less:
   assumes [simp]: "integrable M X"
-  assumes gt: "\<forall>x\<in>space M. X x < b"
+  assumes gt: "AE x in M. X x < b"
   shows "expectation X < b"
 proof -
   have "expectation X < expectation (\<lambda>x. b)"
@@ -327,7 +327,7 @@
 
 lemma (in prob_space) expectation_greater:
   assumes [simp]: "integrable M X"
-  assumes gt: "\<forall>x\<in>space M. a < X x"
+  assumes gt: "AE x in M. a < X x"
   shows "a < expectation X"
 proof -
   have "expectation (\<lambda>x. a) < expectation X"
@@ -338,13 +338,13 @@
 
 lemma (in prob_space) jensens_inequality:
   fixes a b :: real
-  assumes X: "integrable M X" "X ` space M \<subseteq> I"
+  assumes X: "integrable M X" "AE x in M. X x \<in> I"
   assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV"
   assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q"
   shows "q (expectation X) \<le> expectation (\<lambda>x. q (X x))"
 proof -
   let ?F = "\<lambda>x. Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I))"
-  from not_empty X(2) have "I \<noteq> {}" by auto
+  from X(2) AE_False have "I \<noteq> {}" by auto
 
   from I have "open I" by auto
 
@@ -376,9 +376,12 @@
       using prob_space by (simp add: X)
     also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
       using `x \<in> I` `open I` X(2)
-      by (intro integral_mono integral_add integral_cmult integral_diff
-                lebesgue_integral_const X q convex_le_Inf_differential)
-         (auto simp: interior_open)
+      apply (intro integral_mono_AE integral_add integral_cmult integral_diff
+                lebesgue_integral_const X q)
+      apply (elim eventually_elim1)
+      apply (intro convex_le_Inf_differential)
+      apply (auto simp: interior_open q)
+      done
     finally show "k \<le> expectation (\<lambda>w. q (X w))" using x by auto
   qed
   finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" .
@@ -407,7 +410,7 @@
 sublocale pair_prob_space \<subseteq> P: prob_space "M1 \<Otimes>\<^isub>M M2"
 proof
   show "emeasure (M1 \<Otimes>\<^isub>M M2) (space (M1 \<Otimes>\<^isub>M M2)) = 1"
-    by (simp add: emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure)
+    by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure)
 qed
 
 locale product_prob_space = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
@@ -503,6 +506,16 @@
   shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
   using subdensity[OF T, of M X "\<lambda>x. ereal (f x)" Y "\<lambda>x. ereal (g x)"] assms by auto
 
+lemma distributed_emeasure:
+  "distributed M N X f \<Longrightarrow> A \<in> sets N \<Longrightarrow> emeasure M (X -` A \<inter> space M) = (\<integral>\<^isup>+x. f x * indicator A x \<partial>N)"
+  by (auto simp: distributed_measurable distributed_AE distributed_borel_measurable
+                 distributed_distr_eq_density[symmetric] emeasure_density[symmetric] emeasure_distr)
+
+lemma distributed_positive_integral:
+  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^isup>+x. f x * g x \<partial>N) = (\<integral>\<^isup>+x. g (X x) \<partial>M)"
+  by (auto simp: distributed_measurable distributed_AE distributed_borel_measurable
+                 distributed_distr_eq_density[symmetric] positive_integral_density[symmetric] positive_integral_distr)
+
 lemma distributed_integral:
   "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>x. f x * g x \<partial>N) = (\<integral>x. g (X x) \<partial>M)"
   by (auto simp: distributed_real_measurable distributed_real_AE distributed_measurable
@@ -523,54 +536,116 @@
   finally show ?thesis .
 qed
 
-lemma distributed_marginal_eq_joint:
-  assumes T: "sigma_finite_measure T"
-  assumes S: "sigma_finite_measure S"
+lemma (in prob_space) distributed_unique:
   assumes Px: "distributed M S X Px"
-  assumes Py: "distributed M T Y Py"
-  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
-  shows "AE y in T. Py y = (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S)"
-proof (rule sigma_finite_measure.density_unique[OF T])
-  interpret ST: pair_sigma_finite S T using S T unfolding pair_sigma_finite_def by simp
-  show "Py \<in> borel_measurable T" "AE y in T. 0 \<le> Py y"
-    "(\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S) \<in> borel_measurable T" "AE y in T. 0 \<le> \<integral>\<^isup>+ x. Pxy (x, y) \<partial>S"
-    using Pxy[THEN distributed_borel_measurable]
-    by (auto intro!: Py[THEN distributed_borel_measurable] Py[THEN distributed_AE]
-                     ST.positive_integral_snd_measurable' positive_integral_positive)
+  assumes Py: "distributed M S X Py"
+  shows "AE x in S. Px x = Py x"
+proof -
+  interpret X: prob_space "distr M S X"
+    using distributed_measurable[OF Px] by (rule prob_space_distr)
+  have "sigma_finite_measure (distr M S X)" ..
+  with sigma_finite_density_unique[of Px S Py ] Px Py
+  show ?thesis
+    by (auto simp: distributed_def)
+qed
+
+lemma (in prob_space) distributed_jointI:
+  assumes "sigma_finite_measure S" "sigma_finite_measure T"
+  assumes X[simp]: "X \<in> measurable M S" and Y[simp]: "Y \<in> measurable M T"
+  assumes f[simp]: "f \<in> borel_measurable (S \<Otimes>\<^isub>M T)" "AE x in S \<Otimes>\<^isub>M T. 0 \<le> f x"
+  assumes eq: "\<And>A B. A \<in> sets S \<Longrightarrow> B \<in> sets T \<Longrightarrow> 
+    emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B} = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. f (x, y) * indicator B y \<partial>T) * indicator A x \<partial>S)"
+  shows "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) f"
+  unfolding distributed_def
+proof safe
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T by default
 
-  show "density T Py = density T (\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S)"
-  proof (rule measure_eqI)
-    fix A assume A: "A \<in> sets (density T Py)"
-    have *: "\<And>x y. x \<in> space S \<Longrightarrow> indicator (space S \<times> A) (x, y) = indicator A y"
-      by (auto simp: indicator_def)
-    have "emeasure (density T Py) A = emeasure (distr M T Y) A"
-      unfolding Py[THEN distributed_distr_eq_density] ..
-    also have "\<dots> = emeasure (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) (space S \<times> A)"
-      using A Px Py Pxy
-      by (subst (1 2) emeasure_distr)
-         (auto dest: measurable_space distributed_measurable intro!: arg_cong[where f="emeasure M"])
-    also have "\<dots> = emeasure (density (S \<Otimes>\<^isub>M T) Pxy) (space S \<times> A)"
-      unfolding Pxy[THEN distributed_distr_eq_density] ..
-    also have "\<dots> = (\<integral>\<^isup>+ x. Pxy x * indicator (space S \<times> A) x \<partial>(S \<Otimes>\<^isub>M T))"
-      using A Pxy by (simp add: emeasure_density distributed_borel_measurable)
-    also have "\<dots> = (\<integral>\<^isup>+y. \<integral>\<^isup>+x. Pxy (x, y) * indicator (space S \<times> A) (x, y) \<partial>S \<partial>T)"
-      using A Pxy
-      by (subst ST.positive_integral_snd_measurable) (simp_all add: emeasure_density distributed_borel_measurable)
-    also have "\<dots> = (\<integral>\<^isup>+y. (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S) * indicator A y \<partial>T)"
-      using measurable_comp[OF measurable_Pair1[OF measurable_identity] distributed_borel_measurable[OF Pxy]]
-      using distributed_borel_measurable[OF Pxy] distributed_AE[OF Pxy, THEN ST.AE_pair]
-      by (subst (asm) ST.AE_commute) (auto intro!: positive_integral_cong_AE positive_integral_multc cong: positive_integral_cong simp: * comp_def)
-    also have "\<dots> = emeasure (density T (\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S)) A"
-      using A by (intro emeasure_density[symmetric])  (auto intro!: ST.positive_integral_snd_measurable' Pxy[THEN distributed_borel_measurable])
-    finally show "emeasure (density T Py) A = emeasure (density T (\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S)) A" .
-  qed simp
+  from ST.sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('b \<times> 'c) set" .. note F = this
+  let ?E = "{a \<times> b |a b. a \<in> sets S \<and> b \<in> sets T}"
+  let ?P = "S \<Otimes>\<^isub>M T"
+  show "distr M ?P (\<lambda>x. (X x, Y x)) = density ?P f" (is "?L = ?R")
+  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of S T]])
+    show "?E \<subseteq> Pow (space ?P)"
+      using space_closed[of S] space_closed[of T] by (auto simp: space_pair_measure)
+    show "sets ?L = sigma_sets (space ?P) ?E"
+      by (simp add: sets_pair_measure space_pair_measure)
+    then show "sets ?R = sigma_sets (space ?P) ?E"
+      by simp
+  next
+    interpret L: prob_space ?L
+      by (rule prob_space_distr) (auto intro!: measurable_Pair)
+    show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?L (F i) \<noteq> \<infinity>"
+      using F by (auto simp: space_pair_measure)
+  next
+    fix E assume "E \<in> ?E"
+    then obtain A B where E[simp]: "E = A \<times> B" and A[simp]: "A \<in> sets S" and B[simp]: "B \<in> sets T" by auto
+    have "emeasure ?L E = emeasure M {x \<in> space M. X x \<in> A \<and> Y x \<in> B}"
+      by (auto intro!: arg_cong[where f="emeasure M"] simp add: emeasure_distr measurable_Pair)
+    also have "\<dots> = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. (f (x, y) * indicator B y) * indicator A x \<partial>T) \<partial>S)"
+      by (auto simp add: eq measurable_Pair measurable_compose[OF _ f(1)] positive_integral_multc
+               intro!: positive_integral_cong)
+    also have "\<dots> = emeasure ?R E"
+      by (auto simp add: emeasure_density ST.positive_integral_fst_measurable(2)[symmetric]
+               intro!: positive_integral_cong split: split_indicator)
+    finally show "emeasure ?L E = emeasure ?R E" .
+  qed
+qed (auto intro!: measurable_Pair)
+
+lemma (in prob_space) distributed_swap:
+  assumes "sigma_finite_measure S" "sigma_finite_measure T"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  shows "distributed M (T \<Otimes>\<^isub>M S) (\<lambda>x. (Y x, X x)) (\<lambda>(x, y). Pxy (y, x))"
+proof -
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T by default
+  interpret TS: pair_sigma_finite T S by default
+
+  note measurable_Pxy = measurable_compose[OF _ distributed_borel_measurable[OF Pxy]]
+  show ?thesis 
+    apply (subst TS.distr_pair_swap)
+    unfolding distributed_def
+  proof safe
+    let ?D = "distr (S \<Otimes>\<^isub>M T) (T \<Otimes>\<^isub>M S) (\<lambda>(x, y). (y, x))"
+    show 1: "(\<lambda>(x, y). Pxy (y, x)) \<in> borel_measurable ?D"
+      by (auto simp: measurable_split_conv intro!: measurable_Pair measurable_Pxy)
+    with Pxy
+    show "AE x in distr (S \<Otimes>\<^isub>M T) (T \<Otimes>\<^isub>M S) (\<lambda>(x, y). (y, x)). 0 \<le> (case x of (x, y) \<Rightarrow> Pxy (y, x))"
+      by (subst AE_distr_iff)
+         (auto dest!: distributed_AE
+               simp: measurable_split_conv split_beta
+               intro!: measurable_Pair borel_measurable_ereal_le)
+    show 2: "random_variable (distr (S \<Otimes>\<^isub>M T) (T \<Otimes>\<^isub>M S) (\<lambda>(x, y). (y, x))) (\<lambda>x. (Y x, X x))"
+      using measurable_compose[OF distributed_measurable[OF Pxy] measurable_fst]
+      using measurable_compose[OF distributed_measurable[OF Pxy] measurable_snd]
+      by (auto intro!: measurable_Pair)
+    { fix A assume A: "A \<in> sets (T \<Otimes>\<^isub>M S)"
+      let ?B = "(\<lambda>(x, y). (y, x)) -` A \<inter> space (S \<Otimes>\<^isub>M T)"
+      from sets_into_space[OF A]
+      have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) =
+        emeasure M ((\<lambda>x. (X x, Y x)) -` ?B \<inter> space M)"
+        by (auto intro!: arg_cong2[where f=emeasure] simp: space_pair_measure)
+      also have "\<dots> = (\<integral>\<^isup>+ x. Pxy x * indicator ?B x \<partial>(S \<Otimes>\<^isub>M T))"
+        using Pxy A by (intro distributed_emeasure measurable_sets) (auto simp: measurable_split_conv measurable_Pair)
+      finally have "emeasure M ((\<lambda>x. (Y x, X x)) -` A \<inter> space M) =
+        (\<integral>\<^isup>+ x. Pxy x * indicator A (snd x, fst x) \<partial>(S \<Otimes>\<^isub>M T))"
+        by (auto intro!: positive_integral_cong split: split_indicator) }
+    note * = this
+    show "distr M ?D (\<lambda>x. (Y x, X x)) = density ?D (\<lambda>(x, y). Pxy (y, x))"
+      apply (intro measure_eqI)
+      apply (simp_all add: emeasure_distr[OF 2] emeasure_density[OF 1])
+      apply (subst positive_integral_distr)
+      apply (auto intro!: measurable_pair measurable_Pxy * simp: comp_def split_beta)
+      done
+  qed
 qed
 
 lemma (in prob_space) distr_marginal1:
-  fixes Pxy :: "('b \<times> 'c) \<Rightarrow> real"
   assumes "sigma_finite_measure S" "sigma_finite_measure T"
   assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
-  defines "Px \<equiv> \<lambda>x. real (\<integral>\<^isup>+z. Pxy (x, z) \<partial>T)"
+  defines "Px \<equiv> \<lambda>x. (\<integral>\<^isup>+z. Pxy (x, z) \<partial>T)"
   shows "distributed M S X Px"
   unfolding distributed_def
 proof safe
@@ -585,18 +660,15 @@
   from XY have Y: "Y \<in> measurable M T"
     unfolding measurable_pair_iff by (simp add: comp_def)
 
-  from Pxy show borel: "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S"
-    by (auto intro!: ST.positive_integral_fst_measurable borel_measurable_real_of_ereal dest!: distributed_real_measurable simp: Px_def)
+  from Pxy show borel: "Px \<in> borel_measurable S"
+    by (auto intro!: ST.positive_integral_fst_measurable dest!: distributed_borel_measurable simp: Px_def)
 
   interpret Pxy: prob_space "distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
     using XY by (rule prob_space_distr)
-  have "(\<integral>\<^isup>+ x. max 0 (ereal (- Pxy x)) \<partial>(S \<Otimes>\<^isub>M T)) = (\<integral>\<^isup>+ x. 0 \<partial>(S \<Otimes>\<^isub>M T))"
+  have "(\<integral>\<^isup>+ x. max 0 (- Pxy x) \<partial>(S \<Otimes>\<^isub>M T)) = (\<integral>\<^isup>+ x. 0 \<partial>(S \<Otimes>\<^isub>M T))"
     using Pxy
-    by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_real_measurable distributed_real_AE)
-  then have Pxy_integrable: "integrable (S \<Otimes>\<^isub>M T) Pxy"
-    using Pxy Pxy.emeasure_space_1
-    by (simp add: integrable_def emeasure_density positive_integral_max_0 distributed_def borel_measurable_ereal_iff cong: positive_integral_cong)
-    
+    by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_borel_measurable distributed_AE)
+
   show "distr M S X = density S Px"
   proof (rule measure_eqI)
     fix A assume A: "A \<in> sets (distr M S X)"
@@ -605,31 +677,100 @@
                intro!: arg_cong[where f="emeasure M"] dest: measurable_space)
     also have "\<dots> = emeasure (density (S \<Otimes>\<^isub>M T) Pxy) (A \<times> space T)"
       using Pxy by (simp add: distributed_def)
-    also have "\<dots> = \<integral>\<^isup>+ x. \<integral>\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S"
+    also have "\<dots> = \<integral>\<^isup>+ x. \<integral>\<^isup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S"
       using A borel Pxy
       by (simp add: emeasure_density ST.positive_integral_fst_measurable(2)[symmetric] distributed_def)
-    also have "\<dots> = \<integral>\<^isup>+ x. ereal (Px x) * indicator A x \<partial>S"
+    also have "\<dots> = \<integral>\<^isup>+ x. Px x * indicator A x \<partial>S"
       apply (rule positive_integral_cong_AE)
-      using Pxy[THEN distributed_real_AE, THEN ST.AE_pair] ST.integrable_fst_measurable(1)[OF Pxy_integrable] AE_space
+      using Pxy[THEN distributed_AE, THEN ST.AE_pair] AE_space
     proof eventually_elim
-      fix x assume "x \<in> space S" "AE y in T. 0 \<le> Pxy (x, y)" and i: "integrable T (\<lambda>y. Pxy (x, y))"
+      fix x assume "x \<in> space S" "AE y in T. 0 \<le> Pxy (x, y)"
       moreover have eq: "\<And>y. y \<in> space T \<Longrightarrow> indicator (A \<times> space T) (x, y) = indicator A x"
         by (auto simp: indicator_def)
-      ultimately have "(\<integral>\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \<times> space T) (x, y) \<partial>T) =
-          (\<integral>\<^isup>+ y. ereal (Pxy (x, y)) \<partial>T) * indicator A x"
-        using Pxy[THEN distributed_real_measurable] by (simp add: eq positive_integral_multc measurable_Pair2 cong: positive_integral_cong)
-      also have "(\<integral>\<^isup>+ y. ereal (Pxy (x, y)) \<partial>T) = Px x"
-        using i by (simp add: Px_def ereal_real integrable_def positive_integral_positive)
-      finally show "(\<integral>\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \<times> space T) (x, y) \<partial>T) = ereal (Px x) * indicator A x" .
+      ultimately have "(\<integral>\<^isup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = (\<integral>\<^isup>+ y. Pxy (x, y) \<partial>T) * indicator A x"
+        using Pxy[THEN distributed_borel_measurable] by (simp add: eq positive_integral_multc measurable_Pair2 cong: positive_integral_cong)
+      also have "(\<integral>\<^isup>+ y. Pxy (x, y) \<partial>T) = Px x"
+        by (simp add: Px_def ereal_real positive_integral_positive)
+      finally show "(\<integral>\<^isup>+ y. Pxy (x, y) * indicator (A \<times> space T) (x, y) \<partial>T) = Px x * indicator A x" .
     qed
     finally show "emeasure (distr M S X) A = emeasure (density S Px) A"
       using A borel Pxy by (simp add: emeasure_density)
   qed simp
   
-  show "AE x in S. 0 \<le> ereal (Px x)"
+  show "AE x in S. 0 \<le> Px x"
     by (simp add: Px_def positive_integral_positive real_of_ereal_pos)
 qed
 
+lemma (in prob_space) distr_marginal2:
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  shows "distributed M T Y (\<lambda>y. (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S))"
+  using distr_marginal1[OF T S distributed_swap[OF S T]] Pxy by simp
+
+lemma (in prob_space) distributed_marginal_eq_joint1:
+  assumes T: "sigma_finite_measure T"
+  assumes S: "sigma_finite_measure S"
+  assumes Px: "distributed M S X Px"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  shows "AE x in S. Px x = (\<integral>\<^isup>+y. Pxy (x, y) \<partial>T)"
+  using Px distr_marginal1[OF S T Pxy] by (rule distributed_unique)
+
+lemma (in prob_space) distributed_marginal_eq_joint2:
+  assumes T: "sigma_finite_measure T"
+  assumes S: "sigma_finite_measure S"
+  assumes Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  shows "AE y in T. Py y = (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S)"
+  using Py distr_marginal2[OF S T Pxy] by (rule distributed_unique)
+
+lemma (in prob_space) distributed_joint_indep':
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes X: "distributed M S X Px" and Y: "distributed M T Y Py"
+  assumes indep: "distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+  shows "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) (\<lambda>(x, y). Px x * Py y)"
+  unfolding distributed_def
+proof safe
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T by default
+
+  interpret X: prob_space "density S Px"
+    unfolding distributed_distr_eq_density[OF X, symmetric]
+    using distributed_measurable[OF X]
+    by (rule prob_space_distr)
+  have sf_X: "sigma_finite_measure (density S Px)" ..
+
+  interpret Y: prob_space "density T Py"
+    unfolding distributed_distr_eq_density[OF Y, symmetric]
+    using distributed_measurable[OF Y]
+    by (rule prob_space_distr)
+  have sf_Y: "sigma_finite_measure (density T Py)" ..
+
+  show "distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = density (S \<Otimes>\<^isub>M T) (\<lambda>(x, y). Px x * Py y)"
+    unfolding indep[symmetric] distributed_distr_eq_density[OF X] distributed_distr_eq_density[OF Y]
+    using distributed_borel_measurable[OF X] distributed_AE[OF X]
+    using distributed_borel_measurable[OF Y] distributed_AE[OF Y]
+    by (rule pair_measure_density[OF _ _ _ _ S T sf_X sf_Y])
+
+  show "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+    using distributed_measurable[OF X] distributed_measurable[OF Y]
+    by (auto intro: measurable_Pair)
+
+  show Pxy: "(\<lambda>(x, y). Px x * Py y) \<in> borel_measurable (S \<Otimes>\<^isub>M T)"
+    by (auto simp: split_beta' 
+             intro!: measurable_compose[OF _ distributed_borel_measurable[OF X]]
+                     measurable_compose[OF _ distributed_borel_measurable[OF Y]])
+
+  show "AE x in S \<Otimes>\<^isub>M T. 0 \<le> (case x of (x, y) \<Rightarrow> Px x * Py y)"
+    apply (intro ST.AE_pair_measure borel_measurable_ereal_le Pxy borel_measurable_const)
+    using distributed_AE[OF X]
+    apply eventually_elim
+    using distributed_AE[OF Y]
+    apply eventually_elim
+    apply auto
+    done
+qed
+
 definition
   "simple_distributed M X f \<longleftrightarrow> distributed M (count_space (X`space M)) X (\<lambda>x. ereal (f x)) \<and>
     finite (X`space M)"
@@ -792,19 +933,21 @@
   note Px = simple_distributedI[OF Px refl]
   have *: "\<And>f A. setsum (\<lambda>x. max 0 (ereal (f x))) A = ereal (setsum (\<lambda>x. max 0 (f x)) A)"
     by (simp add: setsum_ereal[symmetric] zero_ereal_def)
-  from distributed_marginal_eq_joint[OF sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite
-    simple_distributed[OF Px] simple_distributed[OF Py] simple_distributed_joint[OF Pxy],
+  from distributed_marginal_eq_joint2[OF
+    sigma_finite_measure_count_space_finite
+    sigma_finite_measure_count_space_finite
+    simple_distributed[OF Py] simple_distributed_joint[OF Pxy],
     OF Py[THEN simple_distributed_finite] Px[THEN simple_distributed_finite]]
-    y Px[THEN simple_distributed_finite] Py[THEN simple_distributed_finite]
+    y
+    Px[THEN simple_distributed_finite]
+    Py[THEN simple_distributed_finite]
     Pxy[THEN simple_distributed, THEN distributed_real_AE]
   show ?thesis
     unfolding AE_count_space
-    apply (elim ballE[where x=y])
     apply (auto simp add: positive_integral_count_space_finite * intro!: setsum_cong split: split_max)
     done
 qed
 
-
 lemma prob_space_uniform_measure:
   assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>"
   shows "prob_space (uniform_measure M A)"
--- a/src/HOL/Probability/Radon_Nikodym.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -776,15 +776,16 @@
   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\<^isup>P M f \<noteq> \<infinity>"
-  shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. g x * indicator A x \<partial>M))
-    \<longleftrightarrow> (AE x in M. f x = g x)"
-    (is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _")
+  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"
-  then show "?P f A = ?P g A"
-    by (auto intro: positive_integral_cong_AE)
+  with borel show "density M f = density M g"
+    by (auto intro: density_cong)
 next
-  assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+  let ?P = "\<lambda>f A. \<integral>\<^isup>+ 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 top] fin
   have g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)
   { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
@@ -844,7 +845,7 @@
     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!: borel_measurable_ereal_times f Int simp: *)
+       (auto intro!: borel_measurable_ereal_times f measure_eqI Int simp: emeasure_density * subset_eq)
   moreover have "AE x in M. ?f Q0 x = ?f' Q0 x"
   proof (rule AE_I')
     { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
@@ -933,7 +934,42 @@
   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 (in sigma_finite_measure) sigma_finite_iff_density_finite:
+lemma sigma_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: "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>\<^isup>+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 pos cover(1) pos
+      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" "AE x in M. 0 \<le> f x"
   shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
     (is "sigma_finite_measure ?N \<longleftrightarrow> _")
@@ -1019,6 +1055,13 @@
   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>)"
+  apply (subst density_max_0)
+  apply (subst sigma_finite_iff_density_finite')
+  apply (auto simp: max_def intro!: measurable_If)
+  done
+
 section "Radon-Nikodym derivative"
 
 definition
@@ -1069,15 +1112,17 @@
   assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
   and eq: "density M f = N"
   shows "AE x in M. f x = RN_deriv M N x"
-proof (rule density_unique)
-  have ac: "absolutely_continuous M N"
-    using f(1) unfolding eq[symmetric] by (rule absolutely_continuousI_density)
-  have eq2: "sets N = sets M"
-    unfolding eq[symmetric] by simp
-  show "RN_deriv M N \<in> borel_measurable M" "AE x in M. 0 \<le> RN_deriv M N x"
-    "density M f = density M (RN_deriv M N)"
-    using RN_deriv[OF ac eq2] eq by auto
-qed fact+
+  unfolding eq[symmetric]
+  by (intro density_unique_iff[THEN iffD1] f borel_measurable_RN_deriv_density
+            RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric])
+
+lemma RN_deriv_unique_sigma_finite:
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  and eq: "density M f = N" and fin: "sigma_finite_measure N"
+  shows "AE x in M. f x = RN_deriv M N x"
+  using fin unfolding eq[symmetric]
+  by (intro sigma_finite_density_unique[THEN iffD1] f borel_measurable_RN_deriv_density
+            RN_deriv_density_nonneg[THEN AE_I2] density_RN_deriv_density[symmetric])
 
 lemma (in sigma_finite_measure) RN_deriv_distr:
   fixes T :: "'a \<Rightarrow> 'b"
--- a/src/HOL/Probability/Sigma_Algebra.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -1292,11 +1292,11 @@
 
 lemma measurable_If_set:
   assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
-  assumes P: "A \<in> sets 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" using `A \<in> sets M` sets_into_space by auto
-  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<in> sets M` by auto
+  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 `A \<inter> space M \<in> sets M` by auto
 qed
 
 lemma measurable_ident[intro, simp]: "id \<in> measurable M M"
@@ -1310,6 +1310,10 @@
   apply force+
   done
 
+lemma measurable_compose:
+  "f \<in> measurable M N \<Longrightarrow> g \<in> measurable N L \<Longrightarrow> (\<lambda>x. g (f x)) \<in> measurable M L"
+  using measurable_comp[of f M N g L] by (simp add: comp_def)
+
 lemma measurable_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"
@@ -2037,4 +2041,25 @@
     using assms by (auto simp: dynkin_def)
 qed
 
+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 default (auto dest: sets_into_space intro!: space compl union)
+  have "sigma_sets \<Omega> G = ?D"
+    by (rule dynkin_lemma) (auto simp: basic `Int_stable G`)
+  with A show ?thesis by auto
+qed
+
 end
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Oct 10 15:16:44 2012 +0200
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Oct 10 15:17:18 2012 +0200
@@ -145,7 +145,7 @@
   "msgs = keys \<times> {ms. set ms \<subseteq> messages \<and> length ms = n}"
 
 definition P :: "('key \<times> 'message list) \<Rightarrow> real" where
-  "P = (\<lambda>(k, ms). K k * (\<Prod>i<length ms. M (ms ! i)))"
+  "P = (\<lambda>(k, ms). K k * (\<Prod>i<n. M (ms ! i)))"
 
 end
 
@@ -191,7 +191,10 @@
   "OB = (\<lambda>(k, ms). map (observe k) ms)"
 
 definition t :: "'observation list \<Rightarrow> 'observation \<Rightarrow> nat" where
-  "t seq obs = length (filter (op = obs) seq)"
+  t_def2: "t seq obs = card { i. i < length seq \<and> seq ! i = obs}"
+
+lemma t_def: "t seq obs = length (filter (op = obs) seq)"
+  unfolding t_def2 length_filter_conv_card by (subst eq_commute) simp
 
 lemma card_T_bound: "card ((t\<circ>OB)`msgs) \<le> (n+1)^card observations"
 proof -
@@ -306,7 +309,6 @@
      log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
   apply (subst conditional_entropy_eq[OF
     simple_distributedI[OF simple_function_finite refl]
-    simple_function_finite
     simple_distributedI[OF simple_function_finite refl]])
   unfolding space_point_measure
 proof -
@@ -325,7 +327,7 @@
     -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))" .
 qed
 
-lemma ce_OB_eq_ce_t: "\<H>(fst | OB) = \<H>(fst | t\<circ>OB)"
+lemma ce_OB_eq_ce_t: "\<I>(fst ; OB) = \<I>(fst ; t\<circ>OB)"
 proof -
   txt {* Lemma 2 *}
 
@@ -434,7 +436,6 @@
   note P_t_sum_P_O = this
 
   txt {* Lemma 3 *}
-  txt {* Lemma 3 *}
   have "\<H>(fst | OB) = -(\<Sum>obs\<in>OB`msgs. \<P>(OB) {obs} * ?Ht (t obs))"
     unfolding conditional_entropy_eq_ce_with_hypothesis using * by simp
   also have "\<dots> = -(\<Sum>obs\<in>t`OB`msgs. \<P>(t\<circ>OB) {obs} * ?Ht obs)"
@@ -452,16 +453,15 @@
   also have "\<dots> = \<H>(fst | t\<circ>OB)"
     unfolding conditional_entropy_eq_ce_with_hypothesis
     by (simp add: comp_def image_image[symmetric])
-  finally show ?thesis .
+  finally show ?thesis
+    by (subst (1 2) mutual_information_eq_entropy_conditional_entropy) simp_all
 qed
 
 theorem "\<I>(fst ; OB) \<le> real (card observations) * log b (real n + 1)"
 proof -
-  from simple_function_finite simple_function_finite
-  have "\<I>(fst ; OB) = \<H>(fst) - \<H>(fst | OB)"
-    by (rule mutual_information_eq_entropy_conditional_entropy)
-  also have "\<dots> = \<H>(fst) - \<H>(fst | t\<circ>OB)"
-    unfolding ce_OB_eq_ce_t ..
+  have "\<I>(fst ; OB) = \<H>(fst) - \<H>(fst | t\<circ>OB)"
+    unfolding ce_OB_eq_ce_t
+    by (rule mutual_information_eq_entropy_conditional_entropy simple_function_finite)+
   also have "\<dots> = \<H>(t\<circ>OB) - \<H>(t\<circ>OB | fst)"
     unfolding entropy_chain_rule[symmetric, OF simple_function_finite simple_function_finite] sign_simps
     by (subst entropy_commute) simp