Support product spaces on sigma finite measures.
authorhoelzl
Wed, 01 Dec 2010 19:20:30 +0100
changeset 40859 de0b30e6c2d2
parent 40854 f2c9ebbe04aa
child 40860 a6df324752da
Support product spaces on sigma finite measures. Introduce the almost everywhere quantifier. Introduce 'morphism' theorems for most constants. Prove uniqueness of measures on cut stable generators. Prove uniqueness of the Radon-Nikodym derivative. Removed misleading suffix from borel_space and lebesgue_space. Use product spaces to introduce Fubini on the Lebesgue integral. Combine Euclidean_Lebesgue and Lebesgue_Measure. Generalize theorems about mutual information and entropy to arbitrary probability spaces. Remove simproc mult_log as it does not work with interpretations. Introduce completion of measure spaces. Change type of sigma. Introduce dynkin systems.
src/HOL/Probability/Borel.thy
src/HOL/Probability/Borel_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Complete_Measure.thy
src/HOL/Probability/Euclidean_Lebesgue.thy
src/HOL/Probability/Information.thy
src/HOL/Probability/Lebesgue_Integration.thy
src/HOL/Probability/Lebesgue_Measure.thy
src/HOL/Probability/Measure.thy
src/HOL/Probability/Positive_Infinite_Real.thy
src/HOL/Probability/Probability.thy
src/HOL/Probability/Probability_Space.thy
src/HOL/Probability/Product_Measure.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/Probability/Sigma_Algebra.thy
src/HOL/Probability/ex/Dining_Cryptographers.thy
src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy
--- a/src/HOL/Probability/Borel.thy	Wed Dec 01 06:50:54 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1390 +0,0 @@
-(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *)
-
-header {*Borel spaces*}
-
-theory Borel
-  imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
-begin
-
-lemma LIMSEQ_max:
-  "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
-  by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
-
-section "Generic Borel spaces"
-
-definition "borel_space = sigma (UNIV::'a::topological_space set) open"
-abbreviation "borel_measurable M \<equiv> measurable M borel_space"
-
-interpretation borel_space: sigma_algebra borel_space
-  using sigma_algebra_sigma by (auto simp: borel_space_def)
-
-lemma in_borel_measurable:
-   "f \<in> borel_measurable M \<longleftrightarrow>
-    (\<forall>S \<in> sets (sigma UNIV open).
-      f -` S \<inter> space M \<in> sets M)"
-  by (auto simp add: measurable_def borel_space_def)
-
-lemma in_borel_measurable_borel_space:
-   "f \<in> borel_measurable M \<longleftrightarrow>
-    (\<forall>S \<in> sets borel_space.
-      f -` S \<inter> space M \<in> sets M)"
-  by (auto simp add: measurable_def borel_space_def)
-
-lemma space_borel_space[simp]: "space borel_space = UNIV"
-  unfolding borel_space_def by auto
-
-lemma borel_space_open[simp]:
-  assumes "open A" shows "A \<in> sets borel_space"
-proof -
-  have "A \<in> open" unfolding mem_def using assms .
-  thus ?thesis unfolding borel_space_def sigma_def by (auto intro!: sigma_sets.Basic)
-qed
-
-lemma borel_space_closed[simp]:
-  assumes "closed A" shows "A \<in> sets borel_space"
-proof -
-  have "space borel_space - (- A) \<in> sets borel_space"
-    using assms unfolding closed_def by (blast intro: borel_space_open)
-  thus ?thesis by simp
-qed
-
-lemma (in sigma_algebra) borel_measurable_vimage:
-  fixes f :: "'a \<Rightarrow> 'x::t2_space"
-  assumes borel: "f \<in> borel_measurable M"
-  shows "f -` {x} \<inter> space M \<in> sets M"
-proof (cases "x \<in> f ` space M")
-  case True then obtain y where "x = f y" by auto
-  from closed_sing[of "f y"]
-  have "{f y} \<in> sets borel_space" by (rule borel_space_closed)
-  with assms show ?thesis
-    unfolding in_borel_measurable_borel_space `x = f y` by auto
-next
-  case False hence "f -` {x} \<inter> space M = {}" by auto
-  thus ?thesis by auto
-qed
-
-lemma (in sigma_algebra) borel_measurableI:
-  fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
-  assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
-  shows "f \<in> borel_measurable M"
-  unfolding borel_space_def
-proof (rule measurable_sigma)
-  fix S :: "'x set" assume "S \<in> open" thus "f -` S \<inter> space M \<in> sets M"
-    using assms[of S] by (simp add: mem_def)
-qed simp_all
-
-lemma borel_space_singleton[simp, intro]:
-  fixes x :: "'a::t1_space"
-  shows "A \<in> sets borel_space \<Longrightarrow> insert x A \<in> sets borel_space"
-  proof (rule borel_space.insert_in_sets)
-    show "{x} \<in> sets borel_space"
-      using closed_sing[of x] by (rule borel_space_closed)
-  qed simp
-
-lemma (in sigma_algebra) borel_measurable_const[simp, intro]:
-  "(\<lambda>x. c) \<in> borel_measurable M"
-  by (auto intro!: measurable_const)
-
-lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]:
-  assumes A: "A \<in> sets M"
-  shows "indicator A \<in> borel_measurable M"
-  unfolding indicator_def_raw using A
-  by (auto intro!: measurable_If_set borel_measurable_const)
-
-lemma borel_measurable_translate:
-  assumes "A \<in> sets borel_space" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel_space"
-  shows "f -` A \<in> sets borel_space"
-proof -
-  have "A \<in> sigma_sets UNIV open" using assms
-    by (simp add: borel_space_def sigma_def)
-  thus ?thesis
-  proof (induct rule: sigma_sets.induct)
-    case (Basic a) thus ?case using trans[of a] by (simp add: mem_def)
-  next
-    case (Compl a)
-    moreover have "UNIV \<in> sets borel_space"
-      by (metis borel_space.top borel_space_def_raw mem_def space_sigma)
-    ultimately show ?case
-      by (auto simp: vimage_Diff borel_space.Diff)
-  qed (auto simp add: vimage_UN)
-qed
-
-lemma (in sigma_algebra) borel_measurable_restricted:
-  fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
-  shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
-    (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
-    (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
-proof -
-  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
-  have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
-    by (auto intro!: measurable_cong)
-  show ?thesis unfolding *
-    unfolding in_borel_measurable_borel_space
-  proof (simp, safe)
-    fix S :: "'x set" assume "S \<in> sets borel_space"
-      "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
-    then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
-    then have f: "?f -` S \<inter> A \<in> sets M"
-      using `A \<in> sets M` sets_into_space by fastsimp
-    show "?f -` S \<inter> space M \<in> sets M"
-    proof cases
-      assume "0 \<in> S"
-      then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
-        using `A \<in> sets M` sets_into_space by auto
-      then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
-    next
-      assume "0 \<notin> S"
-      then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
-        using `A \<in> sets M` sets_into_space
-        by (auto simp: indicator_def split: split_if_asm)
-      then show ?thesis using f by auto
-    qed
-  next
-    fix S :: "'x set" assume "S \<in> sets borel_space"
-      "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M"
-    then have f: "?f -` S \<inter> space M \<in> sets M" by auto
-    then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
-      using `A \<in> sets M` sets_into_space
-      apply (simp add: image_iff)
-      apply (rule bexI[OF _ f])
-      by auto
-  qed
-qed
-
-lemma (in sigma_algebra) borel_measurable_subalgebra:
-  assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)"
-  shows "f \<in> borel_measurable M"
-  using assms unfolding measurable_def by auto
-
-section "Borel spaces on euclidean spaces"
-
-lemma lessThan_borel[simp, intro]:
-  fixes a :: "'a\<Colon>ordered_euclidean_space"
-  shows "{..< a} \<in> sets borel_space"
-  by (blast intro: borel_space_open)
-
-lemma greaterThan_borel[simp, intro]:
-  fixes a :: "'a\<Colon>ordered_euclidean_space"
-  shows "{a <..} \<in> sets borel_space"
-  by (blast intro: borel_space_open)
-
-lemma greaterThanLessThan_borel[simp, intro]:
-  fixes a b :: "'a\<Colon>ordered_euclidean_space"
-  shows "{a<..<b} \<in> sets borel_space"
-  by (blast intro: borel_space_open)
-
-lemma atMost_borel[simp, intro]:
-  fixes a :: "'a\<Colon>ordered_euclidean_space"
-  shows "{..a} \<in> sets borel_space"
-  by (blast intro: borel_space_closed)
-
-lemma atLeast_borel[simp, intro]:
-  fixes a :: "'a\<Colon>ordered_euclidean_space"
-  shows "{a..} \<in> sets borel_space"
-  by (blast intro: borel_space_closed)
-
-lemma atLeastAtMost_borel[simp, intro]:
-  fixes a b :: "'a\<Colon>ordered_euclidean_space"
-  shows "{a..b} \<in> sets borel_space"
-  by (blast intro: borel_space_closed)
-
-lemma greaterThanAtMost_borel[simp, intro]:
-  fixes a b :: "'a\<Colon>ordered_euclidean_space"
-  shows "{a<..b} \<in> sets borel_space"
-  unfolding greaterThanAtMost_def by blast
-
-lemma atLeastLessThan_borel[simp, intro]:
-  fixes a b :: "'a\<Colon>ordered_euclidean_space"
-  shows "{a..<b} \<in> sets borel_space"
-  unfolding atLeastLessThan_def by blast
-
-lemma hafspace_less_borel[simp, intro]:
-  fixes a :: real
-  shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel_space"
-  by (auto intro!: borel_space_open open_halfspace_component_gt)
-
-lemma hafspace_greater_borel[simp, intro]:
-  fixes a :: real
-  shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel_space"
-  by (auto intro!: borel_space_open open_halfspace_component_lt)
-
-lemma hafspace_less_eq_borel[simp, intro]:
-  fixes a :: real
-  shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel_space"
-  by (auto intro!: borel_space_closed closed_halfspace_component_ge)
-
-lemma hafspace_greater_eq_borel[simp, intro]:
-  fixes a :: real
-  shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel_space"
-  by (auto intro!: borel_space_closed closed_halfspace_component_le)
-
-lemma (in sigma_algebra) borel_measurable_less[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  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 "{w \<in> space M. f w < g w} =
-        (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
-    using Rats_dense_in_real by (auto simp add: Rats_def)
-  then show ?thesis using f g
-    by simp (blast intro: measurable_sets)
-qed
-
-lemma (in sigma_algebra) borel_measurable_le[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes f: "f \<in> borel_measurable M"
-  assumes g: "g \<in> borel_measurable M"
-  shows "{w \<in> space M. f w \<le> g w} \<in> sets M"
-proof -
-  have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
-    by auto
-  thus ?thesis using f g
-    by simp blast
-qed
-
-lemma (in sigma_algebra) borel_measurable_eq[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  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 "{w \<in> space M. f w = g w} =
-        {w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}"
-    by auto
-  thus ?thesis using f g by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_neq[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  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
-
-subsection "Borel space equals sigma algebras over intervals"
-
-lemma rational_boxes:
-  fixes x :: "'a\<Colon>ordered_euclidean_space"
-  assumes "0 < e"
-  shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
-proof -
-  def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
-  then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos)
-  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i")
-  proof
-    fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e
-    show "?th i" by auto
-  qed
-  from choice[OF this] guess a .. note a = this
-  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i")
-  proof
-    fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e
-    show "?th i" by auto
-  qed
-  from choice[OF this] guess b .. note b = this
-  { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
-    have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
-      unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
-    also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
-    proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
-      fix i assume i: "i \<in> {..<DIM('a)}"
-      have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto
-      moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto
-      moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto
-      ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto
-      then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
-        unfolding e'_def by (auto simp: dist_real_def)
-      then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
-        by (rule power_strict_mono) auto
-      then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
-        by (simp add: power_divide)
-    qed auto
-    also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive)
-    finally have "dist x y < e" . }
-  with a b show ?thesis
-    apply (rule_tac exI[of _ "Chi a"])
-    apply (rule_tac exI[of _ "Chi b"])
-    using eucl_less[where 'a='a] by auto
-qed
-
-lemma ex_rat_list:
-  fixes x :: "'a\<Colon>ordered_euclidean_space"
-  assumes "\<And> i. x $$ i \<in> \<rat>"
-  shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
-proof -
-  have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast
-  from choice[OF this] guess r ..
-  then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
-qed
-
-lemma open_UNION:
-  fixes M :: "'a\<Colon>ordered_euclidean_space set"
-  assumes "open M"
-  shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
-                   (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
-    (is "M = UNION ?idx ?box")
-proof safe
-  fix x assume "x \<in> M"
-  obtain e where e: "e > 0" "ball x e \<subseteq> M"
-    using openE[OF assms `x \<in> M`] by auto
-  then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
-    using rational_boxes[OF e(1)] by blast
-  then obtain p q where pq: "length p = DIM ('a)"
-                            "length q = DIM ('a)"
-                            "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
-    using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast
-  hence p: "Chi (of_rat \<circ> op ! p) = a"
-    using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a]
-    unfolding o_def by auto
-  from pq have q: "Chi (of_rat \<circ> op ! q) = b"
-    using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b]
-    unfolding o_def by auto
-  have "x \<in> ?box (p, q)"
-    using p q ab by auto
-  thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto
-qed auto
-
-lemma halfspace_span_open:
-  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))
-    \<subseteq> sets borel_space"
-  by (auto intro!: borel_space.sigma_sets_subset[simplified] borel_space_open
-                   open_halfspace_component_lt simp: sets_sigma)
-
-lemma halfspace_lt_in_halfspace:
-  "{x\<Colon>'a. x $$ i < a} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
-  unfolding sets_sigma by (rule sigma_sets.Basic) auto
-
-lemma halfspace_gt_in_halfspace:
-  "{x\<Colon>'a. a < x $$ i} \<in> sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})))"
-    (is "?set \<in> sets ?SIGMA")
-proof -
-  interpret sigma_algebra ?SIGMA by (rule sigma_algebra_sigma) simp
-  have *: "?set = (\<Union>n. space ?SIGMA - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
-  proof (safe, simp_all add: not_less)
-    fix x assume "a < x $$ i"
-    with reals_Archimedean[of "x $$ i - a"]
-    obtain n where "a + 1 / real (Suc n) < x $$ i"
-      by (auto simp: inverse_eq_divide field_simps)
-    then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i"
-      by (blast intro: less_imp_le)
-  next
-    fix x n
-    have "a < a + 1 / real (Suc n)" by auto
-    also assume "\<dots> \<le> x"
-    finally show "a < x" .
-  qed
-  show "?set \<in> sets ?SIGMA" unfolding *
-    by (safe intro!: countable_UN Diff halfspace_lt_in_halfspace)
-qed
-
-lemma (in sigma_algebra) sets_sigma_subset:
-  assumes "A = space M"
-  assumes "B \<subseteq> sets M"
-  shows "sets (sigma A B) \<subseteq> sets M"
-  by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms)
-
-lemma open_span_halfspace:
-  "sets borel_space \<subseteq> sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})))"
-    (is "_ \<subseteq> sets ?SIGMA")
-proof (unfold borel_space_def, rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp
-  then interpret sigma_algebra ?SIGMA .
-  fix S :: "'a set" assume "S \<in> open" then have "open S" unfolding mem_def .
-  from open_UNION[OF this]
-  obtain I where *: "S =
-    (\<Union>(a, b)\<in>I.
-        (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
-        (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
-    unfolding greaterThanLessThan_def
-    unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]
-    unfolding eucl_lessThan_eq_halfspaces[where 'a='a]
-    by blast
-  show "S \<in> sets ?SIGMA"
-    unfolding *
-    by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace)
-qed auto
-
-lemma halfspace_span_halfspace_le:
-  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq>
-   sets (sigma UNIV (range (\<lambda> (a, i). {x. x $$ i \<le> a})))"
-  (is "_ \<subseteq> sets ?SIGMA")
-proof (rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-  then interpret sigma_algebra ?SIGMA .
-  fix a i
-  have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
-  proof (safe, simp_all)
-    fix x::'a assume *: "x$$i < a"
-    with reals_Archimedean[of "a - x$$i"]
-    obtain n where "x $$ i < a - 1 / (real (Suc n))"
-      by (auto simp: field_simps inverse_eq_divide)
-    then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"
-      by (blast intro: less_imp_le)
-  next
-    fix x::'a and n
-    assume "x$$i \<le> a - 1 / real (Suc n)"
-    also have "\<dots> < a" by auto
-    finally show "x$$i < a" .
-  qed
-  show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
-    by (safe intro!: countable_UN)
-       (auto simp: sets_sigma intro!: sigma_sets.Basic)
-qed auto
-
-lemma halfspace_span_halfspace_ge:
-  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))) \<subseteq> 
-   sets (sigma UNIV (range (\<lambda> (a, i). {x. a \<le> x $$ i})))"
-  (is "_ \<subseteq> sets ?SIGMA")
-proof (rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-  then interpret sigma_algebra ?SIGMA .
-  fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
-  show "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
-    by (safe intro!: Diff)
-       (auto simp: sets_sigma intro!: sigma_sets.Basic)
-qed auto
-
-lemma halfspace_le_span_halfspace_gt:
-  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq> 
-   sets (sigma UNIV (range (\<lambda> (a, i). {x. a < x $$ i})))"
-  (is "_ \<subseteq> sets ?SIGMA")
-proof (rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-  then interpret sigma_algebra ?SIGMA .
-  fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
-  show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
-    by (safe intro!: Diff)
-       (auto simp: sets_sigma intro!: sigma_sets.Basic)
-qed auto
-
-lemma halfspace_le_span_atMost:
-  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
-   sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})))"
-  (is "_ \<subseteq> sets ?SIGMA")
-proof (rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-  then interpret sigma_algebra ?SIGMA .
-  fix a i
-  show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
-  proof cases
-    assume "i < DIM('a)"
-    then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
-    proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
-      fix x
-      from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
-      then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
-        by (subst (asm) Max_le_iff) auto
-      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
-        by (auto intro!: exI[of _ k])
-    qed
-    show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
-      by (safe intro!: countable_UN)
-         (auto simp: sets_sigma intro!: sigma_sets.Basic)
-  next
-    assume "\<not> i < DIM('a)"
-    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
-      using top by auto
-  qed
-qed auto
-
-lemma halfspace_le_span_greaterThan:
-  "sets (sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a}))) \<subseteq>
-   sets (sigma UNIV (range (\<lambda>a. {a<..})))"
-  (is "_ \<subseteq> sets ?SIGMA")
-proof (rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-  then interpret sigma_algebra ?SIGMA .
-  fix a i
-  show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
-  proof cases
-    assume "i < DIM('a)"
-    have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
-    also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
-    proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
-      fix x
-      from real_arch_lt[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
-      guess k::nat .. note k = this
-      { fix i assume "i < DIM('a)"
-        then have "-x$$i < real k"
-          using k by (subst (asm) Max_less_iff) auto
-        then have "- real k < x$$i" by simp }
-      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
-        by (auto intro!: exI[of _ k])
-    qed
-    finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
-      apply (simp only:)
-      apply (safe intro!: countable_UN Diff)
-      by (auto simp: sets_sigma intro!: sigma_sets.Basic)
-  next
-    assume "\<not> i < DIM('a)"
-    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
-      using top by auto
-  qed
-qed auto
-
-lemma atMost_span_atLeastAtMost:
-  "sets (sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))) \<subseteq>
-   sets (sigma UNIV (range (\<lambda>(a,b). {a..b})))"
-  (is "_ \<subseteq> sets ?SIGMA")
-proof (rule sigma_algebra.sets_sigma_subset, safe)
-  show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-  then interpret sigma_algebra ?SIGMA .
-  fix a::'a
-  have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
-  proof (safe, simp_all add: eucl_le[where 'a='a])
-    fix x
-    from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
-    guess k::nat .. note k = this
-    { fix i assume "i < DIM('a)"
-      with k have "- x$$i \<le> real k"
-        by (subst (asm) Max_le_iff) (auto simp: field_simps)
-      then have "- real k \<le> x$$i" by simp }
-    then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
-      by (auto intro!: exI[of _ k])
-  qed
-  show "{..a} \<in> sets ?SIGMA" unfolding *
-    by (safe intro!: countable_UN)
-       (auto simp: sets_sigma intro!: sigma_sets.Basic)
-qed auto
-
-lemma borel_space_eq_greaterThanLessThan:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})))"
-    (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    unfolding borel_space_def
-  proof (rule sigma_algebra.sets_sigma_subset, safe)
-    show "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
-    then interpret sigma_algebra ?SIGMA .
-    fix M :: "'a set" assume "M \<in> open"
-    then have "open M" by (simp add: mem_def)
-    show "M \<in> sets ?SIGMA"
-      apply (subst open_UNION[OF `open M`])
-      apply (safe intro!: countable_UN)
-      by (auto simp add: sigma_def intro!: sigma_sets.Basic)
-  qed auto
-qed
-
-lemma borel_space_eq_atMost:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})))"
-    (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace
-    by auto
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma borel_space_eq_atLeastAtMost:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})))"
-   (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using atMost_span_atLeastAtMost halfspace_le_span_atMost
-      halfspace_span_halfspace_le open_span_halfspace
-    by auto
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma borel_space_eq_greaterThan:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})))"
-   (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using halfspace_le_span_greaterThan
-      halfspace_span_halfspace_le open_span_halfspace
-    by auto
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma borel_space_eq_halfspace_le:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})))"
-   (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using open_span_halfspace halfspace_span_halfspace_le by auto
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma borel_space_eq_halfspace_less:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})))"
-   (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using open_span_halfspace .
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma borel_space_eq_halfspace_gt:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})))"
-   (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma borel_space_eq_halfspace_ge:
-  "sets borel_space = sets (sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})))"
-   (is "_ = sets ?SIGMA")
-proof (rule antisym)
-  show "sets borel_space \<subseteq> sets ?SIGMA"
-    using halfspace_span_halfspace_ge open_span_halfspace by auto
-  show "sets ?SIGMA \<subseteq> sets borel_space"
-    by (rule borel_space.sets_sigma_subset) auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_halfspacesI:
-  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
-  assumes "sets borel_space = sets (sigma UNIV (range F))"
-  and "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
-  and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
-  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
-proof safe
-  fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
-  then show "S a i \<in> sets M" unfolding assms
-    by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def)
-next
-  assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
-  { fix a i have "S a i \<in> sets M"
-    proof cases
-      assume "i < DIM('c)"
-      with a show ?thesis unfolding assms(2) by simp
-    next
-      assume "\<not> i < DIM('c)"
-      from assms(3)[OF this] show ?thesis .
-    qed }
-  then have "f \<in> measurable M (sigma UNIV (range F))"
-    by (auto intro!: measurable_sigma simp: assms(2))
-  then show "f \<in> borel_measurable M" unfolding measurable_def
-    unfolding assms(1) by simp
-qed
-
-lemma (in sigma_algebra) borel_measurable_iff_halfspace_le:
-  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
-  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
-  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_le]) auto
-
-lemma (in sigma_algebra) borel_measurable_iff_halfspace_less:
-  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
-  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
-  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_less]) auto
-
-lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge:
-  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
-  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
-  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_ge]) auto
-
-lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater:
-  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
-  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
-  by (rule borel_measurable_halfspacesI[OF borel_space_eq_halfspace_gt]) auto
-
-lemma (in sigma_algebra) borel_measurable_iff_le:
-  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
-  using borel_measurable_iff_halfspace_le[where 'c=real] by simp
-
-lemma (in sigma_algebra) borel_measurable_iff_less:
-  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
-  using borel_measurable_iff_halfspace_less[where 'c=real] by simp
-
-lemma (in sigma_algebra) borel_measurable_iff_ge:
-  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
-  using borel_measurable_iff_halfspace_ge[where 'c=real] by simp
-
-lemma (in sigma_algebra) borel_measurable_iff_greater:
-  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
-  using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
-
-lemma borel_measureable_euclidean_component:
-  "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel_space"
-  unfolding borel_space_def[where 'a=real]
-proof (rule borel_space.measurable_sigma)
-  fix S::"real set" assume "S \<in> open" then have "open S" unfolding mem_def .
-  from open_vimage_euclidean_component[OF this]
-  show "(\<lambda>x. x $$ i) -` S \<inter> space borel_space \<in> sets borel_space"
-    by (auto intro: borel_space_open)
-qed auto
-
-lemma (in sigma_algebra) borel_measureable_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_measureable_euclidean_component)
-next
-  assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
-  then show "f \<in> borel_measurable M"
-    unfolding borel_measurable_iff_halfspace_le by auto
-qed
-
-subsection "Borel measurable operators"
-
-lemma (in sigma_algebra) affine_borel_measurable_vector:
-  fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
-  assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
-proof (rule borel_measurableI)
-  fix S :: "'x set" assume "open S"
-  show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
-  proof cases
-    assume "b \<noteq> 0"
-    with `open S` have "((\<lambda>x. (- a + x) /\<^sub>R b) ` S) \<in> open" (is "?S \<in> open")
-      by (auto intro!: open_affinity simp: scaleR.add_right mem_def)
-    hence "?S \<in> sets borel_space"
-      unfolding borel_space_def by (auto simp: sigma_def intro!: sigma_sets.Basic)
-    moreover
-    from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
-      apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
-    ultimately show ?thesis using assms unfolding in_borel_measurable_borel_space
-      by auto
-  qed simp
-qed
-
-lemma (in sigma_algebra) affine_borel_measurable:
-  fixes g :: "'a \<Rightarrow> real"
-  assumes g: "g \<in> borel_measurable M"
-  shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
-  using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)
-
-lemma (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) borel_measurable_inverse[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes "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)
-qed
-
-lemma (in sigma_algebra) borel_measurable_divide[simp, intro]:
-  fixes f :: "'a \<Rightarrow> real"
-  assumes "f \<in> borel_measurable M"
-  and "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
-  unfolding field_divide_inverse
-  by (rule borel_measurable_inverse borel_measurable_times assms)+
-
-lemma (in sigma_algebra) borel_measurable_max[intro, simp]:
-  fixes f g :: "'a \<Rightarrow> real"
-  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
-
-lemma (in sigma_algebra) 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
-
-lemma (in sigma_algebra) borel_measurable_abs[simp, intro]:
-  assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
-proof -
-  have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def)
-  show ?thesis unfolding * using assms by auto
-qed
-
-section "Borel space over the real line with infinity"
-
-lemma borel_space_Real_measurable:
-  "A \<in> sets borel_space \<Longrightarrow> Real -` A \<in> sets borel_space"
-proof (rule borel_measurable_translate)
-  fix B :: "pinfreal set" assume "open B"
-  then obtain T x where T: "open T" "Real ` (T \<inter> {0..}) = B - {\<omega>}" and
-    x: "\<omega> \<in> B \<Longrightarrow> 0 \<le> x" "\<omega> \<in> B \<Longrightarrow> {Real x <..} \<subseteq> B"
-    unfolding open_pinfreal_def by blast
-
-  have "Real -` B = Real -` (B - {\<omega>})" by auto
-  also have "\<dots> = Real -` (Real ` (T \<inter> {0..}))" using T by simp
-  also have "\<dots> = (if 0 \<in> T then T \<union> {.. 0} else T \<inter> {0..})"
-    apply (auto simp add: Real_eq_Real image_iff)
-    apply (rule_tac x="max 0 x" in bexI)
-    by (auto simp: max_def)
-  finally show "Real -` B \<in> sets borel_space"
-    using `open T` by auto
-qed simp
-
-lemma borel_space_real_measurable:
-  "A \<in> sets borel_space \<Longrightarrow> (real -` A :: pinfreal set) \<in> sets borel_space"
-proof (rule borel_measurable_translate)
-  fix B :: "real set" assume "open B"
-  { fix x have "0 < real x \<longleftrightarrow> (\<exists>r>0. x = Real r)" by (cases x) auto }
-  note Ex_less_real = this
-  have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
-    by (force simp: Ex_less_real)
-
-  have "open (real -` (B \<inter> {0 <..}) :: pinfreal set)"
-    unfolding open_pinfreal_def using `open B`
-    by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
-  then show "(real -` B :: pinfreal set) \<in> sets borel_space" unfolding * by auto
-qed simp
-
-lemma (in sigma_algebra) borel_measurable_Real[intro, simp]:
-  assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
-  unfolding in_borel_measurable_borel_space
-proof safe
-  fix S :: "pinfreal set" assume "S \<in> sets borel_space"
-  from borel_space_Real_measurable[OF this]
-  have "(Real \<circ> f) -` S \<inter> space M \<in> sets M"
-    using assms
-    unfolding vimage_compose in_borel_measurable_borel_space
-    by auto
-  thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
-qed
-
-lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  assumes "f \<in> borel_measurable M"
-  shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
-  unfolding in_borel_measurable_borel_space
-proof safe
-  fix S :: "real set" assume "S \<in> sets borel_space"
-  from borel_space_real_measurable[OF this]
-  have "(real \<circ> f) -` S \<inter> space M \<in> sets M"
-    using assms
-    unfolding vimage_compose in_borel_measurable_borel_space
-    by auto
-  thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
-qed
-
-lemma (in sigma_algebra) borel_measurable_Real_eq:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
-  shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
-proof
-  have [simp]: "(\<lambda>x. Real (f x)) -` {\<omega>} \<inter> space M = {}"
-    by auto
-  assume "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
-  hence "(\<lambda>x. real (Real (f x))) \<in> borel_measurable M"
-    by (rule borel_measurable_real)
-  moreover have "\<And>x. x \<in> space M \<Longrightarrow> real (Real (f x)) = f x"
-    using assms by auto
-  ultimately show "f \<in> borel_measurable M"
-    by (simp cong: measurable_cong)
-qed auto
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_eq_real:
-  "f \<in> borel_measurable M \<longleftrightarrow>
-    ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<omega>} \<inter> space M \<in> sets M)"
-proof safe
-  assume "f \<in> borel_measurable M"
-  then show "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
-    by (auto intro: borel_measurable_vimage borel_measurable_real)
-next
-  assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
-  have "f -` {\<omega>} \<inter> space M = {x\<in>space M. f x = \<omega>}" by auto
-  with * have **: "{x\<in>space M. f x = \<omega>} \<in> sets M" by simp
-  have f: "f = (\<lambda>x. if f x = \<omega> then \<omega> else Real (real (f x)))"
-    by (simp add: fun_eq_iff Real_real)
-  show "f \<in> borel_measurable M"
-    apply (subst f)
-    apply (rule measurable_If)
-    using * ** by auto
-qed
-
-lemma (in sigma_algebra) less_eq_ge_measurable:
-  fixes f :: "'a \<Rightarrow> 'c::linorder"
-  shows "{x\<in>space M. a < f x} \<in> sets M \<longleftrightarrow> {x\<in>space M. f x \<le> a} \<in> sets M"
-proof
-  assume "{x\<in>space M. f x \<le> a} \<in> sets M"
-  moreover have "{x\<in>space M. a < f x} = space M - {x\<in>space M. f x \<le> a}" by auto
-  ultimately show "{x\<in>space M. a < f x} \<in> sets M" by auto
-next
-  assume "{x\<in>space M. a < f x} \<in> sets M"
-  moreover have "{x\<in>space M. f x \<le> a} = space M - {x\<in>space M. a < f x}" by auto
-  ultimately show "{x\<in>space M. f x \<le> a} \<in> sets M" by auto
-qed
-
-lemma (in sigma_algebra) greater_eq_le_measurable:
-  fixes f :: "'a \<Rightarrow> 'c::linorder"
-  shows "{x\<in>space M. f x < a} \<in> sets M \<longleftrightarrow> {x\<in>space M. a \<le> f x} \<in> sets M"
-proof
-  assume "{x\<in>space M. a \<le> f x} \<in> sets M"
-  moreover have "{x\<in>space M. f x < a} = space M - {x\<in>space M. a \<le> f x}" by auto
-  ultimately show "{x\<in>space M. f x < a} \<in> sets M" by auto
-next
-  assume "{x\<in>space M. f x < a} \<in> sets M"
-  moreover have "{x\<in>space M. a \<le> f x} = space M - {x\<in>space M. f x < a}" by auto
-  ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
-qed
-
-lemma (in sigma_algebra) less_eq_le_pinfreal_measurable:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  shows "(\<forall>a. {x\<in>space M. a < f x} \<in> sets M) \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
-proof
-  assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
-  show "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
-  proof
-    fix a show "{x \<in> space M. a < f x} \<in> sets M"
-    proof (cases a)
-      case (preal r)
-      have "{x\<in>space M. a < f x} = (\<Union>i. {x\<in>space M. a + inverse (of_nat (Suc i)) \<le> f x})"
-      proof safe
-        fix x assume "a < f x" and [simp]: "x \<in> space M"
-        with ex_pinfreal_inverse_of_nat_Suc_less[of "f x - a"]
-        obtain n where "a + inverse (of_nat (Suc n)) < f x"
-          by (cases "f x", auto simp: pinfreal_minus_order)
-        then have "a + inverse (of_nat (Suc n)) \<le> f x" by simp
-        then show "x \<in> (\<Union>i. {x \<in> space M. a + inverse (of_nat (Suc i)) \<le> f x})"
-          by auto
-      next
-        fix i x assume [simp]: "x \<in> space M"
-        have "a < a + inverse (of_nat (Suc i))" using preal by auto
-        also assume "a + inverse (of_nat (Suc i)) \<le> f x"
-        finally show "a < f x" .
-      qed
-      with a show ?thesis by auto
-    qed simp
-  qed
-next
-  assume a': "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
-  then have a: "\<forall>a. {x \<in> space M. f x \<le> a} \<in> sets M" unfolding less_eq_ge_measurable .
-  show "\<forall>a. {x \<in> space M. a \<le> f x} \<in> sets M" unfolding greater_eq_le_measurable[symmetric]
-  proof
-    fix a show "{x \<in> space M. f x < a} \<in> sets M"
-    proof (cases a)
-      case (preal r)
-      show ?thesis
-      proof cases
-        assume "a = 0" then show ?thesis by simp
-      next
-        assume "a \<noteq> 0"
-        have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
-        proof safe
-          fix x assume "f x < a" and [simp]: "x \<in> space M"
-          with ex_pinfreal_inverse_of_nat_Suc_less[of "a - f x"]
-          obtain n where "inverse (of_nat (Suc n)) < a - f x"
-            using preal by (cases "f x") auto
-          then have "f x \<le> a - inverse (of_nat (Suc n)) "
-            using preal by (cases "f x") (auto split: split_if_asm)
-          then show "x \<in> (\<Union>i. {x \<in> space M. f x \<le> a - inverse (of_nat (Suc i))})"
-            by auto
-        next
-          fix i x assume [simp]: "x \<in> space M"
-          assume "f x \<le> a - inverse (of_nat (Suc i))"
-          also have "\<dots> < a" using `a \<noteq> 0` preal by auto
-          finally show "f x < a" .
-        qed
-        with a show ?thesis by auto
-      qed
-    next
-      case infinite
-      have "f -` {\<omega>} \<inter> space M = (\<Inter>n. {x\<in>space M. of_nat n < f x})"
-      proof (safe, simp_all, safe)
-        fix x assume *: "\<forall>n::nat. Real (real n) < f x"
-        show "f x = \<omega>"    proof (rule ccontr)
-          assume "f x \<noteq> \<omega>"
-          with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n"
-            by (auto simp: pinfreal_noteq_omega_Ex)
-          with *[THEN spec, of n] show False by auto
-        qed
-      qed
-      with a' have \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" by auto
-      moreover have "{x \<in> space M. f x < a} = space M - f -` {\<omega>} \<inter> space M"
-        using infinite by auto
-      ultimately show ?thesis by auto
-    qed
-  qed
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_iff_greater:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
-proof safe
-  fix a assume f: "f \<in> borel_measurable M"
-  have "{x\<in>space M. a < f x} = f -` {a <..} \<inter> space M" by auto
-  with f show "{x\<in>space M. a < f x} \<in> sets M"
-    by (auto intro!: measurable_sets)
-next
-  assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
-  hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
-    unfolding less_eq_le_pinfreal_measurable
-    unfolding greater_eq_le_measurable .
-
-  show "f \<in> borel_measurable M" unfolding borel_measurable_pinfreal_eq_real borel_measurable_iff_greater
-  proof safe
-    have "f -` {\<omega>} \<inter> space M = space M - {x\<in>space M. f x < \<omega>}" by auto
-    then show \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" using ** by auto
-
-    fix a
-    have "{w \<in> space M. a < real (f w)} =
-      (if 0 \<le> a then {w\<in>space M. Real a < f w} - (f -` {\<omega>} \<inter> space M) else space M)"
-    proof (split split_if, safe del: notI)
-      fix x assume "0 \<le> a"
-      { assume "a < real (f x)" then show "Real a < f x" "x \<notin> f -` {\<omega>} \<inter> space M"
-          using `0 \<le> a` by (cases "f x", auto) }
-      { assume "Real a < f x" "x \<notin> f -` {\<omega>}" then show "a < real (f x)"
-          using `0 \<le> a` by (cases "f x", auto) }
-    next
-      fix x assume "\<not> 0 \<le> a" then show "a < real (f x)" by (cases "f x") auto
-    qed
-    then show "{w \<in> space M. a < real (f w)} \<in> sets M"
-      using \<omega> * by (auto intro!: Diff)
-  qed
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_iff_less:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
-  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable greater_eq_le_measurable .
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_iff_le:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
-  using borel_measurable_pinfreal_iff_greater unfolding less_eq_ge_measurable .
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_iff_ge:
-  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
-  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable .
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_eq_const:
-  fixes f :: "'a \<Rightarrow> pinfreal" 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 (in sigma_algebra) borel_measurable_pinfreal_neq_const:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  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 (in sigma_algebra) borel_measurable_pinfreal_less[intro,simp]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
-  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 "(\<lambda>x. real (f x)) \<in> borel_measurable M"
-    "(\<lambda>x. real (g x)) \<in> borel_measurable M"
-    using assms by (auto intro!: borel_measurable_real)
-  from borel_measurable_less[OF this]
-  have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
-  moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pinfreal_neq_const)
-  moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_eq_const)
-  moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_neq_const)
-  moreover have "{x \<in> space M. f x < g x} = ({x \<in> space M. g x = \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>}) \<union>
-    ({x \<in> space M. g x \<noteq> \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>} \<inter> {x \<in> space M. real (f x) < real (g x)})"
-    by (auto simp: real_of_pinfreal_strict_mono_iff)
-  ultimately show ?thesis by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_le[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  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"
-proof -
-  have "{x \<in> space M. f x \<le> g x} = space M - {x \<in> space M. g x < f x}" by auto
-  then show ?thesis using g f by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_eq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  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 (in sigma_algebra) borel_measurable_pinfreal_neq[intro,simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  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 (in sigma_algebra) borel_measurable_pinfreal_add[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  assumes measure: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
-proof -
-  have *: "(\<lambda>x. f x + g x) =
-     (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
-     by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex)
-  show ?thesis using assms unfolding *
-    by (auto intro!: measurable_If)
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_times[intro, simp]:
-  fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
-proof -
-  have *: "(\<lambda>x. f x * g x) =
-     (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else
-      Real (real (f x) * real (g x)))"
-     by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex)
-  show ?thesis using assms unfolding *
-    by (auto intro!: measurable_If)
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_setsum[simp, intro]:
-  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pinfreal"
-  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 add: borel_measurable_const)
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_min[intro, simp]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
-  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 (in sigma_algebra) borel_measurable_pinfreal_max[intro]:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
-  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 (in sigma_algebra) borel_measurable_SUP[simp, intro]:
-  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
-  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
-  shows "(SUP i : A. f i) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
-  unfolding borel_measurable_pinfreal_iff_greater
-proof safe
-  fix a
-  have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
-    by (auto simp: less_Sup_iff SUPR_def[where 'a=pinfreal] SUPR_fun_expand[where 'c=pinfreal])
-  then show "{x\<in>space M. a < ?sup x} \<in> sets M"
-    using assms by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
-  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
-  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
-  shows "(INF i : A. f i) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
-  unfolding borel_measurable_pinfreal_iff_less
-proof safe
-  fix a
-  have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
-    by (auto simp: Inf_less_iff INFI_def[where 'a=pinfreal] INFI_fun_expand)
-  then show "{x\<in>space M. ?inf x < a} \<in> sets M"
-    using assms by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_pinfreal_diff:
-  fixes f g :: "'a \<Rightarrow> pinfreal"
-  assumes "f \<in> borel_measurable M"
-  assumes "g \<in> borel_measurable M"
-  shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
-  unfolding borel_measurable_pinfreal_iff_greater
-proof safe
-  fix a
-  have "{x \<in> space M. a < f x - g x} = {x \<in> space M. g x + a < f x}"
-    by (simp add: pinfreal_less_minus_iff)
-  then show "{x \<in> space M. a < f x - g x} \<in> sets M"
-    using assms by auto
-qed
-
-lemma (in sigma_algebra) borel_measurable_psuminf:
-  assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
-  using assms unfolding psuminf_def
-  by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
-
-section "LIMSEQ is borel measurable"
-
-lemma (in sigma_algebra) borel_measurable_LIMSEQ:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
-  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 -
-  let "?pu x i" = "max (u i x) 0"
-  let "?nu x i" = "max (- u i x) 0"
-
-  { fix x assume x: "x \<in> space M"
-    have "(?pu x) ----> max (u' x) 0"
-      "(?nu x) ----> max (- u' x) 0"
-      using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
-    from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
-    have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
-      "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
-      by (simp_all add: Real_max'[symmetric]) }
-  note eq = this
-
-  have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
-    by auto
-
-  have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
-       "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
-    using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
-  with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
-  have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
-       "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
-    unfolding SUPR_fun_expand INFI_fun_expand by auto
-  note this[THEN borel_measurable_real]
-  from borel_measurable_diff[OF this]
-  show ?thesis unfolding * .
-qed
-
-end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Borel_Space.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -0,0 +1,1466 @@
+(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *)
+
+header {*Borel spaces*}
+
+theory Borel_Space
+  imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
+begin
+
+lemma (in sigma_algebra) sets_sigma_subset:
+  assumes "space N = space M"
+  assumes "sets N \<subseteq> sets M"
+  shows "sets (sigma N) \<subseteq> sets M"
+  by (unfold assms sets_sigma, rule sigma_sets_subset, rule assms)
+
+lemma LIMSEQ_max:
+  "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
+  by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
+
+section "Generic Borel spaces"
+
+definition "borel = sigma \<lparr> space = UNIV::'a::topological_space set, sets = open\<rparr>"
+abbreviation "borel_measurable M \<equiv> measurable M borel"
+
+interpretation borel: sigma_algebra borel
+  by (auto simp: borel_def intro!: sigma_algebra_sigma)
+
+lemma in_borel_measurable:
+   "f \<in> borel_measurable M \<longleftrightarrow>
+    (\<forall>S \<in> sets (sigma \<lparr> space = UNIV, sets = open\<rparr>).
+      f -` S \<inter> space M \<in> sets M)"
+  by (auto simp add: measurable_def borel_def)
+
+lemma in_borel_measurable_borel:
+   "f \<in> borel_measurable M \<longleftrightarrow>
+    (\<forall>S \<in> sets borel.
+      f -` S \<inter> space M \<in> sets M)"
+  by (auto simp add: measurable_def borel_def)
+
+lemma space_borel[simp]: "space borel = UNIV"
+  unfolding borel_def by auto
+
+lemma borel_open[simp]:
+  assumes "open A" shows "A \<in> sets borel"
+proof -
+  have "A \<in> open" unfolding mem_def using assms .
+  thus ?thesis unfolding borel_def sigma_def by (auto intro!: sigma_sets.Basic)
+qed
+
+lemma borel_closed[simp]:
+  assumes "closed A" shows "A \<in> sets borel"
+proof -
+  have "space borel - (- A) \<in> sets borel"
+    using assms unfolding closed_def by (blast intro: borel_open)
+  thus ?thesis by simp
+qed
+
+lemma (in sigma_algebra) borel_measurable_vimage:
+  fixes f :: "'a \<Rightarrow> 'x::t2_space"
+  assumes borel: "f \<in> borel_measurable M"
+  shows "f -` {x} \<inter> space M \<in> sets M"
+proof (cases "x \<in> f ` space M")
+  case True then obtain y where "x = f y" by auto
+  from closed_sing[of "f y"]
+  have "{f y} \<in> sets borel" by (rule borel_closed)
+  with assms show ?thesis
+    unfolding in_borel_measurable_borel `x = f y` by auto
+next
+  case False hence "f -` {x} \<inter> space M = {}" by auto
+  thus ?thesis by auto
+qed
+
+lemma (in sigma_algebra) borel_measurableI:
+  fixes f :: "'a \<Rightarrow> 'x\<Colon>topological_space"
+  assumes "\<And>S. open S \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+  shows "f \<in> borel_measurable M"
+  unfolding borel_def
+proof (rule measurable_sigma, simp_all)
+  fix S :: "'x set" assume "S \<in> open" thus "f -` S \<inter> space M \<in> sets M"
+    using assms[of S] by (simp add: mem_def)
+qed
+
+lemma borel_singleton[simp, intro]:
+  fixes x :: "'a::t1_space"
+  shows "A \<in> sets borel \<Longrightarrow> insert x A \<in> sets borel"
+  proof (rule borel.insert_in_sets)
+    show "{x} \<in> sets borel"
+      using closed_sing[of x] by (rule borel_closed)
+  qed simp
+
+lemma (in sigma_algebra) borel_measurable_const[simp, intro]:
+  "(\<lambda>x. c) \<in> borel_measurable M"
+  by (auto intro!: measurable_const)
+
+lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]:
+  assumes A: "A \<in> sets M"
+  shows "indicator A \<in> borel_measurable M"
+  unfolding indicator_def_raw using A
+  by (auto intro!: measurable_If_set borel_measurable_const)
+
+lemma (in sigma_algebra) borel_measurable_indicator_iff:
+  "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
+    (is "?I \<in> borel_measurable M \<longleftrightarrow> _")
+proof
+  assume "?I \<in> borel_measurable M"
+  then have "?I -` {1} \<inter> space M \<in> sets M"
+    unfolding measurable_def by auto
+  also have "?I -` {1} \<inter> space M = A \<inter> space M"
+    unfolding indicator_def_raw by auto
+  finally show "A \<inter> space M \<in> sets M" .
+next
+  assume "A \<inter> space M \<in> sets M"
+  moreover have "?I \<in> borel_measurable M \<longleftrightarrow>
+    (indicator (A \<inter> space M) :: 'a \<Rightarrow> 'x) \<in> borel_measurable M"
+    by (intro measurable_cong) (auto simp: indicator_def)
+  ultimately show "?I \<in> borel_measurable M" by auto
+qed
+
+lemma borel_measurable_translate:
+  assumes "A \<in> sets borel" and trans: "\<And>B. open B \<Longrightarrow> f -` B \<in> sets borel"
+  shows "f -` A \<in> sets borel"
+proof -
+  have "A \<in> sigma_sets UNIV open" using assms
+    by (simp add: borel_def sigma_def)
+  thus ?thesis
+  proof (induct rule: sigma_sets.induct)
+    case (Basic a) thus ?case using trans[of a] by (simp add: mem_def)
+  next
+    case (Compl a)
+    moreover have "UNIV \<in> sets borel"
+      using borel.top by simp
+    ultimately show ?case
+      by (auto simp: vimage_Diff borel.Diff)
+  qed (auto simp add: vimage_UN)
+qed
+
+lemma (in sigma_algebra) borel_measurable_restricted:
+  fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
+  shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
+    (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
+    (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
+proof -
+  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
+  have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
+    by (auto intro!: measurable_cong)
+  show ?thesis unfolding *
+    unfolding in_borel_measurable_borel
+  proof (simp, safe)
+    fix S :: "'x set" assume "S \<in> sets borel"
+      "\<forall>S\<in>sets borel. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+    then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
+    then have f: "?f -` S \<inter> A \<in> sets M"
+      using `A \<in> sets M` sets_into_space by fastsimp
+    show "?f -` S \<inter> space M \<in> sets M"
+    proof cases
+      assume "0 \<in> S"
+      then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
+        using `A \<in> sets M` sets_into_space by auto
+      then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
+    next
+      assume "0 \<notin> S"
+      then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
+        using `A \<in> sets M` sets_into_space
+        by (auto simp: indicator_def split: split_if_asm)
+      then show ?thesis using f by auto
+    qed
+  next
+    fix S :: "'x set" assume "S \<in> sets borel"
+      "\<forall>S\<in>sets borel. ?f -` S \<inter> space M \<in> sets M"
+    then have f: "?f -` S \<inter> space M \<in> sets M" by auto
+    then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+      using `A \<in> sets M` sets_into_space
+      apply (simp add: image_iff)
+      apply (rule bexI[OF _ f])
+      by auto
+  qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_subalgebra:
+  assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)"
+  shows "f \<in> borel_measurable M"
+  using assms unfolding measurable_def by auto
+
+section "Borel spaces on euclidean spaces"
+
+lemma lessThan_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{..< a} \<in> sets borel"
+  by (blast intro: borel_open)
+
+lemma greaterThan_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a <..} \<in> sets borel"
+  by (blast intro: borel_open)
+
+lemma greaterThanLessThan_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a<..<b} \<in> sets borel"
+  by (blast intro: borel_open)
+
+lemma atMost_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{..a} \<in> sets borel"
+  by (blast intro: borel_closed)
+
+lemma atLeast_borel[simp, intro]:
+  fixes a :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a..} \<in> sets borel"
+  by (blast intro: borel_closed)
+
+lemma atLeastAtMost_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a..b} \<in> sets borel"
+  by (blast intro: borel_closed)
+
+lemma greaterThanAtMost_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a<..b} \<in> sets borel"
+  unfolding greaterThanAtMost_def by blast
+
+lemma atLeastLessThan_borel[simp, intro]:
+  fixes a b :: "'a\<Colon>ordered_euclidean_space"
+  shows "{a..<b} \<in> sets borel"
+  unfolding atLeastLessThan_def by blast
+
+lemma hafspace_less_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. a < x $$ i} \<in> sets borel"
+  by (auto intro!: borel_open open_halfspace_component_gt)
+
+lemma hafspace_greater_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. x $$ i < a} \<in> sets borel"
+  by (auto intro!: borel_open open_halfspace_component_lt)
+
+lemma hafspace_less_eq_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. a \<le> x $$ i} \<in> sets borel"
+  by (auto intro!: borel_closed closed_halfspace_component_ge)
+
+lemma hafspace_greater_eq_borel[simp, intro]:
+  fixes a :: real
+  shows "{x::'a::euclidean_space. x $$ i \<le> a} \<in> sets borel"
+  by (auto intro!: borel_closed closed_halfspace_component_le)
+
+lemma (in sigma_algebra) borel_measurable_less[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  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 "{w \<in> space M. f w < g w} =
+        (\<Union>r. (f -` {..< of_rat r} \<inter> space M) \<inter> (g -` {of_rat r <..} \<inter> space M))"
+    using Rats_dense_in_real by (auto simp add: Rats_def)
+  then show ?thesis using f g
+    by simp (blast intro: measurable_sets)
+qed
+
+lemma (in sigma_algebra) borel_measurable_le[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable M"
+  assumes g: "g \<in> borel_measurable M"
+  shows "{w \<in> space M. f w \<le> g w} \<in> sets M"
+proof -
+  have "{w \<in> space M. f w \<le> g w} = space M - {w \<in> space M. g w < f w}"
+    by auto
+  thus ?thesis using f g
+    by simp blast
+qed
+
+lemma (in sigma_algebra) borel_measurable_eq[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  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 "{w \<in> space M. f w = g w} =
+        {w \<in> space M. f w \<le> g w} \<inter> {w \<in> space M. g w \<le> f w}"
+    by auto
+  thus ?thesis using f g by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_neq[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  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
+
+subsection "Borel space equals sigma algebras over intervals"
+
+lemma rational_boxes:
+  fixes x :: "'a\<Colon>ordered_euclidean_space"
+  assumes "0 < e"
+  shows "\<exists>a b. (\<forall>i. a $$ i \<in> \<rat>) \<and> (\<forall>i. b $$ i \<in> \<rat>) \<and> x \<in> {a <..< b} \<and> {a <..< b} \<subseteq> ball x e"
+proof -
+  def e' \<equiv> "e / (2 * sqrt (real (DIM ('a))))"
+  then have e: "0 < e'" using assms by (auto intro!: divide_pos_pos)
+  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> y < x $$ i \<and> x $$ i - y < e'" (is "\<forall>i. ?th i")
+  proof
+    fix i from Rats_dense_in_real[of "x $$ i - e'" "x $$ i"] e
+    show "?th i" by auto
+  qed
+  from choice[OF this] guess a .. note a = this
+  have "\<forall>i. \<exists>y. y \<in> \<rat> \<and> x $$ i < y \<and> y - x $$ i < e'" (is "\<forall>i. ?th i")
+  proof
+    fix i from Rats_dense_in_real[of "x $$ i" "x $$ i + e'"] e
+    show "?th i" by auto
+  qed
+  from choice[OF this] guess b .. note b = this
+  { fix y :: 'a assume *: "Chi a < y" "y < Chi b"
+    have "dist x y = sqrt (\<Sum>i<DIM('a). (dist (x $$ i) (y $$ i))\<twosuperior>)"
+      unfolding setL2_def[symmetric] by (rule euclidean_dist_l2)
+    also have "\<dots> < sqrt (\<Sum>i<DIM('a). e^2 / real (DIM('a)))"
+    proof (rule real_sqrt_less_mono, rule setsum_strict_mono)
+      fix i assume i: "i \<in> {..<DIM('a)}"
+      have "a i < y$$i \<and> y$$i < b i" using * i eucl_less[where 'a='a] by auto
+      moreover have "a i < x$$i" "x$$i - a i < e'" using a by auto
+      moreover have "x$$i < b i" "b i - x$$i < e'" using b by auto
+      ultimately have "\<bar>x$$i - y$$i\<bar> < 2 * e'" by auto
+      then have "dist (x $$ i) (y $$ i) < e/sqrt (real (DIM('a)))"
+        unfolding e'_def by (auto simp: dist_real_def)
+      then have "(dist (x $$ i) (y $$ i))\<twosuperior> < (e/sqrt (real (DIM('a))))\<twosuperior>"
+        by (rule power_strict_mono) auto
+      then show "(dist (x $$ i) (y $$ i))\<twosuperior> < e\<twosuperior> / real DIM('a)"
+        by (simp add: power_divide)
+    qed auto
+    also have "\<dots> = e" using `0 < e` by (simp add: real_eq_of_nat DIM_positive)
+    finally have "dist x y < e" . }
+  with a b show ?thesis
+    apply (rule_tac exI[of _ "Chi a"])
+    apply (rule_tac exI[of _ "Chi b"])
+    using eucl_less[where 'a='a] by auto
+qed
+
+lemma ex_rat_list:
+  fixes x :: "'a\<Colon>ordered_euclidean_space"
+  assumes "\<And> i. x $$ i \<in> \<rat>"
+  shows "\<exists> r. length r = DIM('a) \<and> (\<forall> i < DIM('a). of_rat (r ! i) = x $$ i)"
+proof -
+  have "\<forall>i. \<exists>r. x $$ i = of_rat r" using assms unfolding Rats_def by blast
+  from choice[OF this] guess r ..
+  then show ?thesis by (auto intro!: exI[of _ "map r [0 ..< DIM('a)]"])
+qed
+
+lemma open_UNION:
+  fixes M :: "'a\<Colon>ordered_euclidean_space set"
+  assumes "open M"
+  shows "M = UNION {(a, b) | a b. {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)} \<subseteq> M}
+                   (\<lambda> (a, b). {Chi (of_rat \<circ> op ! a) <..< Chi (of_rat \<circ> op ! b)})"
+    (is "M = UNION ?idx ?box")
+proof safe
+  fix x assume "x \<in> M"
+  obtain e where e: "e > 0" "ball x e \<subseteq> M"
+    using openE[OF assms `x \<in> M`] by auto
+  then obtain a b where ab: "x \<in> {a <..< b}" "\<And>i. a $$ i \<in> \<rat>" "\<And>i. b $$ i \<in> \<rat>" "{a <..< b} \<subseteq> ball x e"
+    using rational_boxes[OF e(1)] by blast
+  then obtain p q where pq: "length p = DIM ('a)"
+                            "length q = DIM ('a)"
+                            "\<forall> i < DIM ('a). of_rat (p ! i) = a $$ i \<and> of_rat (q ! i) = b $$ i"
+    using ex_rat_list[OF ab(2)] ex_rat_list[OF ab(3)] by blast
+  hence p: "Chi (of_rat \<circ> op ! p) = a"
+    using euclidean_eq[of "Chi (of_rat \<circ> op ! p)" a]
+    unfolding o_def by auto
+  from pq have q: "Chi (of_rat \<circ> op ! q) = b"
+    using euclidean_eq[of "Chi (of_rat \<circ> op ! q)" b]
+    unfolding o_def by auto
+  have "x \<in> ?box (p, q)"
+    using p q ab by auto
+  thus "x \<in> UNION ?idx ?box" using ab e p q exI[of _ p] exI[of _ q] by auto
+qed auto
+
+lemma halfspace_span_open:
+  "sigma_sets UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))
+    \<subseteq> sets borel"
+  by (auto intro!: borel.sigma_sets_subset[simplified] borel_open
+                   open_halfspace_component_lt)
+
+lemma halfspace_lt_in_halfspace:
+  "{x\<Colon>'a. x $$ i < a} \<in> sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>)"
+  by (auto intro!: sigma_sets.Basic simp: sets_sigma)
+
+lemma halfspace_gt_in_halfspace:
+  "{x\<Colon>'a. a < x $$ i} \<in> sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>)"
+  (is "?set \<in> sets ?SIGMA")
+proof -
+  interpret sigma_algebra "?SIGMA"
+    by (intro sigma_algebra_sigma_sets) (simp_all add: sets_sigma)
+  have *: "?set = (\<Union>n. space ?SIGMA - {x\<Colon>'a. x $$ i < a + 1 / real (Suc n)})"
+  proof (safe, simp_all add: not_less)
+    fix x assume "a < x $$ i"
+    with reals_Archimedean[of "x $$ i - a"]
+    obtain n where "a + 1 / real (Suc n) < x $$ i"
+      by (auto simp: inverse_eq_divide field_simps)
+    then show "\<exists>n. a + 1 / real (Suc n) \<le> x $$ i"
+      by (blast intro: less_imp_le)
+  next
+    fix x n
+    have "a < a + 1 / real (Suc n)" by auto
+    also assume "\<dots> \<le> x"
+    finally show "a < x" .
+  qed
+  show "?set \<in> sets ?SIGMA" unfolding *
+    by (safe intro!: countable_UN Diff halfspace_lt_in_halfspace)
+qed
+
+lemma open_span_halfspace:
+  "sets borel \<subseteq> sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i < a})\<rparr>)"
+    (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) simp
+  then interpret sigma_algebra ?SIGMA .
+  { fix S :: "'a set" assume "S \<in> open" then have "open S" unfolding mem_def .
+    from open_UNION[OF this]
+    obtain I where *: "S =
+      (\<Union>(a, b)\<in>I.
+          (\<Inter> i<DIM('a). {x. (Chi (real_of_rat \<circ> op ! a)::'a) $$ i < x $$ i}) \<inter>
+          (\<Inter> i<DIM('a). {x. x $$ i < (Chi (real_of_rat \<circ> op ! b)::'a) $$ i}))"
+      unfolding greaterThanLessThan_def
+      unfolding eucl_greaterThan_eq_halfspaces[where 'a='a]
+      unfolding eucl_lessThan_eq_halfspaces[where 'a='a]
+      by blast
+    have "S \<in> sets ?SIGMA"
+      unfolding *
+      by (auto intro!: countable_UN Int countable_INT halfspace_lt_in_halfspace halfspace_gt_in_halfspace) }
+  then show ?thesis unfolding borel_def
+    by (intro sets_sigma_subset) auto
+qed
+
+lemma halfspace_span_halfspace_le:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. x $$ i \<le> a})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  { fix a i
+    have *: "{x::'a. x$$i < a} = (\<Union>n. {x. x$$i \<le> a - 1/real (Suc n)})"
+    proof (safe, simp_all)
+      fix x::'a assume *: "x$$i < a"
+      with reals_Archimedean[of "a - x$$i"]
+      obtain n where "x $$ i < a - 1 / (real (Suc n))"
+        by (auto simp: field_simps inverse_eq_divide)
+      then show "\<exists>n. x $$ i \<le> a - 1 / (real (Suc n))"
+        by (blast intro: less_imp_le)
+    next
+      fix x::'a and n
+      assume "x$$i \<le> a - 1 / real (Suc n)"
+      also have "\<dots> < a" by auto
+      finally show "x$$i < a" .
+    qed
+    have "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
+      by (safe intro!: countable_UN)
+         (auto simp: sets_sigma intro!: sigma_sets.Basic) }
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma halfspace_span_halfspace_ge:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. a \<le> x $$ i})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  { fix a i have *: "{x::'a. x$$i < a} = space ?SIGMA - {x::'a. a \<le> x$$i}" by auto
+    have "{x. x$$i < a} \<in> sets ?SIGMA" unfolding *
+      by (safe intro!: Diff)
+         (auto simp: sets_sigma intro!: sigma_sets.Basic) }
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma halfspace_le_span_halfspace_gt:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x. a < x $$ i})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  { fix a i have *: "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
+    have "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
+      by (safe intro!: Diff)
+         (auto simp: sets_sigma intro!: sigma_sets.Basic) }
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma halfspace_le_span_atMost:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  have "\<And>a i. {x. x$$i \<le> a} \<in> sets ?SIGMA"
+  proof cases
+    fix a i assume "i < DIM('a)"
+    then have *: "{x::'a. x$$i \<le> a} = (\<Union>k::nat. {.. (\<chi>\<chi> n. if n = i then a else real k)})"
+    proof (safe, simp_all add: eucl_le[where 'a='a] split: split_if_asm)
+      fix x
+      from real_arch_simple[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"] guess k::nat ..
+      then have "\<And>i. i < DIM('a) \<Longrightarrow> x$$i \<le> real k"
+        by (subst (asm) Max_le_iff) auto
+      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia \<le> real k"
+        by (auto intro!: exI[of _ k])
+    qed
+    show "{x. x$$i \<le> a} \<in> sets ?SIGMA" unfolding *
+      by (safe intro!: countable_UN)
+         (auto simp: sets_sigma intro!: sigma_sets.Basic)
+  next
+    fix a i assume "\<not> i < DIM('a)"
+    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+      using top by auto
+  qed
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma halfspace_le_span_greaterThan:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i \<le> a})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {a<..})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  have "\<And>a i. {x. x$$i \<le> a} \<in> sets ?SIGMA"
+  proof cases
+    fix a i assume "i < DIM('a)"
+    have "{x::'a. x$$i \<le> a} = space ?SIGMA - {x::'a. a < x$$i}" by auto
+    also have *: "{x::'a. a < x$$i} = (\<Union>k::nat. {(\<chi>\<chi> n. if n = i then a else -real k) <..})" using `i <DIM('a)`
+    proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
+      fix x
+      from real_arch_lt[of "Max ((\<lambda>i. -x$$i)`{..<DIM('a)})"]
+      guess k::nat .. note k = this
+      { fix i assume "i < DIM('a)"
+        then have "-x$$i < real k"
+          using k by (subst (asm) Max_less_iff) auto
+        then have "- real k < x$$i" by simp }
+      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> -real k < x $$ ia"
+        by (auto intro!: exI[of _ k])
+    qed
+    finally show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+      apply (simp only:)
+      apply (safe intro!: countable_UN Diff)
+      by (auto simp: sets_sigma intro!: sigma_sets.Basic)
+  next
+    fix a i assume "\<not> i < DIM('a)"
+    then show "{x. x$$i \<le> a} \<in> sets ?SIGMA"
+      using top by auto
+  qed
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma halfspace_le_span_lessThan:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..<a})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  have "\<And>a i. {x. a \<le> x$$i} \<in> sets ?SIGMA"
+  proof cases
+    fix a i assume "i < DIM('a)"
+    have "{x::'a. a \<le> x$$i} = space ?SIGMA - {x::'a. x$$i < a}" by auto
+    also have *: "{x::'a. x$$i < a} = (\<Union>k::nat. {..< (\<chi>\<chi> n. if n = i then a else real k)})" using `i <DIM('a)`
+    proof (safe, simp_all add: eucl_less[where 'a='a] split: split_if_asm)
+      fix x
+      from real_arch_lt[of "Max ((\<lambda>i. x$$i)`{..<DIM('a)})"]
+      guess k::nat .. note k = this
+      { fix i assume "i < DIM('a)"
+        then have "x$$i < real k"
+          using k by (subst (asm) Max_less_iff) auto
+        then have "x$$i < real k" by simp }
+      then show "\<exists>k::nat. \<forall>ia. ia \<noteq> i \<longrightarrow> ia < DIM('a) \<longrightarrow> x $$ ia < real k"
+        by (auto intro!: exI[of _ k])
+    qed
+    finally show "{x. a \<le> x$$i} \<in> sets ?SIGMA"
+      apply (simp only:)
+      apply (safe intro!: countable_UN Diff)
+      by (auto simp: sets_sigma intro!: sigma_sets.Basic)
+  next
+    fix a i assume "\<not> i < DIM('a)"
+    then show "{x. a \<le> x$$i} \<in> sets ?SIGMA"
+      using top by auto
+  qed
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma atMost_span_atLeastAtMost:
+  "sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space})\<rparr>) \<subseteq>
+   sets (sigma \<lparr>space=UNIV, sets=range (\<lambda>(a,b). {a..b})\<rparr>)"
+  (is "_ \<subseteq> sets ?SIGMA")
+proof -
+  have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+  then interpret sigma_algebra ?SIGMA .
+  { fix a::'a
+    have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
+    proof (safe, simp_all add: eucl_le[where 'a='a])
+      fix x
+      from real_arch_simple[of "Max ((\<lambda>i. - x$$i)`{..<DIM('a)})"]
+      guess k::nat .. note k = this
+      { fix i assume "i < DIM('a)"
+        with k have "- x$$i \<le> real k"
+          by (subst (asm) Max_le_iff) (auto simp: field_simps)
+        then have "- real k \<le> x$$i" by simp }
+      then show "\<exists>n::nat. \<forall>i<DIM('a). - real n \<le> x $$ i"
+        by (auto intro!: exI[of _ k])
+    qed
+    have "{..a} \<in> sets ?SIGMA" unfolding *
+      by (safe intro!: countable_UN)
+         (auto simp: sets_sigma intro!: sigma_sets.Basic) }
+  then show ?thesis by (intro sets_sigma_subset) auto
+qed
+
+lemma algebra_eqI: assumes "sets A = sets (B::'a algebra)" "space A = space B"
+  shows "A = B" using assms by auto
+
+lemma borel_eq_atMost:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})\<rparr>)"
+    (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_atMost halfspace_span_halfspace_le open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_atLeastAtMost:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space, b). {a .. b})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using atMost_span_atLeastAtMost halfspace_le_span_atMost
+      halfspace_span_halfspace_le open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_greaterThan:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {a <..})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_greaterThan
+      halfspace_span_halfspace_le open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_lessThan:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a :: 'a\<Colon>ordered_euclidean_space). {..< a})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_lessThan
+      halfspace_span_halfspace_ge open_span_halfspace
+    by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_greaterThanLessThan:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})\<rparr>)"
+    (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+  show "sets borel \<subseteq> sets ?SIGMA"
+  proof -
+    have "sigma_algebra ?SIGMA" by (rule sigma_algebra_sigma) auto
+    then interpret sigma_algebra ?SIGMA .
+    { fix M :: "'a set" assume "M \<in> open"
+      then have "open M" by (simp add: mem_def)
+      have "M \<in> sets ?SIGMA"
+        apply (subst open_UNION[OF `open M`])
+        apply (safe intro!: countable_UN)
+        by (auto simp add: sigma_def intro!: sigma_sets.Basic) }
+    then show ?thesis
+      unfolding borel_def by (intro sets_sigma_subset) auto
+  qed
+qed auto
+
+lemma borel_eq_halfspace_le:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i \<le> a})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using open_span_halfspace halfspace_span_halfspace_le by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_halfspace_less:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x$$i < a})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using open_span_halfspace .
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_halfspace_gt:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a < x$$i})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using halfspace_le_span_halfspace_gt open_span_halfspace halfspace_span_halfspace_le by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma borel_eq_halfspace_ge:
+  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. a \<le> x$$i})\<rparr>)"
+   (is "_ = ?SIGMA")
+proof (rule algebra_eqI, rule antisym)
+  show "sets borel \<subseteq> sets ?SIGMA"
+    using halfspace_span_halfspace_ge open_span_halfspace by auto
+  show "sets ?SIGMA \<subseteq> sets borel"
+    by (rule borel.sets_sigma_subset) auto
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_halfspacesI:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  assumes "borel = (sigma \<lparr>space=UNIV, sets=range F\<rparr>)"
+  and "\<And>a i. S a i = f -` F (a,i) \<inter> space M"
+  and "\<And>a i. \<not> i < DIM('c) \<Longrightarrow> S a i \<in> sets M"
+  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a::real. S a i \<in> sets M)"
+proof safe
+  fix a :: real and i assume i: "i < DIM('c)" and f: "f \<in> borel_measurable M"
+  then show "S a i \<in> sets M" unfolding assms
+    by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1) sigma_def)
+next
+  assume a: "\<forall>i<DIM('c). \<forall>a. S a i \<in> sets M"
+  { fix a i have "S a i \<in> sets M"
+    proof cases
+      assume "i < DIM('c)"
+      with a show ?thesis unfolding assms(2) by simp
+    next
+      assume "\<not> i < DIM('c)"
+      from assms(3)[OF this] show ?thesis .
+    qed }
+  then have "f \<in> measurable M (sigma \<lparr>space=UNIV, sets=range F\<rparr>)"
+    by (auto intro!: measurable_sigma simp: assms(2))
+  then show "f \<in> borel_measurable M" unfolding measurable_def
+    unfolding assms(1) by simp
+qed
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_le:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i \<le> a} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_le]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_less:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. f w $$ i < a} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_less]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_ge:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M = (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a \<le> f w $$ i} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_ge]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_halfspace_greater:
+  fixes f :: "'a \<Rightarrow> 'c\<Colon>ordered_euclidean_space"
+  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). \<forall>a. {w \<in> space M. a < f w $$ i} \<in> sets M)"
+  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_gt]) auto
+
+lemma (in sigma_algebra) borel_measurable_iff_le:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
+  using borel_measurable_iff_halfspace_le[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_less:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
+  using borel_measurable_iff_halfspace_less[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_ge:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
+  using borel_measurable_iff_halfspace_ge[where 'c=real] by simp
+
+lemma (in sigma_algebra) borel_measurable_iff_greater:
+  "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
+  using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
+
+lemma borel_measureable_euclidean_component:
+  "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel"
+  unfolding borel_def[where 'a=real]
+proof (rule borel.measurable_sigma, simp_all)
+  fix S::"real set" assume "S \<in> open" then have "open S" unfolding mem_def .
+  from open_vimage_euclidean_component[OF this]
+  show "(\<lambda>x. x $$ i) -` S \<in> sets borel"
+    by (auto intro: borel_open)
+qed
+
+lemma (in sigma_algebra) borel_measureable_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_measureable_euclidean_component)
+next
+  assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
+  then show "f \<in> borel_measurable M"
+    unfolding borel_measurable_iff_halfspace_le by auto
+qed
+
+subsection "Borel measurable operators"
+
+lemma (in sigma_algebra) affine_borel_measurable_vector:
+  fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
+proof (rule borel_measurableI)
+  fix S :: "'x set" assume "open S"
+  show "(\<lambda>x. a + b *\<^sub>R f x) -` S \<inter> space M \<in> sets M"
+  proof cases
+    assume "b \<noteq> 0"
+    with `open S` have "((\<lambda>x. (- a + x) /\<^sub>R b) ` S) \<in> open" (is "?S \<in> open")
+      by (auto intro!: open_affinity simp: scaleR.add_right mem_def)
+    hence "?S \<in> sets borel"
+      unfolding borel_def by (auto simp: sigma_def intro!: sigma_sets.Basic)
+    moreover
+    from `b \<noteq> 0` have "(\<lambda>x. a + b *\<^sub>R f x) -` S = f -` ?S"
+      apply auto by (rule_tac x="a + b *\<^sub>R f x" in image_eqI, simp_all)
+    ultimately show ?thesis using assms unfolding in_borel_measurable_borel
+      by auto
+  qed simp
+qed
+
+lemma (in sigma_algebra) affine_borel_measurable:
+  fixes g :: "'a \<Rightarrow> real"
+  assumes g: "g \<in> borel_measurable M"
+  shows "(\<lambda>x. a + (g x) * b) \<in> borel_measurable M"
+  using affine_borel_measurable_vector[OF assms] by (simp add: mult_commute)
+
+lemma (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) 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 (in sigma_algebra) borel_measurable_inverse[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes "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)
+qed
+
+lemma (in sigma_algebra) borel_measurable_divide[simp, intro]:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes "f \<in> borel_measurable M"
+  and "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x / g x) \<in> borel_measurable M"
+  unfolding field_divide_inverse
+  by (rule borel_measurable_inverse borel_measurable_times assms)+
+
+lemma (in sigma_algebra) borel_measurable_max[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> real"
+  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
+
+lemma (in sigma_algebra) 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
+
+lemma (in sigma_algebra) borel_measurable_abs[simp, intro]:
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
+proof -
+  have *: "\<And>x. \<bar>f x\<bar> = max 0 (f x) + max 0 (- f x)" by (simp add: max_def)
+  show ?thesis unfolding * using assms by auto
+qed
+
+section "Borel space over the real line with infinity"
+
+lemma borel_Real_measurable:
+  "A \<in> sets borel \<Longrightarrow> Real -` A \<in> sets borel"
+proof (rule borel_measurable_translate)
+  fix B :: "pinfreal set" assume "open B"
+  then obtain T x where T: "open T" "Real ` (T \<inter> {0..}) = B - {\<omega>}" and
+    x: "\<omega> \<in> B \<Longrightarrow> 0 \<le> x" "\<omega> \<in> B \<Longrightarrow> {Real x <..} \<subseteq> B"
+    unfolding open_pinfreal_def by blast
+
+  have "Real -` B = Real -` (B - {\<omega>})" by auto
+  also have "\<dots> = Real -` (Real ` (T \<inter> {0..}))" using T by simp
+  also have "\<dots> = (if 0 \<in> T then T \<union> {.. 0} else T \<inter> {0..})"
+    apply (auto simp add: Real_eq_Real image_iff)
+    apply (rule_tac x="max 0 x" in bexI)
+    by (auto simp: max_def)
+  finally show "Real -` B \<in> sets borel"
+    using `open T` by auto
+qed simp
+
+lemma borel_real_measurable:
+  "A \<in> sets borel \<Longrightarrow> (real -` A :: pinfreal set) \<in> sets borel"
+proof (rule borel_measurable_translate)
+  fix B :: "real set" assume "open B"
+  { fix x have "0 < real x \<longleftrightarrow> (\<exists>r>0. x = Real r)" by (cases x) auto }
+  note Ex_less_real = this
+  have *: "real -` B = (if 0 \<in> B then real -` (B \<inter> {0 <..}) \<union> {0, \<omega>} else real -` (B \<inter> {0 <..}))"
+    by (force simp: Ex_less_real)
+
+  have "open (real -` (B \<inter> {0 <..}) :: pinfreal set)"
+    unfolding open_pinfreal_def using `open B`
+    by (auto intro!: open_Int exI[of _ "B \<inter> {0 <..}"] simp: image_iff Ex_less_real)
+  then show "(real -` B :: pinfreal set) \<in> sets borel" unfolding * by auto
+qed simp
+
+lemma (in sigma_algebra) borel_measurable_Real[intro, simp]:
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+  unfolding in_borel_measurable_borel
+proof safe
+  fix S :: "pinfreal set" assume "S \<in> sets borel"
+  from borel_Real_measurable[OF this]
+  have "(Real \<circ> f) -` S \<inter> space M \<in> sets M"
+    using assms
+    unfolding vimage_compose in_borel_measurable_borel
+    by auto
+  thus "(\<lambda>x. Real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
+qed
+
+lemma (in sigma_algebra) borel_measurable_real[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  shows "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+  unfolding in_borel_measurable_borel
+proof safe
+  fix S :: "real set" assume "S \<in> sets borel"
+  from borel_real_measurable[OF this]
+  have "(real \<circ> f) -` S \<inter> space M \<in> sets M"
+    using assms
+    unfolding vimage_compose in_borel_measurable_borel
+    by auto
+  thus "(\<lambda>x. real (f x)) -` S \<inter> space M \<in> sets M" by (simp add: comp_def)
+qed
+
+lemma (in sigma_algebra) borel_measurable_Real_eq:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
+  shows "(\<lambda>x. Real (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
+proof
+  have [simp]: "(\<lambda>x. Real (f x)) -` {\<omega>} \<inter> space M = {}"
+    by auto
+  assume "(\<lambda>x. Real (f x)) \<in> borel_measurable M"
+  hence "(\<lambda>x. real (Real (f x))) \<in> borel_measurable M"
+    by (rule borel_measurable_real)
+  moreover have "\<And>x. x \<in> space M \<Longrightarrow> real (Real (f x)) = f x"
+    using assms by auto
+  ultimately show "f \<in> borel_measurable M"
+    by (simp cong: measurable_cong)
+qed auto
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq_real:
+  "f \<in> borel_measurable M \<longleftrightarrow>
+    ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<omega>} \<inter> space M \<in> sets M)"
+proof safe
+  assume "f \<in> borel_measurable M"
+  then show "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
+    by (auto intro: borel_measurable_vimage borel_measurable_real)
+next
+  assume *: "(\<lambda>x. real (f x)) \<in> borel_measurable M" "f -` {\<omega>} \<inter> space M \<in> sets M"
+  have "f -` {\<omega>} \<inter> space M = {x\<in>space M. f x = \<omega>}" by auto
+  with * have **: "{x\<in>space M. f x = \<omega>} \<in> sets M" by simp
+  have f: "f = (\<lambda>x. if f x = \<omega> then \<omega> else Real (real (f x)))"
+    by (simp add: fun_eq_iff Real_real)
+  show "f \<in> borel_measurable M"
+    apply (subst f)
+    apply (rule measurable_If)
+    using * ** by auto
+qed
+
+lemma (in sigma_algebra) less_eq_ge_measurable:
+  fixes f :: "'a \<Rightarrow> 'c::linorder"
+  shows "{x\<in>space M. a < f x} \<in> sets M \<longleftrightarrow> {x\<in>space M. f x \<le> a} \<in> sets M"
+proof
+  assume "{x\<in>space M. f x \<le> a} \<in> sets M"
+  moreover have "{x\<in>space M. a < f x} = space M - {x\<in>space M. f x \<le> a}" by auto
+  ultimately show "{x\<in>space M. a < f x} \<in> sets M" by auto
+next
+  assume "{x\<in>space M. a < f x} \<in> sets M"
+  moreover have "{x\<in>space M. f x \<le> a} = space M - {x\<in>space M. a < f x}" by auto
+  ultimately show "{x\<in>space M. f x \<le> a} \<in> sets M" by auto
+qed
+
+lemma (in sigma_algebra) greater_eq_le_measurable:
+  fixes f :: "'a \<Rightarrow> 'c::linorder"
+  shows "{x\<in>space M. f x < a} \<in> sets M \<longleftrightarrow> {x\<in>space M. a \<le> f x} \<in> sets M"
+proof
+  assume "{x\<in>space M. a \<le> f x} \<in> sets M"
+  moreover have "{x\<in>space M. f x < a} = space M - {x\<in>space M. a \<le> f x}" by auto
+  ultimately show "{x\<in>space M. f x < a} \<in> sets M" by auto
+next
+  assume "{x\<in>space M. f x < a} \<in> sets M"
+  moreover have "{x\<in>space M. a \<le> f x} = space M - {x\<in>space M. f x < a}" by auto
+  ultimately show "{x\<in>space M. a \<le> f x} \<in> sets M" by auto
+qed
+
+lemma (in sigma_algebra) less_eq_le_pinfreal_measurable:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  shows "(\<forall>a. {x\<in>space M. a < f x} \<in> sets M) \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+proof
+  assume a: "\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M"
+  show "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
+  proof
+    fix a show "{x \<in> space M. a < f x} \<in> sets M"
+    proof (cases a)
+      case (preal r)
+      have "{x\<in>space M. a < f x} = (\<Union>i. {x\<in>space M. a + inverse (of_nat (Suc i)) \<le> f x})"
+      proof safe
+        fix x assume "a < f x" and [simp]: "x \<in> space M"
+        with ex_pinfreal_inverse_of_nat_Suc_less[of "f x - a"]
+        obtain n where "a + inverse (of_nat (Suc n)) < f x"
+          by (cases "f x", auto simp: pinfreal_minus_order)
+        then have "a + inverse (of_nat (Suc n)) \<le> f x" by simp
+        then show "x \<in> (\<Union>i. {x \<in> space M. a + inverse (of_nat (Suc i)) \<le> f x})"
+          by auto
+      next
+        fix i x assume [simp]: "x \<in> space M"
+        have "a < a + inverse (of_nat (Suc i))" using preal by auto
+        also assume "a + inverse (of_nat (Suc i)) \<le> f x"
+        finally show "a < f x" .
+      qed
+      with a show ?thesis by auto
+    qed simp
+  qed
+next
+  assume a': "\<forall>a. {x \<in> space M. a < f x} \<in> sets M"
+  then have a: "\<forall>a. {x \<in> space M. f x \<le> a} \<in> sets M" unfolding less_eq_ge_measurable .
+  show "\<forall>a. {x \<in> space M. a \<le> f x} \<in> sets M" unfolding greater_eq_le_measurable[symmetric]
+  proof
+    fix a show "{x \<in> space M. f x < a} \<in> sets M"
+    proof (cases a)
+      case (preal r)
+      show ?thesis
+      proof cases
+        assume "a = 0" then show ?thesis by simp
+      next
+        assume "a \<noteq> 0"
+        have "{x\<in>space M. f x < a} = (\<Union>i. {x\<in>space M. f x \<le> a - inverse (of_nat (Suc i))})"
+        proof safe
+          fix x assume "f x < a" and [simp]: "x \<in> space M"
+          with ex_pinfreal_inverse_of_nat_Suc_less[of "a - f x"]
+          obtain n where "inverse (of_nat (Suc n)) < a - f x"
+            using preal by (cases "f x") auto
+          then have "f x \<le> a - inverse (of_nat (Suc n)) "
+            using preal by (cases "f x") (auto split: split_if_asm)
+          then show "x \<in> (\<Union>i. {x \<in> space M. f x \<le> a - inverse (of_nat (Suc i))})"
+            by auto
+        next
+          fix i x assume [simp]: "x \<in> space M"
+          assume "f x \<le> a - inverse (of_nat (Suc i))"
+          also have "\<dots> < a" using `a \<noteq> 0` preal by auto
+          finally show "f x < a" .
+        qed
+        with a show ?thesis by auto
+      qed
+    next
+      case infinite
+      have "f -` {\<omega>} \<inter> space M = (\<Inter>n. {x\<in>space M. of_nat n < f x})"
+      proof (safe, simp_all, safe)
+        fix x assume *: "\<forall>n::nat. Real (real n) < f x"
+        show "f x = \<omega>"    proof (rule ccontr)
+          assume "f x \<noteq> \<omega>"
+          with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n"
+            by (auto simp: pinfreal_noteq_omega_Ex)
+          with *[THEN spec, of n] show False by auto
+        qed
+      qed
+      with a' have \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" by auto
+      moreover have "{x \<in> space M. f x < a} = space M - f -` {\<omega>} \<inter> space M"
+        using infinite by auto
+      ultimately show ?thesis by auto
+    qed
+  qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_greater:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a < f x} \<in> sets M)"
+proof safe
+  fix a assume f: "f \<in> borel_measurable M"
+  have "{x\<in>space M. a < f x} = f -` {a <..} \<inter> space M" by auto
+  with f show "{x\<in>space M. a < f x} \<in> sets M"
+    by (auto intro!: measurable_sets)
+next
+  assume *: "\<forall>a. {x\<in>space M. a < f x} \<in> sets M"
+  hence **: "\<forall>a. {x\<in>space M. f x < a} \<in> sets M"
+    unfolding less_eq_le_pinfreal_measurable
+    unfolding greater_eq_le_measurable .
+
+  show "f \<in> borel_measurable M" unfolding borel_measurable_pinfreal_eq_real borel_measurable_iff_greater
+  proof safe
+    have "f -` {\<omega>} \<inter> space M = space M - {x\<in>space M. f x < \<omega>}" by auto
+    then show \<omega>: "f -` {\<omega>} \<inter> space M \<in> sets M" using ** by auto
+
+    fix a
+    have "{w \<in> space M. a < real (f w)} =
+      (if 0 \<le> a then {w\<in>space M. Real a < f w} - (f -` {\<omega>} \<inter> space M) else space M)"
+    proof (split split_if, safe del: notI)
+      fix x assume "0 \<le> a"
+      { assume "a < real (f x)" then show "Real a < f x" "x \<notin> f -` {\<omega>} \<inter> space M"
+          using `0 \<le> a` by (cases "f x", auto) }
+      { assume "Real a < f x" "x \<notin> f -` {\<omega>}" then show "a < real (f x)"
+          using `0 \<le> a` by (cases "f x", auto) }
+    next
+      fix x assume "\<not> 0 \<le> a" then show "a < real (f x)" by (cases "f x") auto
+    qed
+    then show "{w \<in> space M. a < real (f w)} \<in> sets M"
+      using \<omega> * by (auto intro!: Diff)
+  qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_less:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x < a} \<in> sets M)"
+  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable greater_eq_le_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_le:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. f x \<le> a} \<in> sets M)"
+  using borel_measurable_pinfreal_iff_greater unfolding less_eq_ge_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_iff_ge:
+  "(f::'a \<Rightarrow> pinfreal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. {x\<in>space M. a \<le> f x} \<in> sets M)"
+  using borel_measurable_pinfreal_iff_greater unfolding less_eq_le_pinfreal_measurable .
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq_const:
+  fixes f :: "'a \<Rightarrow> pinfreal" 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 (in sigma_algebra) borel_measurable_pinfreal_neq_const:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  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 (in sigma_algebra) borel_measurable_pinfreal_less[intro,simp]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  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 "(\<lambda>x. real (f x)) \<in> borel_measurable M"
+    "(\<lambda>x. real (g x)) \<in> borel_measurable M"
+    using assms by (auto intro!: borel_measurable_real)
+  from borel_measurable_less[OF this]
+  have "{x \<in> space M. real (f x) < real (g x)} \<in> sets M" .
+  moreover have "{x \<in> space M. f x \<noteq> \<omega>} \<in> sets M" using f by (rule borel_measurable_pinfreal_neq_const)
+  moreover have "{x \<in> space M. g x = \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_eq_const)
+  moreover have "{x \<in> space M. g x \<noteq> \<omega>} \<in> sets M" using g by (rule borel_measurable_pinfreal_neq_const)
+  moreover have "{x \<in> space M. f x < g x} = ({x \<in> space M. g x = \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>}) \<union>
+    ({x \<in> space M. g x \<noteq> \<omega>} \<inter> {x \<in> space M. f x \<noteq> \<omega>} \<inter> {x \<in> space M. real (f x) < real (g x)})"
+    by (auto simp: real_of_pinfreal_strict_mono_iff)
+  ultimately show ?thesis by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_le[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  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"
+proof -
+  have "{x \<in> space M. f x \<le> g x} = space M - {x \<in> space M. g x < f x}" by auto
+  then show ?thesis using g f by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_eq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  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 (in sigma_algebra) borel_measurable_pinfreal_neq[intro,simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  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 (in sigma_algebra) borel_measurable_pinfreal_add[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  assumes measure: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
+proof -
+  have *: "(\<lambda>x. f x + g x) =
+     (\<lambda>x. if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else Real (real (f x) + real (g x)))"
+     by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex)
+  show ?thesis using assms unfolding *
+    by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_times[intro, simp]:
+  fixes f :: "'a \<Rightarrow> pinfreal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
+proof -
+  have *: "(\<lambda>x. f x * g x) =
+     (\<lambda>x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \<omega> then \<omega> else if g x = \<omega> then \<omega> else
+      Real (real (f x) * real (g x)))"
+     by (auto simp: fun_eq_iff pinfreal_noteq_omega_Ex)
+  show ?thesis using assms unfolding *
+    by (auto intro!: measurable_If)
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_setsum[simp, intro]:
+  fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  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 add: borel_measurable_const)
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_min[intro, simp]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  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 (in sigma_algebra) borel_measurable_pinfreal_max[intro]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  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 (in sigma_algebra) borel_measurable_SUP[simp, intro]:
+  fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(SUP i : A. f i) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
+  unfolding borel_measurable_pinfreal_iff_greater
+proof safe
+  fix a
+  have "{x\<in>space M. a < ?sup x} = (\<Union>i\<in>A. {x\<in>space M. a < f i x})"
+    by (auto simp: less_Sup_iff SUPR_def[where 'a=pinfreal] SUPR_fun_expand[where 'c=pinfreal])
+  then show "{x\<in>space M. a < ?sup x} \<in> sets M"
+    using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
+  fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
+  shows "(INF i : A. f i) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
+  unfolding borel_measurable_pinfreal_iff_less
+proof safe
+  fix a
+  have "{x\<in>space M. ?inf x < a} = (\<Union>i\<in>A. {x\<in>space M. f i x < a})"
+    by (auto simp: Inf_less_iff INFI_def[where 'a=pinfreal] INFI_fun_expand)
+  then show "{x\<in>space M. ?inf x < a} \<in> sets M"
+    using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_pinfreal_diff[simp, intro]:
+  fixes f g :: "'a \<Rightarrow> pinfreal"
+  assumes "f \<in> borel_measurable M"
+  assumes "g \<in> borel_measurable M"
+  shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
+  unfolding borel_measurable_pinfreal_iff_greater
+proof safe
+  fix a
+  have "{x \<in> space M. a < f x - g x} = {x \<in> space M. g x + a < f x}"
+    by (simp add: pinfreal_less_minus_iff)
+  then show "{x \<in> space M. a < f x - g x} \<in> sets M"
+    using assms by auto
+qed
+
+lemma (in sigma_algebra) borel_measurable_psuminf:
+  assumes "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
+  using assms unfolding psuminf_def
+  by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
+
+section "LIMSEQ is borel measurable"
+
+lemma (in sigma_algebra) borel_measurable_LIMSEQ:
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+  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 -
+  let "?pu x i" = "max (u i x) 0"
+  let "?nu x i" = "max (- u i x) 0"
+
+  { fix x assume x: "x \<in> space M"
+    have "(?pu x) ----> max (u' x) 0"
+      "(?nu x) ----> max (- u' x) 0"
+      using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
+    from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
+    have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
+      "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
+      by (simp_all add: Real_max'[symmetric]) }
+  note eq = this
+
+  have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
+    by auto
+
+  have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
+       "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
+    using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
+  with eq[THEN measurable_cong, of M "\<lambda>x. x" borel]
+  have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
+       "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
+    unfolding SUPR_fun_expand INFI_fun_expand by auto
+  note this[THEN borel_measurable_real]
+  from borel_measurable_diff[OF this]
+  show ?thesis unfolding * .
+qed
+
+end
--- a/src/HOL/Probability/Caratheodory.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Caratheodory.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -760,7 +760,7 @@
 
 theorem (in algebra) caratheodory:
   assumes posf: "positive f" and ca: "countably_additive M f"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> pinfreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma (space M) (sets M)) \<mu>"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> pinfreal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and> measure_space (sigma M) \<mu>"
   proof -
     have inc: "increasing M f"
       by (metis additive_increasing ca countably_additive_additive posf)
@@ -778,7 +778,7 @@
     hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
       using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
       by simp
-    have "measure_space (sigma (space M) (sets M)) ?infm"
+    have "measure_space (sigma M) ?infm"
       unfolding sigma_def
       by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
          (simp_all add: sgs_sb space_closed)
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Complete_Measure.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -0,0 +1,317 @@
+(*  Title:      Complete_Measure.thy
+    Author:     Robert Himmelmann, Johannes Hoelzl, TU Muenchen
+*)
+theory Complete_Measure
+imports Product_Measure
+begin
+
+locale completeable_measure_space = measure_space
+
+definition (in completeable_measure_space) completion :: "'a algebra" where
+  "completion = \<lparr> space = space M,
+    sets = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N' } \<rparr>"
+
+lemma (in completeable_measure_space) space_completion[simp]:
+  "space completion = space M" unfolding completion_def by simp
+
+lemma (in completeable_measure_space) sets_completionE:
+  assumes "A \<in> sets completion"
+  obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets" "S \<in> sets M"
+  using assms unfolding completion_def by auto
+
+lemma (in completeable_measure_space) sets_completionI:
+  assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets" "S \<in> sets M"
+  shows "A \<in> sets completion"
+  using assms unfolding completion_def by auto
+
+lemma (in completeable_measure_space) sets_completionI_sets[intro]:
+  "A \<in> sets M \<Longrightarrow> A \<in> sets completion"
+  unfolding completion_def by force
+
+lemma (in completeable_measure_space) null_sets_completion:
+  assumes "N' \<in> null_sets" "N \<subseteq> N'" shows "N \<in> sets completion"
+  apply(rule sets_completionI[of N "{}" N N'])
+  using assms by auto
+
+sublocale completeable_measure_space \<subseteq> completion!: sigma_algebra completion
+proof (unfold sigma_algebra_iff2, safe)
+  fix A x assume "A \<in> sets completion" "x \<in> A"
+  with sets_into_space show "x \<in> space completion"
+    by (auto elim!: sets_completionE)
+next
+  fix A assume "A \<in> sets completion"
+  from this[THEN sets_completionE] guess S N N' . note A = this
+  let ?C = "space completion"
+  show "?C - A \<in> sets completion" using A
+    by (intro sets_completionI[of _ "(?C - S) \<inter> (?C - N')" "(?C - S) \<inter> N' \<inter> (?C - N)"])
+       auto
+next
+  fix A ::"nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion"
+  then have "\<forall>n. \<exists>S N N'. A n = S \<union> N \<and> S \<in> sets M \<and> N' \<in> null_sets \<and> N \<subseteq> N'"
+    unfolding completion_def by (auto simp: image_subset_iff)
+  from choice[OF this] guess S ..
+  from choice[OF this] guess N ..
+  from choice[OF this] guess N' ..
+  then show "UNION UNIV A \<in> sets completion"
+    using null_sets_UN[of N']
+    by (intro sets_completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"])
+       auto
+qed auto
+
+definition (in completeable_measure_space)
+  "split_completion A p = (\<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and>
+    fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets)"
+
+definition (in completeable_measure_space)
+  "main_part A = fst (Eps (split_completion A))"
+
+definition (in completeable_measure_space)
+  "null_part A = snd (Eps (split_completion A))"
+
+lemma (in completeable_measure_space) split_completion:
+  assumes "A \<in> sets completion"
+  shows "split_completion A (main_part A, null_part A)"
+  unfolding main_part_def null_part_def
+proof (rule someI2_ex)
+  from assms[THEN sets_completionE] guess S N N' . note A = this
+  let ?P = "(S, N - S)"
+  show "\<exists>p. split_completion A p"
+    unfolding split_completion_def using A
+  proof (intro exI conjI)
+    show "A = fst ?P \<union> snd ?P" using A by auto
+    show "snd ?P \<subseteq> N'" using A by auto
+  qed auto
+qed auto
+
+lemma (in completeable_measure_space)
+  assumes "S \<in> sets completion"
+  shows main_part_sets[intro, simp]: "main_part S \<in> sets M"
+    and main_part_null_part_Un[simp]: "main_part S \<union> null_part S = S"
+    and main_part_null_part_Int[simp]: "main_part S \<inter> null_part S = {}"
+  using split_completion[OF assms] by (auto simp: split_completion_def)
+
+lemma (in completeable_measure_space) null_part:
+  assumes "S \<in> sets completion" shows "\<exists>N. N\<in>null_sets \<and> null_part S \<subseteq> N"
+  using split_completion[OF assms] by (auto simp: split_completion_def)
+
+lemma (in completeable_measure_space) null_part_sets[intro, simp]:
+  assumes "S \<in> sets M" shows "null_part S \<in> sets M" "\<mu> (null_part S) = 0"
+proof -
+  have S: "S \<in> sets completion" using assms by auto
+  have "S - main_part S \<in> sets M" using assms by auto
+  moreover
+  from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S]
+  have "S - main_part S = null_part S" by auto
+  ultimately show sets: "null_part S \<in> sets M" by auto
+  from null_part[OF S] guess N ..
+  with measure_eq_0[of N "null_part S"] sets
+  show "\<mu> (null_part S) = 0" by auto
+qed
+
+definition (in completeable_measure_space) "\<mu>' A = \<mu> (main_part A)"
+
+lemma (in completeable_measure_space) \<mu>'_set[simp]:
+  assumes "S \<in> sets M" shows "\<mu>' S = \<mu> S"
+proof -
+  have S: "S \<in> sets completion" using assms by auto
+  then have "\<mu> S = \<mu> (main_part S \<union> null_part S)" by simp
+  also have "\<dots> = \<mu> (main_part S)"
+    using S assms measure_additive[of "main_part S" "null_part S"]
+    by (auto simp: measure_additive)
+  finally show ?thesis unfolding \<mu>'_def by simp
+qed
+
+lemma (in completeable_measure_space) sets_completionI_sub:
+  assumes N: "N' \<in> null_sets" "N \<subseteq> N'"
+  shows "N \<in> sets completion"
+  using assms by (intro sets_completionI[of _ "{}" N N']) auto
+
+lemma (in completeable_measure_space) \<mu>_main_part_UN:
+  fixes S :: "nat \<Rightarrow> 'a set"
+  assumes "range S \<subseteq> sets completion"
+  shows "\<mu>' (\<Union>i. (S i)) = \<mu> (\<Union>i. main_part (S i))"
+proof -
+  have S: "\<And>i. S i \<in> sets completion" using assms by auto
+  then have UN: "(\<Union>i. S i) \<in> sets completion" by auto
+  have "\<forall>i. \<exists>N. N \<in> null_sets \<and> null_part (S i) \<subseteq> N"
+    using null_part[OF S] by auto
+  from choice[OF this] guess N .. note N = this
+  then have UN_N: "(\<Union>i. N i) \<in> null_sets" by (intro null_sets_UN) auto
+  have "(\<Union>i. S i) \<in> sets completion" using S by auto
+  from null_part[OF this] guess N' .. note N' = this
+  let ?N = "(\<Union>i. N i) \<union> N'"
+  have null_set: "?N \<in> null_sets" using N' UN_N by (intro null_sets_Un) auto
+  have "main_part (\<Union>i. S i) \<union> ?N = (main_part (\<Union>i. S i) \<union> null_part (\<Union>i. S i)) \<union> ?N"
+    using N' by auto
+  also have "\<dots> = (\<Union>i. main_part (S i) \<union> null_part (S i)) \<union> ?N"
+    unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto
+  also have "\<dots> = (\<Union>i. main_part (S i)) \<union> ?N"
+    using N by auto
+  finally have *: "main_part (\<Union>i. S i) \<union> ?N = (\<Union>i. main_part (S i)) \<union> ?N" .
+  have "\<mu> (main_part (\<Union>i. S i)) = \<mu> (main_part (\<Union>i. S i) \<union> ?N)"
+    using null_set UN by (intro measure_Un_null_set[symmetric]) auto
+  also have "\<dots> = \<mu> ((\<Union>i. main_part (S i)) \<union> ?N)"
+    unfolding * ..
+  also have "\<dots> = \<mu> (\<Union>i. main_part (S i))"
+    using null_set S by (intro measure_Un_null_set) auto
+  finally show ?thesis unfolding \<mu>'_def .
+qed
+
+lemma (in completeable_measure_space) \<mu>_main_part_Un:
+  assumes S: "S \<in> sets completion" and T: "T \<in> sets completion"
+  shows "\<mu>' (S \<union> T) = \<mu> (main_part S \<union> main_part T)"
+proof -
+  have UN: "(\<Union>i. binary (main_part S) (main_part T) i) = (\<Union>i. main_part (binary S T i))"
+    unfolding binary_def by (auto split: split_if_asm)
+  show ?thesis
+    using \<mu>_main_part_UN[of "binary S T"] assms
+    unfolding range_binary_eq Un_range_binary UN by auto
+qed
+
+sublocale completeable_measure_space \<subseteq> completion!: measure_space completion \<mu>'
+proof
+  show "\<mu>' {} = 0" by auto
+next
+  show "countably_additive completion \<mu>'"
+  proof (unfold countably_additive_def, intro allI conjI impI)
+    fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets completion" "disjoint_family A"
+    have "disjoint_family (\<lambda>i. main_part (A i))"
+    proof (intro disjoint_family_on_bisimulation[OF A(2)])
+      fix n m assume "A n \<inter> A m = {}"
+      then have "(main_part (A n) \<union> null_part (A n)) \<inter> (main_part (A m) \<union> null_part (A m)) = {}"
+        using A by (subst (1 2) main_part_null_part_Un) auto
+      then show "main_part (A n) \<inter> main_part (A m) = {}" by auto
+    qed
+    then have "(\<Sum>\<^isub>\<infinity>n. \<mu>' (A n)) = \<mu> (\<Union>i. main_part (A i))"
+      unfolding \<mu>'_def using A by (intro measure_countably_additive) auto
+    then show "(\<Sum>\<^isub>\<infinity>n. \<mu>' (A n)) = \<mu>' (UNION UNIV A)"
+      unfolding \<mu>_main_part_UN[OF A(1)] .
+  qed
+qed
+
+lemma (in sigma_algebra) simple_functionD':
+  assumes "simple_function f"
+  shows "f -` {x} \<inter> space M \<in> sets M"
+proof cases
+  assume "x \<in> f`space M" from simple_functionD(2)[OF assms this] show ?thesis .
+next
+  assume "x \<notin> f`space M" then have "f -` {x} \<inter> space M = {}" by auto
+  then show ?thesis by auto
+qed
+
+lemma (in sigma_algebra) simple_function_If:
+  assumes sf: "simple_function f" "simple_function g" and A: "A \<in> sets M"
+  shows "simple_function (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function ?IF")
+proof -
+  def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
+  show ?thesis unfolding simple_function_def
+  proof safe
+    have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto
+    from finite_subset[OF this] assms
+    show "finite (?IF ` space M)" unfolding simple_function_def by auto
+  next
+    fix x assume "x \<in> space M"
+    then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
+      then ((F (f x) \<inter> A) \<union> (G (f x) - (G (f x) \<inter> A)))
+      else ((F (g x) \<inter> A) \<union> (G (g x) - (G (g x) \<inter> A))))"
+      using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def)
+    have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M"
+      unfolding F_def G_def using sf[THEN simple_functionD'] by auto
+    show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
+  qed
+qed
+
+lemma (in measure_space) null_sets_finite_UN:
+  assumes "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> null_sets"
+  shows "(\<Union>i\<in>S. A i) \<in> null_sets"
+proof (intro CollectI conjI)
+  show "(\<Union>i\<in>S. A i) \<in> sets M" using assms by (intro finite_UN) auto
+  have "\<mu> (\<Union>i\<in>S. A i) \<le> (\<Sum>i\<in>S. \<mu> (A i))"
+    using assms by (intro measure_finitely_subadditive) auto
+  then show "\<mu> (\<Union>i\<in>S. A i) = 0"
+    using assms by auto
+qed
+
+lemma (in completeable_measure_space) completion_ex_simple_function:
+  assumes f: "completion.simple_function f"
+  shows "\<exists>f'. simple_function f' \<and> (AE x. f x = f' x)"
+proof -
+  let "?F x" = "f -` {x} \<inter> space M"
+  have F: "\<And>x. ?F x \<in> sets completion" and fin: "finite (f`space M)"
+    using completion.simple_functionD'[OF f]
+      completion.simple_functionD[OF f] by simp_all
+  have "\<forall>x. \<exists>N. N \<in> null_sets \<and> null_part (?F x) \<subseteq> N"
+    using F null_part by auto
+  from choice[OF this] obtain N where
+    N: "\<And>x. null_part (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets" by auto
+  let ?N = "\<Union>x\<in>f`space M. N x" let "?f' x" = "if x \<in> ?N then undefined else f x"
+  have sets: "?N \<in> null_sets" using N fin by (intro null_sets_finite_UN) auto
+  show ?thesis unfolding simple_function_def
+  proof (safe intro!: exI[of _ ?f'])
+    have "?f' ` space M \<subseteq> f`space M \<union> {undefined}" by auto
+    from finite_subset[OF this] completion.simple_functionD(1)[OF f]
+    show "finite (?f' ` space M)" by auto
+  next
+    fix x assume "x \<in> space M"
+    have "?f' -` {?f' x} \<inter> space M =
+      (if x \<in> ?N then ?F undefined \<union> ?N
+       else if f x = undefined then ?F (f x) \<union> ?N
+       else ?F (f x) - ?N)"
+      using N(2) sets_into_space by (auto split: split_if_asm)
+    moreover { fix y have "?F y \<union> ?N \<in> sets M"
+      proof cases
+        assume y: "y \<in> f`space M"
+        have "?F y \<union> ?N = (main_part (?F y) \<union> null_part (?F y)) \<union> ?N"
+          using main_part_null_part_Un[OF F] by auto
+        also have "\<dots> = main_part (?F y) \<union> ?N"
+          using y N by auto
+        finally show ?thesis
+          using F sets by auto
+      next
+        assume "y \<notin> f`space M" then have "?F y = {}" by auto
+        then show ?thesis using sets by auto
+      qed }
+    moreover {
+      have "?F (f x) - ?N = main_part (?F (f x)) \<union> null_part (?F (f x)) - ?N"
+        using main_part_null_part_Un[OF F] by auto
+      also have "\<dots> = main_part (?F (f x)) - ?N"
+        using N `x \<in> space M` by auto
+      finally have "?F (f x) - ?N \<in> sets M"
+        using F sets by auto }
+    ultimately show "?f' -` {?f' x} \<inter> space M \<in> sets M" by auto
+  next
+    show "AE x. f x = ?f' x"
+      by (rule AE_I', rule sets) auto
+  qed
+qed
+
+lemma (in completeable_measure_space) completion_ex_borel_measurable:
+  fixes g :: "'a \<Rightarrow> pinfreal"
+  assumes g: "g \<in> borel_measurable completion"
+  shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
+proof -
+  from g[THEN completion.borel_measurable_implies_simple_function_sequence]
+  obtain f where "\<And>i. completion.simple_function (f i)" "f \<up> g" by auto
+  then have "\<forall>i. \<exists>f'. simple_function f' \<and> (AE x. f i x = f' x)"
+    using completion_ex_simple_function by auto
+  from this[THEN choice] obtain f' where
+    sf: "\<And>i. simple_function (f' i)" and
+    AE: "\<forall>i. AE x. f i x = f' i x" by auto
+  show ?thesis
+  proof (intro bexI)
+    from AE[unfolded all_AE_countable]
+    show "AE x. g x = (SUP i. f' i) x" (is "AE x. g x = ?f x")
+    proof (rule AE_mp, safe intro!: AE_cong)
+      fix x assume eq: "\<forall>i. f i x = f' i x"
+      have "g x = (SUP i. f i x)"
+        using `f \<up> g` unfolding isoton_def SUPR_fun_expand by auto
+      then show "g x = ?f x"
+        using eq unfolding SUPR_fun_expand by auto
+    qed
+    show "?f \<in> borel_measurable M"
+      using sf by (auto intro!: borel_measurable_SUP
+        intro: borel_measurable_simple_function)
+  qed
+qed
+
+end
--- a/src/HOL/Probability/Euclidean_Lebesgue.thy	Wed Dec 01 06:50:54 2010 -0800
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,212 +0,0 @@
-theory Euclidean_Lebesgue
-  imports Lebesgue_Integration Lebesgue_Measure
-begin
-
-lemma simple_function_has_integral:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
-  assumes f:"lebesgue.simple_function f"
-  and f':"\<forall>x. f x \<noteq> \<omega>"
-  and om:"\<forall>x\<in>range f. lmeasure (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
-  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
-  unfolding lebesgue.simple_integral_def
-  apply(subst lebesgue_simple_function_indicator[OF f])
-proof- case goal1
-  have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
-    "\<forall>x\<in>range f. x * lmeasure (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
-    using f' om unfolding indicator_def by auto
-  show ?case unfolding space_lebesgue_space real_of_pinfreal_setsum'[OF *(1),THEN sym]
-    unfolding real_of_pinfreal_setsum'[OF *(2),THEN sym]
-    unfolding real_of_pinfreal_setsum space_lebesgue_space
-    apply(rule has_integral_setsum)
-  proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD)
-    fix y::'a show "((\<lambda>x. real (f y * indicator (f -` {f y}) x)) has_integral
-      real (f y * lmeasure (f -` {f y} \<inter> UNIV))) UNIV"
-    proof(cases "f y = 0") case False
-      have mea:"gmeasurable (f -` {f y})" apply(rule glmeasurable_finite)
-        using assms unfolding lebesgue.simple_function_def using False by auto
-      have *:"\<And>x. real (indicator (f -` {f y}) x::pinfreal) = (if x \<in> f -` {f y} then 1 else 0)" by auto
-      show ?thesis unfolding real_of_pinfreal_mult[THEN sym]
-        apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def])
-        unfolding Int_UNIV_right lmeasure_gmeasure[OF mea,THEN sym]
-        unfolding measure_integral_univ[OF mea] * apply(rule integrable_integral)
-        unfolding gmeasurable_integrable[THEN sym] using mea .
-    qed auto
-  qed qed
-
-lemma (in measure_space) positive_integral_omega:
-  assumes "f \<in> borel_measurable M"
-  and "positive_integral f \<noteq> \<omega>"
-  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
-proof -
-  have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = positive_integral (\<lambda>x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x)"
-    using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp
-  also have "\<dots> \<le> positive_integral f"
-    by (auto intro!: positive_integral_mono simp: indicator_def)
-  finally show ?thesis
-    using assms(2) by (cases ?thesis) auto
-qed
-
-lemma (in measure_space) simple_integral_omega:
-  assumes "simple_function f"
-  and "simple_integral f \<noteq> \<omega>"
-  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
-proof (rule positive_integral_omega)
-  show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
-  show "positive_integral f \<noteq> \<omega>"
-    using assms by (simp add: positive_integral_eq_simple_integral)
-qed
-
-lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
-  unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
-  using assms by auto
-
-lemma simple_function_has_integral':
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
-  assumes f:"lebesgue.simple_function f"
-  and i: "lebesgue.simple_integral f \<noteq> \<omega>"
-  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
-proof- let ?f = "\<lambda>x. if f x = \<omega> then 0 else f x"
-  { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this
-  have **:"{x. f x \<noteq> ?f x} = f -` {\<omega>}" by auto
-  have **:"lmeasure {x\<in>space lebesgue_space. f x \<noteq> ?f x} = 0"
-    using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**)
-  show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **])
-    apply(rule lebesgue.simple_function_compose1[OF f])
-    unfolding * defer apply(rule simple_function_has_integral)
-  proof-
-    show "lebesgue.simple_function ?f"
-      using lebesgue.simple_function_compose1[OF f] .
-    show "\<forall>x. ?f x \<noteq> \<omega>" by auto
-    show "\<forall>x\<in>range ?f. lmeasure (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
-    proof (safe, simp, safe, rule ccontr)
-      fix y assume "f y \<noteq> \<omega>" "f y \<noteq> 0"
-      hence "(\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y} = f -` {f y}"
-        by (auto split: split_if_asm)
-      moreover assume "lmeasure ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
-      ultimately have "lmeasure (f -` {f y}) = \<omega>" by simp
-      moreover
-      have "f y * lmeasure (f -` {f y}) \<noteq> \<omega>" using i f
-        unfolding lebesgue.simple_integral_def setsum_\<omega> lebesgue.simple_function_def
-        by auto
-      ultimately have "f y = 0" by (auto split: split_if_asm)
-      then show False using `f y \<noteq> 0` by simp
-    qed
-  qed
-qed
-
-lemma (in measure_space) positive_integral_monotone_convergence:
-  fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
-  assumes i: "\<And>i. f i \<in> borel_measurable M" and mono: "\<And>x. mono (\<lambda>n. f n x)"
-  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
-  shows "u \<in> borel_measurable M"
-  and "(\<lambda>i. positive_integral (f i)) ----> positive_integral u" (is ?ilim)
-proof -
-  from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u]
-  show ?ilim using mono lim i by auto
-  have "(SUP i. f i) = u" using mono lim SUP_Lim_pinfreal
-    unfolding fun_eq_iff SUPR_fun_expand mono_def by auto
-  moreover have "(SUP i. f i) \<in> borel_measurable M"
-    using i by (rule borel_measurable_SUP)
-  ultimately show "u \<in> borel_measurable M" by simp
-qed
-
-lemma positive_integral_has_integral:
-  fixes f::"'a::ordered_euclidean_space => pinfreal"
-  assumes f:"f \<in> borel_measurable lebesgue_space"
-  and int_om:"lebesgue.positive_integral f \<noteq> \<omega>"
-  and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
-  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.positive_integral f))) UNIV"
-proof- let ?i = "lebesgue.positive_integral f"
-  from lebesgue.borel_measurable_implies_simple_function_sequence[OF f]
-  guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2)
-  let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
-  have u_simple:"\<And>k. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)"
-    apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
-  have int_u_le:"\<And>k. lebesgue.simple_integral (u k) \<le> lebesgue.positive_integral f"
-    unfolding u_simple apply(rule lebesgue.positive_integral_mono)
-    using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
-  have u_int_om:"\<And>i. lebesgue.simple_integral (u i) \<noteq> \<omega>"
-  proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed
-
-  note u_int = simple_function_has_integral'[OF u(1) this]
-  have "(\<lambda>x. real (f x)) integrable_on UNIV \<and>
-    (\<lambda>k. gintegral UNIV (\<lambda>x. real (u k x))) ----> gintegral UNIV (\<lambda>x. real (f x))"
-    apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int)
-  proof safe case goal1 show ?case apply(rule real_of_pinfreal_mono) using u(2,3) by auto
-  next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym])
-      prefer 3 apply(subst Real_real') defer apply(subst Real_real')
-      using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto
-  next case goal3
-    show ?case apply(rule bounded_realI[where B="real (lebesgue.positive_integral f)"])
-      apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int)
-      unfolding integral_unique[OF u_int] defer apply(rule real_of_pinfreal_mono[OF _ int_u_le])
-      using u int_om by auto
-  qed note int = conjunctD2[OF this]
-
-  have "(\<lambda>i. lebesgue.simple_integral (u i)) ----> ?i" unfolding u_simple
-    apply(rule lebesgue.positive_integral_monotone_convergence(2))
-    apply(rule lebesgue.borel_measurable_simple_function[OF u(1)])
-    using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto
-  hence "(\<lambda>i. real (lebesgue.simple_integral (u i))) ----> real ?i" apply-
-    apply(subst lim_Real[THEN sym]) prefer 3
-    apply(subst Real_real') defer apply(subst Real_real')
-    using u f_om int_om u_int_om by auto
-  note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]]
-  show ?thesis unfolding * by(rule integrable_integral[OF int(1)])
-qed
-
-lemma lebesgue_integral_has_integral:
-  fixes f::"'a::ordered_euclidean_space => real"
-  assumes f:"lebesgue.integrable f"
-  shows "(f has_integral (lebesgue.integral f)) UNIV"
-proof- let ?n = "\<lambda>x. - min (f x) 0" and ?p = "\<lambda>x. max (f x) 0"
-  have *:"f = (\<lambda>x. ?p x - ?n x)" apply rule by auto
-  note f = lebesgue.integrableD[OF f]
-  show ?thesis unfolding lebesgue.integral_def apply(subst *)
-  proof(rule has_integral_sub) case goal1
-    have *:"\<forall>x. Real (f x) \<noteq> \<omega>" by auto
-    note lebesgue.borel_measurable_Real[OF f(1)]
-    from positive_integral_has_integral[OF this f(2) *]
-    show ?case unfolding real_Real_max .
-  next case goal2
-    have *:"\<forall>x. Real (- f x) \<noteq> \<omega>" by auto
-    note lebesgue.borel_measurable_uminus[OF f(1)]
-    note lebesgue.borel_measurable_Real[OF this]
-    from positive_integral_has_integral[OF this f(3) *]
-    show ?case unfolding real_Real_max minus_min_eq_max by auto
-  qed
-qed
-
-lemma lmeasurable_imp_borel[dest]: fixes s::"'a::ordered_euclidean_space set"
-  assumes "s \<in> sets borel_space" shows "lmeasurable s"
-proof- let ?S = "range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)})"
-  have *:"?S \<subseteq> sets lebesgue_space" by auto
-  have "s \<in> sigma_sets UNIV ?S" using assms
-    unfolding borel_space_eq_atLeastAtMost by (simp add: sigma_def)
-  thus ?thesis using lebesgue.sigma_subset[of ?S,unfolded sets_sigma,OF *]
-    by auto
-qed
-
-lemma lmeasurable_open[dest]:
-  assumes "open s" shows "lmeasurable s"
-proof- have "s \<in> sets borel_space" using assms by auto
-  thus ?thesis by auto qed
-
-lemma continuous_on_imp_borel_measurable:
-  fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
-  assumes "continuous_on UNIV f"
-  shows "f \<in> borel_measurable lebesgue_space"
-  apply(rule lebesgue.borel_measurableI)
-  unfolding lebesgue_measurable apply(rule lmeasurable_open)
-  using continuous_open_preimage[OF assms] unfolding vimage_def by auto
-
-
-lemma (in measure_space) integral_monotone_convergence_pos':
-  assumes i: "\<And>i. integrable (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 (f i)) ----> x"
-  shows "integrable u \<and> integral u = x"
-  using integral_monotone_convergence_pos[OF assms] by auto
-
-end
--- a/src/HOL/Probability/Information.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Information.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -1,5 +1,5 @@
 theory Information
-imports Probability_Space Product_Measure Convex Radon_Nikodym
+imports Probability_Space Convex Lebesgue_Measure
 begin
 
 lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
@@ -12,43 +12,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
 
-lemma real_of_pinfreal_inverse[simp]:
-  fixes X :: pinfreal
-  shows "real (inverse X) = 1 / real X"
-  by (cases X) (auto simp: inverse_eq_divide)
-
-lemma (in finite_prob_space) finite_product_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr>
-                     (joint_distribution X Y)"
-  (is "finite_prob_space ?S _")
-proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
-  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
-  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
-    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
-qed
-
-lemma (in finite_prob_space) finite_measure_space_prod:
-  assumes X: "finite_measure_space MX (distribution X)"
-  assumes Y: "finite_measure_space MY (distribution Y)"
-  shows "finite_measure_space (prod_measure_space MX MY) (joint_distribution X Y)"
-    (is "finite_measure_space ?M ?D")
-proof (intro finite_measure_spaceI)
-  interpret X: finite_measure_space MX "distribution X" by fact
-  interpret Y: finite_measure_space MY "distribution Y" by fact
-  note finite_measure_space.finite_prod_measure_space[OF X Y, simp]
-  show "finite (space ?M)" using X.finite_space Y.finite_space by auto
-  show "joint_distribution X Y {} = 0" by simp
-  show "sets ?M = Pow (space ?M)" by simp
-  { fix x show "?D (space ?M) \<noteq> \<omega>" by (rule distribution_finite) }
-  { fix A B assume "A \<subseteq> space ?M" "B \<subseteq> space ?M" "A \<inter> B = {}"
-    have *: "(\<lambda>t. (X t, Y t)) -` (A \<union> B) \<inter> space M =
-             (\<lambda>t. (X t, Y t)) -` A \<inter> space M \<union> (\<lambda>t. (X t, Y t)) -` B \<inter> space M"
-      by auto
-    show "?D (A \<union> B) = ?D A + ?D B" unfolding distribution_def *
-      apply (rule measure_additive[symmetric])
-      using `A \<inter> B = {}` by (auto simp: sets_eq_Pow) }
-qed
-
 section "Convex theory"
 
 lemma log_setsum:
@@ -148,82 +111,48 @@
 qed
 
 lemma split_pairs:
-  shows
-    "((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
+  "((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
 
 section "Information theory"
 
-locale finite_information_space = finite_prob_space +
+locale information_space = prob_space +
   fixes b :: real assumes b_gt_1: "1 < b"
 
-context finite_information_space
+context information_space
 begin
 
-lemma
-  assumes "0 \<le> A" and pos: "0 < A \<Longrightarrow> 0 < B" "0 < A \<Longrightarrow> 0 < C"
-  shows mult_log_mult: "A * log b (B * C) = A * log b B + A * log b C" (is "?mult")
-  and mult_log_divide: "A * log b (B / C) = A * log b B - A * log b C" (is "?div")
+text {* Introduce some simplification rules for logarithm of base @{term b}. *}
+
+lemma log_neg_const:
+  assumes "x \<le> 0"
+  shows "log b x = log b 0"
 proof -
-  have "?mult \<and> ?div"
-  proof (cases "A = 0")
-    case False
-    hence "0 < A" using `0 \<le> A` by auto
-      with pos[OF this] show "?mult \<and> ?div" using b_gt_1
-        by (auto simp: log_divide log_mult field_simps)
-  qed simp
-  thus ?mult and ?div by auto
+  { fix u :: real
+    have "x \<le> 0" by fact
+    also have "0 < exp u"
+      using exp_gt_zero .
+    finally have "exp u \<noteq> x"
+      by auto }
+  then show "log b x = log b 0"
+    by (simp add: log_def ln_def)
 qed
 
-ML {*
-
-  (* tactic to solve equations of the form @{term "W * log b (X / (Y * Z)) = W * log b X - W * log b (Y * Z)"}
-     where @{term W} is a joint distribution of @{term X}, @{term Y}, and @{term Z}. *)
-
-  val mult_log_intros = [@{thm mult_log_divide}, @{thm mult_log_mult}]
-  val intros = [@{thm divide_pos_pos}, @{thm mult_pos_pos}, @{thm real_pinfreal_nonneg},
-    @{thm real_distribution_divide_pos_pos},
-    @{thm real_distribution_mult_inverse_pos_pos},
-    @{thm real_distribution_mult_pos_pos}]
-
-  val distribution_gt_0_tac = (rtac @{thm distribution_mono_gt_0}
-    THEN' assume_tac
-    THEN' clarsimp_tac (clasimpset_of @{context} addsimps2 @{thms split_pairs}))
-
-  val distr_mult_log_eq_tac = REPEAT_ALL_NEW (CHANGED o TRY o
-    (resolve_tac (mult_log_intros @ intros)
-      ORELSE' distribution_gt_0_tac
-      ORELSE' clarsimp_tac (clasimpset_of @{context})))
-
-  fun instanciate_term thy redex intro =
-    let
-      val intro_concl = Thm.concl_of intro
+lemma log_mult_eq:
+  "log b (A * B) = (if 0 < A * B then log b \<bar>A\<bar> + log b \<bar>B\<bar> else log b 0)"
+  using log_mult[of b "\<bar>A\<bar>" "\<bar>B\<bar>"] b_gt_1 log_neg_const[of "A * B"]
+  by (auto simp: zero_less_mult_iff mult_le_0_iff)
 
-      val lhs = intro_concl |> HOLogic.dest_Trueprop |> HOLogic.dest_eq |> fst
-
-      val m = SOME (Pattern.match thy (lhs, redex) (Vartab.empty, Vartab.empty))
-        handle Pattern.MATCH => NONE
-
-    in
-      Option.map (fn m => Envir.subst_term m intro_concl) m
-    end
+lemma log_inverse_eq:
+  "log b (inverse B) = (if 0 < B then - log b B else log b 0)"
+  using log_inverse[of b B] log_neg_const[of "inverse B"] b_gt_1 by simp
 
-  fun mult_log_simproc simpset redex =
-  let
-    val ctxt = Simplifier.the_context simpset
-    val thy = ProofContext.theory_of ctxt
-    fun prove (SOME thm) = (SOME
-          (Goal.prove ctxt [] [] thm (K (distr_mult_log_eq_tac 1))
-           |> mk_meta_eq)
-            handle THM _ => NONE)
-      | prove NONE = NONE
-  in
-    get_first (instanciate_term thy (term_of redex) #> prove) mult_log_intros
-  end
-*}
+lemma log_divide_eq:
+  "log b (A / B) = (if 0 < A * B then log b \<bar>A\<bar> - log b \<bar>B\<bar> else log b 0)"
+  unfolding divide_inverse log_mult_eq log_inverse_eq abs_inverse
+  by (auto simp: zero_less_mult_iff mult_le_0_iff)
 
-simproc_setup mult_log ("real (distribution X x) * log b (A * B)" |
-                        "real (distribution X x) * log b (A / B)") = {* K mult_log_simproc *}
+lemmas log_simps = log_mult_eq log_inverse_eq log_divide_eq
 
 end
 
@@ -236,15 +165,49 @@
   "KL_divergence b M \<mu> \<nu> =
     measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))"
 
+lemma (in sigma_finite_measure) KL_divergence_cong:
+  assumes "measure_space M \<nu>"
+  and cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A" "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
+  shows "KL_divergence b M \<nu>' \<mu>' = KL_divergence b M \<nu> \<mu>"
+proof -
+  interpret \<nu>: measure_space M \<nu> by fact
+  show ?thesis
+    unfolding KL_divergence_def
+    using RN_deriv_cong[OF cong, of "\<lambda>A. A"]
+    by (simp add: cong \<nu>.integral_cong_measure[OF cong(2)])
+qed
+
+lemma (in sigma_finite_measure) KL_divergence_vimage:
+  assumes f: "bij_betw f S (space M)"
+  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
+  shows "KL_divergence b (vimage_algebra S f) (\<lambda>A. \<nu> (f ` A)) (\<lambda>A. \<mu> (f ` A)) = KL_divergence b M \<nu> \<mu>"
+    (is "KL_divergence b ?M ?\<nu> ?\<mu> = _")
+proof -
+  interpret \<nu>: measure_space M \<nu> by fact
+  interpret v: measure_space ?M ?\<nu>
+    using f by (rule \<nu>.measure_space_isomorphic)
+
+  let ?RN = "sigma_finite_measure.RN_deriv ?M ?\<mu> ?\<nu>"
+  from RN_deriv_vimage[OF f \<nu>]
+  have *: "\<nu>.almost_everywhere (\<lambda>x. ?RN (the_inv_into S f x) = RN_deriv \<nu> x)"
+    by (rule absolutely_continuous_AE[OF \<nu>])
+
+  show ?thesis
+    unfolding KL_divergence_def \<nu>.integral_vimage_inv[OF f]
+    apply (rule \<nu>.integral_cong_AE)
+    apply (rule \<nu>.AE_mp[OF *])
+    apply (rule \<nu>.AE_cong)
+    apply simp
+    done
+qed
+
 lemma (in finite_measure_space) KL_divergence_eq_finite:
   assumes v: "finite_measure_space M \<nu>"
-  assumes ac: "\<forall>x\<in>space M. \<mu> {x} = 0 \<longrightarrow> \<nu> {x} = 0"
+  assumes ac: "absolutely_continuous \<nu>"
   shows "KL_divergence b M \<nu> \<mu> = (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x})))" (is "_ = ?sum")
 proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v])
   interpret v: finite_measure_space M \<nu> by fact
   have ms: "measure_space M \<nu>" by fact
-  have ac: "absolutely_continuous \<nu>"
-    using ac by (auto intro!: absolutely_continuousI[OF v])
   show "(\<Sum>x \<in> space M. log b (real (RN_deriv \<nu> x)) * real (\<nu> {x})) = ?sum"
     using RN_deriv_finite_measure[OF ms ac]
     by (auto intro!: setsum_cong simp: field_simps real_of_pinfreal_mult[symmetric])
@@ -252,32 +215,27 @@
 
 lemma (in finite_prob_space) KL_divergence_positive_finite:
   assumes v: "finite_prob_space M \<nu>"
-  assumes ac: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
+  assumes ac: "absolutely_continuous \<nu>"
   and "1 < b"
   shows "0 \<le> KL_divergence b M \<nu> \<mu>"
 proof -
   interpret v: finite_prob_space M \<nu> using v .
-
-  have *: "space M \<noteq> {}" using not_empty by simp
+  have ms: "finite_measure_space M \<nu>" by default
 
-  hence "- (KL_divergence b M \<nu> \<mu>) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
-  proof (subst KL_divergence_eq_finite)
-    show "finite_measure_space  M \<nu>" by fact
+  have "- (KL_divergence b M \<nu> \<mu>) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
+  proof (subst KL_divergence_eq_finite[OF ms ac], safe intro!: log_setsum_divide not_empty)
+    show "finite (space M)" using finite_space by simp
+    show "1 < b" by fact
+    show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1" using v.finite_sum_over_space_eq_1 by simp
 
-    show "\<forall>x\<in>space M. \<mu> {x} = 0 \<longrightarrow> \<nu> {x} = 0" using ac by auto
-    show "- (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {x}))) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {x}))"
-    proof (safe intro!: log_setsum_divide *)
-      show "finite (space M)" using finite_space by simp
-      show "1 < b" by fact
-      show "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1" using v.finite_sum_over_space_eq_1 by simp
-
-      fix x assume x: "x \<in> space M"
-      { assume "0 < real (\<nu> {x})"
-        hence "\<mu> {x} \<noteq> 0" using ac[OF x] by auto
-        thus "0 < prob {x}" using finite_measure[of "{x}"] sets_eq_Pow x
-          by (cases "\<mu> {x}") simp_all }
-    qed auto
-  qed
+    fix x assume "x \<in> space M"
+    then have x: "{x} \<in> sets M" unfolding sets_eq_Pow by auto
+    { assume "0 < real (\<nu> {x})"
+      then have "\<nu> {x} \<noteq> 0" by auto
+      then have "\<mu> {x} \<noteq> 0"
+        using ac[unfolded absolutely_continuous_def, THEN bspec, of "{x}"] x by auto
+      thus "0 < prob {x}" using finite_measure[of "{x}"] x by auto }
+  qed auto
   thus "0 \<le> KL_divergence b M \<nu> \<mu>" using finite_sum_over_space_eq_1 by simp
 qed
 
@@ -285,174 +243,175 @@
 
 definition (in prob_space)
   "mutual_information b S T X Y =
-    KL_divergence b (prod_measure_space S T)
+    KL_divergence b (sigma (pair_algebra S T))
       (joint_distribution X Y)
-      (prod_measure S (distribution X) T (distribution Y))"
+      (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y))"
 
-abbreviation (in finite_information_space)
-  finite_mutual_information ("\<I>'(_ ; _')") where
+definition (in prob_space)
+  "entropy b s X = mutual_information b s s X X"
+
+abbreviation (in information_space)
+  mutual_information_Pow ("\<I>'(_ ; _')") where
   "\<I>(X ; Y) \<equiv> mutual_information b
     \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr>
     \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y"
 
-lemma (in finite_information_space) mutual_information_generic_eq:
-  assumes MX: "finite_measure_space MX (distribution X)"
-  assumes MY: "finite_measure_space MY (distribution Y)"
-  shows "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
-      real (joint_distribution X Y {(x,y)}) *
-      log b (real (joint_distribution X Y {(x,y)}) /
-      (real (distribution X {x}) * real (distribution Y {y}))))"
+lemma (in information_space) mutual_information_commute_generic:
+  assumes X: "random_variable S X" and Y: "random_variable T Y"
+  assumes ac: "measure_space.absolutely_continuous (sigma (pair_algebra S T))
+   (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y)) (joint_distribution X Y)"
+  shows "mutual_information b S T X Y = mutual_information b T S Y X"
 proof -
-  let ?P = "prod_measure_space MX MY"
-  let ?\<mu> = "prod_measure MX (distribution X) MY (distribution Y)"
-  let ?\<nu> = "joint_distribution X Y"
-  interpret X: finite_measure_space MX "distribution X" by fact
-  moreover interpret Y: finite_measure_space MY "distribution Y" by fact
-  have fms: "finite_measure_space MX (distribution X)"
-            "finite_measure_space MY (distribution Y)" by fact+
-  have fms_P: "finite_measure_space ?P ?\<mu>"
-    by (rule X.finite_measure_space_finite_prod_measure) fact
-  then interpret P: finite_measure_space ?P ?\<mu> .
-  have fms_P': "finite_measure_space ?P ?\<nu>"
-      using finite_product_measure_space[of "space MX" "space MY"]
-        X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space]
-        X.sets_eq_Pow Y.sets_eq_Pow
-      by (simp add: prod_measure_space_def sigma_def)
-  then interpret P': finite_measure_space ?P ?\<nu> .
-  { fix x assume "x \<in> space ?P"
-    hence in_MX: "{fst x} \<in> sets MX" "{snd x} \<in> sets MY" using X.sets_eq_Pow Y.sets_eq_Pow
-      by (auto simp: prod_measure_space_def)
-    assume "?\<mu> {x} = 0"
-    with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX
-    have "distribution X {fst x} = 0 \<or> distribution Y {snd x} = 0"
-      by (simp add: prod_measure_space_def)
-    hence "joint_distribution X Y {x} = 0"
-      by (cases x) (auto simp: distribution_order) }
-  note measure_0 = this
+  interpret P: prob_space "sigma (pair_algebra S T)" "joint_distribution X Y"
+    using random_variable_pairI[OF X Y] by (rule distribution_prob_space)
+  interpret Q: prob_space "sigma (pair_algebra T S)" "joint_distribution Y X"
+    using random_variable_pairI[OF Y X] by (rule distribution_prob_space)
+  interpret X: prob_space S "distribution X" using X by (rule distribution_prob_space)
+  interpret Y: prob_space T "distribution Y" using Y by (rule distribution_prob_space)
+  interpret ST: pair_sigma_finite S "distribution X" T "distribution Y" by default
+  interpret TS: pair_sigma_finite T "distribution Y" S "distribution X" by default
+
+  have ST: "measure_space (sigma (pair_algebra S T)) (joint_distribution X Y)" by default
+  have TS: "measure_space (sigma (pair_algebra T S)) (joint_distribution Y X)" by default
+
+  have bij_ST: "bij_betw (\<lambda>(x, y). (y, x)) (space (sigma (pair_algebra S T))) (space (sigma (pair_algebra T S)))"
+    by (auto intro!: inj_onI simp: space_pair_algebra bij_betw_def)
+  have bij_TS: "bij_betw (\<lambda>(x, y). (y, x)) (space (sigma (pair_algebra T S))) (space (sigma (pair_algebra S T)))"
+    by (auto intro!: inj_onI simp: space_pair_algebra bij_betw_def)
+
+  { fix A
+    have "joint_distribution X Y ((\<lambda>(x, y). (y, x)) ` A) = joint_distribution Y X A"
+      unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>]) }
+  note jd_commute = this
+
+  { fix A assume A: "A \<in> sets (sigma (pair_algebra T S))"
+    have *: "\<And>x y. indicator ((\<lambda>(x, y). (y, x)) ` A) (x, y) = (indicator A (y, x) :: pinfreal)"
+      unfolding indicator_def by auto
+    have "ST.pair_measure ((\<lambda>(x, y). (y, x)) ` A) = TS.pair_measure A"
+      unfolding ST.pair_measure_def TS.pair_measure_def
+      using A by (auto simp add: TS.Fubini[symmetric] *) }
+  note pair_measure_commute = this
+
   show ?thesis
-    unfolding Let_def mutual_information_def
-    using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def
-    by (subst P.KL_divergence_eq_finite)
-       (auto simp add: prod_measure_space_def prod_measure_times_finite
-         finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
+    unfolding mutual_information_def
+    unfolding ST.KL_divergence_vimage[OF bij_TS ST ac, symmetric]
+    unfolding space_sigma space_pair_algebra jd_commute
+    unfolding ST.pair_sigma_algebra_swap[symmetric]
+    by (simp cong: TS.KL_divergence_cong[OF TS] add: pair_measure_commute)
 qed
 
-lemma (in finite_information_space)
-  assumes MX: "finite_prob_space MX (distribution X)"
-  assumes MY: "finite_prob_space MY (distribution Y)"
-  and X_space: "X ` space M \<subseteq> space MX" and Y_space: "Y ` space M \<subseteq> space MY"
-  shows mutual_information_eq_generic:
+lemma (in prob_space) finite_variables_absolutely_continuous:
+  assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y"
+  shows "measure_space.absolutely_continuous (sigma (pair_algebra S T))
+   (pair_sigma_finite.pair_measure S (distribution X) T (distribution Y)) (joint_distribution X Y)"
+proof -
+  interpret X: finite_prob_space S "distribution X" using X by (rule distribution_finite_prob_space)
+  interpret Y: finite_prob_space T "distribution Y" using Y by (rule distribution_finite_prob_space)
+  interpret XY: pair_finite_prob_space S "distribution X" T "distribution Y" by default
+  interpret P: finite_prob_space XY.P "joint_distribution X Y"
+    using assms by (intro joint_distribution_finite_prob_space)
+  show "XY.absolutely_continuous (joint_distribution X Y)"
+  proof (rule XY.absolutely_continuousI)
+    show "finite_measure_space XY.P (joint_distribution X Y)" by default
+    fix x assume "x \<in> space XY.P" and "XY.pair_measure {x} = 0"
+    then obtain a b where "(a, b) = x" and "a \<in> space S" "b \<in> space T"
+      and distr: "distribution X {a} * distribution Y {b} = 0"
+      by (cases x) (auto simp: pair_algebra_def)
+    with assms[THEN finite_random_variableD]
+      joint_distribution_Times_le_fst[of S X T Y "{a}" "{b}"]
+      joint_distribution_Times_le_snd[of S X T Y "{a}" "{b}"]
+    have "joint_distribution X Y {x} \<le> distribution Y {b}"
+         "joint_distribution X Y {x} \<le> distribution X {a}"
+      by auto
+    with distr show "joint_distribution X Y {x} = 0" by auto
+  qed
+qed
+
+lemma (in information_space) mutual_information_commute:
+  assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y"
+  shows "mutual_information b S T X Y = mutual_information b T S Y X"
+  by (intro finite_random_variableD X Y mutual_information_commute_generic finite_variables_absolutely_continuous)
+
+lemma (in information_space) mutual_information_commute_simple:
+  assumes X: "simple_function X" and Y: "simple_function Y"
+  shows "\<I>(X;Y) = \<I>(Y;X)"
+  by (intro X Y simple_function_imp_finite_random_variable mutual_information_commute)
+
+lemma (in information_space)
+  assumes MX: "finite_random_variable MX X"
+  assumes MY: "finite_random_variable MY Y"
+  shows mutual_information_generic_eq:
     "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
       real (joint_distribution X Y {(x,y)}) *
       log b (real (joint_distribution X Y {(x,y)}) /
       (real (distribution X {x}) * real (distribution Y {y}))))"
-    (is "?equality")
+    (is ?sum)
   and mutual_information_positive_generic:
-    "0 \<le> mutual_information b MX MY X Y" (is "?positive")
+     "0 \<le> mutual_information b MX MY X Y" (is ?positive)
 proof -
-  let ?P = "prod_measure_space MX MY"
-  let ?\<mu> = "prod_measure MX (distribution X) MY (distribution Y)"
-  let ?\<nu> = "joint_distribution X Y"
-
-  interpret X: finite_prob_space MX "distribution X" by fact
-  moreover interpret Y: finite_prob_space MY "distribution Y" by fact
-  have ms_X: "measure_space MX (distribution X)"
-    and ms_Y: "measure_space MY (distribution Y)"
-    and fms: "finite_measure_space MX (distribution X)" "finite_measure_space MY (distribution Y)" by fact+
-  have fms_P: "finite_measure_space ?P ?\<mu>"
-    by (rule X.finite_measure_space_finite_prod_measure) fact
-  then interpret P: finite_measure_space ?P ?\<mu> .
+  interpret X: finite_prob_space MX "distribution X" using MX by (rule distribution_finite_prob_space)
+  interpret Y: finite_prob_space MY "distribution Y" using MY by (rule distribution_finite_prob_space)
+  interpret XY: pair_finite_prob_space MX "distribution X" MY "distribution Y" by default
+  interpret P: finite_prob_space XY.P "joint_distribution X Y"
+    using assms by (intro joint_distribution_finite_prob_space)
 
-  have fms_P': "finite_measure_space ?P ?\<nu>"
-      using finite_product_measure_space[of "space MX" "space MY"]
-        X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space]
-        X.sets_eq_Pow Y.sets_eq_Pow
-      by (simp add: prod_measure_space_def sigma_def)
-  then interpret P': finite_measure_space ?P ?\<nu> .
-
-  { fix x assume "x \<in> space ?P"
-    hence in_MX: "{fst x} \<in> sets MX" "{snd x} \<in> sets MY" using X.sets_eq_Pow Y.sets_eq_Pow
-      by (auto simp: prod_measure_space_def)
+  have P_ms: "finite_measure_space XY.P (joint_distribution X Y)" by default
+  have P_ps: "finite_prob_space XY.P (joint_distribution X Y)" by default
 
-    assume "?\<mu> {x} = 0"
-    with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX
-    have "distribution X {fst x} = 0 \<or> distribution Y {snd x} = 0"
-      by (simp add: prod_measure_space_def)
-
-    hence "joint_distribution X Y {x} = 0"
-      by (cases x) (auto simp: distribution_order) }
-  note measure_0 = this
-
-  show ?equality
+  show ?sum
     unfolding Let_def mutual_information_def
-    using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def
-    by (subst P.KL_divergence_eq_finite)
-       (auto simp add: prod_measure_space_def prod_measure_times_finite
-         finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
+    by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]])
+       (auto simp add: pair_algebra_def setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
 
   show ?positive
-    unfolding Let_def mutual_information_def using measure_0 b_gt_1
-  proof (safe intro!: finite_prob_space.KL_divergence_positive_finite, simp_all)
-    have "?\<mu> (space ?P) = 1"
-      using X.top Y.top X.measure_space_1 Y.measure_space_1 fms
-      by (simp add: prod_measure_space_def X.finite_prod_measure_times)
-    with fms_P show "finite_prob_space ?P ?\<mu>"
-      by (simp add: finite_prob_space_eq)
-
-    from ms_X ms_Y X.top Y.top X.measure_space_1 Y.measure_space_1 Y.not_empty X_space Y_space
-    have "?\<nu> (space ?P) = 1" unfolding measure_space_1[symmetric]
-      by (auto intro!: arg_cong[where f="\<mu>"]
-               simp add: prod_measure_space_def distribution_def vimage_Times comp_def)
-    with fms_P' show "finite_prob_space ?P ?\<nu>"
-      by (simp add: finite_prob_space_eq)
-  qed
+    using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1]
+    unfolding mutual_information_def .
 qed
 
-lemma (in finite_information_space) mutual_information_eq:
-  "\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M.
+lemma (in information_space) mutual_information_eq:
+  assumes "simple_function X" "simple_function Y"
+  shows "\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M.
     real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) * log b (real (distribution (\<lambda>x. (X x, Y x)) {(x,y)}) /
                                                    (real (distribution X {x}) * real (distribution Y {y}))))"
-  by (subst mutual_information_eq_generic) (simp_all add: finite_prob_space_of_images)
+  using assms by (simp add: mutual_information_generic_eq)
 
-lemma (in finite_information_space) mutual_information_cong:
+lemma (in information_space) mutual_information_generic_cong:
   assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
   assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
-  shows "\<I>(X ; Y) = \<I>(X' ; Y')"
-proof -
-  have "X ` space M = X' ` space M" using X by (auto intro!: image_eqI)
-  moreover have "Y ` space M = Y' ` space M" using Y by (auto intro!: image_eqI)
-  ultimately show ?thesis
-  unfolding mutual_information_eq
-    using
-      assms[THEN distribution_cong]
-      joint_distribution_cong[OF assms]
-    by (auto intro!: setsum_cong)
-qed
+  shows "mutual_information b MX MY X Y = mutual_information b MX MY X' Y'"
+  unfolding mutual_information_def using X Y
+  by (simp cong: distribution_cong)
 
-lemma (in finite_information_space) mutual_information_positive: "0 \<le> \<I>(X;Y)"
-  by (subst mutual_information_positive_generic) (simp_all add: finite_prob_space_of_images)
+lemma (in information_space) mutual_information_cong:
+  assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
+  assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
+  shows "\<I>(X; Y) = \<I>(X'; Y')"
+  unfolding mutual_information_def using X Y
+  by (simp cong: distribution_cong image_cong)
+
+lemma (in information_space) mutual_information_positive:
+  assumes "simple_function X" "simple_function Y"
+  shows "0 \<le> \<I>(X;Y)"
+  using assms by (simp add: mutual_information_positive_generic)
 
 subsection {* Entropy *}
 
-definition (in prob_space)
-  "entropy b s X = mutual_information b s s X X"
-
-abbreviation (in finite_information_space)
-  finite_entropy ("\<H>'(_')") where
+abbreviation (in information_space)
+  entropy_Pow ("\<H>'(_')") where
   "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
 
-lemma (in finite_information_space) entropy_generic_eq:
-  assumes MX: "finite_measure_space MX (distribution X)"
+lemma (in information_space) entropy_generic_eq:
+  assumes MX: "finite_random_variable MX X"
   shows "entropy b MX X = -(\<Sum> x \<in> space MX. real (distribution X {x}) * log b (real (distribution X {x})))"
 proof -
+  interpret MX: finite_prob_space MX "distribution X" using MX by (rule distribution_finite_prob_space)
   let "?X x" = "real (distribution X {x})"
   let "?XX x y" = "real (joint_distribution X X {(x, y)})"
-  interpret MX: finite_measure_space MX "distribution X" by fact
   { fix x y
     have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
     then have "?XX x y * log b (?XX x y / (?X x * ?X y)) =
         (if x = y then - ?X y * log b (?X y) else 0)"
-      unfolding distribution_def by (auto simp: mult_log_divide) }
+      unfolding distribution_def by (auto simp: log_simps zero_less_mult_iff) }
   note remove_XX = this
   show ?thesis
     unfolding entropy_def mutual_information_generic_eq[OF MX MX]
@@ -460,201 +419,327 @@
     by (auto simp: setsum_cases MX.finite_space)
 qed
 
-lemma (in finite_information_space) entropy_eq:
-  "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
-  by (simp add: finite_measure_space entropy_generic_eq)
+lemma (in information_space) entropy_eq:
+  assumes "simple_function X"
+  shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
+  using assms by (simp add: entropy_generic_eq)
 
-lemma (in finite_information_space) entropy_positive: "0 \<le> \<H>(X)"
-  unfolding entropy_def using mutual_information_positive .
+lemma (in information_space) entropy_positive:
+  "simple_function X \<Longrightarrow> 0 \<le> \<H>(X)"
+  unfolding entropy_def by (simp add: mutual_information_positive)
 
-lemma (in finite_information_space) entropy_certainty_eq_0:
-  assumes "x \<in> X ` space M" and "distribution X {x} = 1"
+lemma (in information_space) entropy_certainty_eq_0:
+  assumes "simple_function X" and "x \<in> X ` space M" and "distribution X {x} = 1"
   shows "\<H>(X) = 0"
 proof -
   interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
-    by (rule finite_prob_space_of_images)
-
+    using simple_function_imp_finite_random_variable[OF `simple_function X`]
+    by (rule distribution_finite_prob_space)
   have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
     using X.measure_compl[of "{x}"] assms by auto
   also have "\<dots> = 0" using X.prob_space assms by auto
   finally have X0: "distribution X (X ` space M - {x}) = 0" by auto
-
   { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M"
     hence "{y} \<subseteq> X ` space M - {x}" by auto
     from X.measure_mono[OF this] X0 asm
     have "distribution X {y} = 0" by auto }
-
   hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)"
     using assms by auto
-
   have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp
-
-  show ?thesis unfolding entropy_eq by (auto simp: y fi)
+  show ?thesis unfolding entropy_eq[OF `simple_function X`] by (auto simp: y fi)
 qed
 
-lemma (in finite_information_space) entropy_le_card_not_0:
-  "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
+lemma (in information_space) entropy_le_card_not_0:
+  assumes "simple_function X"
+  shows "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
 proof -
   let "?d x" = "distribution X {x}"
   let "?p x" = "real (?d x)"
   have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
-    by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric])
+    by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function X`] setsum_negf[symmetric] log_simps not_less)
   also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
     apply (rule log_setsum')
-    using not_empty b_gt_1 finite_space sum_over_space_real_distribution
-    by auto
+    using not_empty b_gt_1 `simple_function X` sum_over_space_real_distribution
+    by (auto simp: simple_function_def)
   also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)"
-    apply (rule arg_cong[where f="\<lambda>f. log b (\<Sum>x\<in>X`space M. f x)"])
-    using distribution_finite[of X] by (auto simp: fun_eq_iff real_of_pinfreal_eq_0)
+    using distribution_finite[OF `simple_function X`[THEN simple_function_imp_random_variable], simplified]
+    by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) (auto simp: real_of_pinfreal_eq_0)
   finally show ?thesis
-    using finite_space by (auto simp: setsum_cases real_eq_of_nat)
+    using `simple_function X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
 qed
 
-lemma (in finite_information_space) entropy_uniform_max:
+lemma (in information_space) entropy_uniform_max:
+  assumes "simple_function X"
   assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
   shows "\<H>(X) = log b (real (card (X ` space M)))"
 proof -
-  note uniform =
-    finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified]
-
+  interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
+    using simple_function_imp_finite_random_variable[OF `simple_function X`]
+    by (rule distribution_finite_prob_space)
   have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
-    using finite_space not_empty by auto
-
+    using `simple_function X` not_empty by (auto simp: simple_function_def)
   { fix x assume "x \<in> X ` space M"
     hence "real (distribution X {x}) = 1 / real (card (X ` space M))"
-    proof (rule uniform)
+    proof (rule X.uniform_prob[simplified])
       fix x y assume "x \<in> X`space M" "y \<in> X`space M"
-      from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
+      from assms(2)[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
     qed }
   thus ?thesis
-    using not_empty finite_space b_gt_1 card_gt0
-    by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide)
+    using not_empty X.finite_space b_gt_1 card_gt0
+    by (simp add: entropy_eq[OF `simple_function X`] real_eq_of_nat[symmetric] log_simps)
 qed
 
-lemma (in finite_information_space) entropy_le_card:
-  "\<H>(X) \<le> log b (real (card (X ` space M)))"
+lemma (in information_space) entropy_le_card:
+  assumes "simple_function X"
+  shows "\<H>(X) \<le> log b (real (card (X ` space M)))"
 proof cases
   assume "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} = {}"
   then have "\<And>x. x\<in>X`space M \<Longrightarrow> distribution X {x} = 0" by auto
   moreover
   have "0 < card (X`space M)"
-    using finite_space not_empty unfolding card_gt_0_iff by auto
+    using `simple_function X` not_empty
+    by (auto simp: card_gt_0_iff simple_function_def)
   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 unfolding entropy_eq by simp
+  ultimately show ?thesis using assms by (simp add: entropy_eq)
 next
   assume False: "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} \<noteq> {}"
   have "card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}) \<le> card (X ` space M)"
-    (is "?A \<le> ?B") using finite_space not_empty by (auto intro!: card_mono)
-  note entropy_le_card_not_0
+    (is "?A \<le> ?B") using assms not_empty by (auto intro!: card_mono simp: simple_function_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 finite_space not_empty `?A \<le> ?B`
-    by (auto intro!: log_le simp: card_gt_0_iff)
+    using b_gt_1 False not_empty `?A \<le> ?B` assms
+    by (auto intro!: log_le simp: card_gt_0_iff simp: simple_function_def)
   finally show ?thesis .
 qed
 
-lemma (in finite_information_space) entropy_commute:
-  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"
+lemma (in information_space) entropy_commute:
+  assumes "simple_function X" "simple_function Y"
+  shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"
 proof -
+  have sf: "simple_function (\<lambda>x. (X x, Y x))" "simple_function (\<lambda>x. (Y x, X x))"
+    using assms by (auto intro: simple_function_Pair)
   have *: "(\<lambda>x. (Y x, X x))`space M = (\<lambda>(a,b). (b,a))`(\<lambda>x. (X x, Y x))`space M"
     by auto
   have inj: "\<And>X. inj_on (\<lambda>(a,b). (b,a)) X"
     by (auto intro!: inj_onI)
   show ?thesis
-    unfolding entropy_eq unfolding * setsum_reindex[OF inj]
+    unfolding sf[THEN entropy_eq] unfolding * setsum_reindex[OF inj]
     by (simp add: joint_distribution_commute[of Y X] split_beta)
 qed
 
-lemma (in finite_information_space) entropy_eq_cartesian_sum:
-  "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
+lemma (in information_space) entropy_eq_cartesian_product:
+  assumes "simple_function X" "simple_function Y"
+  shows "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
     real (joint_distribution X Y {(x,y)}) *
     log b (real (joint_distribution X Y {(x,y)})))"
 proof -
+  have sf: "simple_function (\<lambda>x. (X x, Y x))"
+    using assms by (auto intro: simple_function_Pair)
   { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M"
     then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto
     then have "joint_distribution X Y {x} = 0"
       unfolding distribution_def by auto }
-  then show ?thesis using finite_space
-    unfolding entropy_eq neg_equal_iff_equal setsum_cartesian_product
-    by (auto intro!: setsum_mono_zero_cong_left)
+  then show ?thesis using sf assms
+    unfolding entropy_eq[OF sf] neg_equal_iff_equal setsum_cartesian_product
+    by (auto intro!: setsum_mono_zero_cong_left simp: simple_function_def)
 qed
 
 subsection {* Conditional Mutual Information *}
 
 definition (in prob_space)
   "conditional_mutual_information b M1 M2 M3 X Y Z \<equiv>
-    mutual_information b M1 (prod_measure_space M2 M3) X (\<lambda>x. (Y x, Z x)) -
+    mutual_information b M1 (sigma (pair_algebra M2 M3)) X (\<lambda>x. (Y x, Z x)) -
     mutual_information b M1 M3 X Z"
 
-abbreviation (in finite_information_space)
-  finite_conditional_mutual_information ("\<I>'( _ ; _ | _ ')") where
+abbreviation (in information_space)
+  conditional_mutual_information_Pow ("\<I>'( _ ; _ | _ ')") where
   "\<I>(X ; Y | Z) \<equiv> conditional_mutual_information b
     \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr>
     \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr>
     \<lparr> space = Z`space M, sets = Pow (Z`space M) \<rparr>
     X Y Z"
 
-lemma (in finite_information_space) conditional_mutual_information_generic_eq:
-  assumes MX: "finite_measure_space MX (distribution X)"
-  assumes MY: "finite_measure_space MY (distribution Y)"
-  assumes MZ: "finite_measure_space MZ (distribution Z)"
-  shows "conditional_mutual_information b MX MY MZ X Y Z =
-    (\<Sum>(x, y, z)\<in>space MX \<times> space MY \<times> space MZ.
-      real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) *
-      log b (real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) /
-                   (real (distribution X {x}) * real (joint_distribution Y Z {(y, z)})))) -
-    (\<Sum>(x, y)\<in>space MX \<times> space MZ.
-      real (joint_distribution X Z {(x, y)}) *
-      log b (real (joint_distribution X Z {(x, y)}) / (real (distribution X {x}) * real (distribution Z {y}))))"
-  using assms finite_measure_space_prod[OF MY MZ]
-  unfolding conditional_mutual_information_def
-  by (subst (1 2) mutual_information_generic_eq)
-     (simp_all add: setsum_cartesian_product' finite_measure_space.finite_prod_measure_space)
 
-
-lemma (in finite_information_space) conditional_mutual_information_eq:
-  "\<I>(X ; Y | Z) = (\<Sum>(x, y, z) \<in> X ` space M \<times> Y ` space M \<times> Z ` space M.
+lemma (in information_space) conditional_mutual_information_generic_eq:
+  assumes MX: "finite_random_variable MX X"
+    and MY: "finite_random_variable MY Y"
+    and MZ: "finite_random_variable MZ Z"
+  shows "conditional_mutual_information b MX MY MZ X Y Z = (\<Sum>(x, y, z) \<in> space MX \<times> space MY \<times> space MZ.
              real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
              log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
     (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
-  by (subst conditional_mutual_information_generic_eq)
-     (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space
-      finite_measure_space finite_product_prob_space_of_images sigma_def
-      setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf
-      setsum_left_distrib[symmetric] setsum_real_distribution setsum_commute[where A="Y`space M"]
-      real_of_pinfreal_mult[symmetric]
-    cong: setsum_cong)
+  (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z)))")
+proof -
+  let ?YZ = "\<lambda>y z. real (joint_distribution Y Z {(y, z)})"
+  let ?X = "\<lambda>x. real (distribution X {x})"
+  let ?Z = "\<lambda>z. real (distribution Z {z})"
+
+  txt {* This proof is actually quiet easy, however we need to show that the
+    distributions are finite and the joint distributions are zero when one of
+    the variables distribution is also zero. *}
+
+  note finite_var = MX MY MZ
+  note random_var = finite_var[THEN finite_random_variableD]
+
+  note space_simps = space_pair_algebra space_sigma algebra.simps
+
+  note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
+  note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
+  note ZX = finite_random_variable_pairI[OF finite_var(3,1)]
+  note YZX = finite_random_variable_pairI[OF finite_var(2) ZX]
+  note order1 =
+    finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps]
+    finite_distribution_order(5,6)[OF finite_var(1,3), simplified space_simps]
+
+  note finite = finite_var(1) YZ finite_var(3) XZ YZX
+  note finite[THEN finite_distribution_finite, simplified space_simps, simp]
+
+  have order2: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
+          \<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
+    unfolding joint_distribution_commute_singleton[of X]
+    unfolding joint_distribution_assoc_singleton[symmetric]
+    using finite_distribution_order(6)[OF finite_var(2) ZX]
+    by (auto simp: space_simps)
 
-lemma (in finite_information_space) conditional_mutual_information_eq_mutual_information:
-  "\<I>(X ; Y) = \<I>(X ; Y | (\<lambda>x. ()))"
+  have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z))) =
+    (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * (log b (?XYZ x y z / (?X x * ?YZ y z)) - log b (?XZ x z / (?X x * ?Z z))))"
+    (is "(\<Sum>(x, y, z)\<in>?S. ?L x y z) = (\<Sum>(x, y, z)\<in>?S. ?R x y z)")
+  proof (safe intro!: setsum_cong)
+    fix x y z assume space: "x \<in> space MX" "y \<in> space MY" "z \<in> space MZ"
+    then have *: "?XYZ x y z / (?XZ x z * ?YZdZ y z) =
+      (?XYZ x y z / (?X x * ?YZ y z)) / (?XZ x z / (?X x * ?Z z))"
+      using order1(3)
+      by (auto simp: real_of_pinfreal_mult[symmetric] real_of_pinfreal_eq_0)
+    show "?L x y z = ?R x y z"
+    proof cases
+      assume "?XYZ x y z \<noteq> 0"
+      with space b_gt_1 order1 order2 show ?thesis unfolding *
+        by (subst log_divide)
+           (auto simp: zero_less_divide_iff zero_less_real_of_pinfreal
+                       real_of_pinfreal_eq_0 zero_less_mult_iff)
+    qed simp
+  qed
+  also have "\<dots> = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) -
+                  (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XZ x z / (?X x * ?Z z)))"
+    by (auto simp add: setsum_subtractf[symmetric] field_simps intro!: setsum_cong)
+  also have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XZ x z / (?X x * ?Z z))) =
+             (\<Sum>(x, z)\<in>space MX \<times> space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z)))"
+    unfolding setsum_cartesian_product[symmetric] setsum_commute[of _ _ "space MY"]
+              setsum_left_distrib[symmetric]
+    unfolding joint_distribution_commute_singleton[of X]
+    unfolding joint_distribution_assoc_singleton[symmetric]
+    using setsum_real_joint_distribution_singleton[OF finite_var(2) ZX, unfolded space_simps]
+    by (intro setsum_cong refl) simp
+  also have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) -
+             (\<Sum>(x, z)\<in>space MX \<times> space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z))) =
+             conditional_mutual_information b MX MY MZ X Y Z"
+    unfolding conditional_mutual_information_def
+    unfolding mutual_information_generic_eq[OF finite_var(1,3)]
+    unfolding mutual_information_generic_eq[OF finite_var(1) YZ]
+    by (simp add: space_sigma space_pair_algebra setsum_cartesian_product')
+  finally show ?thesis by simp
+qed
+
+lemma (in information_space) conditional_mutual_information_eq:
+  assumes "simple_function X" "simple_function Y" "simple_function Z"
+  shows "\<I>(X;Y|Z) = (\<Sum>(x, y, z) \<in> X`space M \<times> Y`space M \<times> Z`space M.
+             real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
+             log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
+    (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
+  using conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]
+  by simp
+
+lemma (in information_space) conditional_mutual_information_eq_mutual_information:
+  assumes X: "simple_function X" and Y: "simple_function Y"
+  shows "\<I>(X ; Y) = \<I>(X ; Y | (\<lambda>x. ()))"
 proof -
   have [simp]: "(\<lambda>x. ()) ` space M = {()}" using not_empty by auto
-
+  have C: "simple_function (\<lambda>x. ())" by auto
   show ?thesis
-    unfolding conditional_mutual_information_eq mutual_information_eq
+    unfolding conditional_mutual_information_eq[OF X Y C]
+    unfolding mutual_information_eq[OF X Y]
     by (simp add: setsum_cartesian_product' distribution_remove_const)
 qed
 
-lemma (in finite_information_space) conditional_mutual_information_positive:
-  "0 \<le> \<I>(X ; Y | Z)"
-proof -
+lemma (in prob_space) distribution_unit[simp]: "distribution (\<lambda>x. ()) {()} = 1"
+  unfolding distribution_def using measure_space_1 by auto
+
+lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\<lambda>x. (X x, ())) {(a, ())} = distribution X {a}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+
+lemma (in prob_space) setsum_distribution:
+  assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. distribution X {a}) = 1"
+  using setsum_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV \<rparr>" "\<lambda>x. ()" "{()}"]
+  using sigma_algebra_Pow[of "UNIV::unit set"] by simp
+
+lemma (in prob_space) setsum_real_distribution:
+  assumes X: "finite_random_variable MX X" shows "(\<Sum>a\<in>space MX. real (distribution X {a})) = 1"
+  using setsum_real_joint_distribution[OF assms, of "\<lparr> space = UNIV, sets = Pow UNIV \<rparr>" "\<lambda>x. ()" "{()}"]
+  using sigma_algebra_Pow[of "UNIV::unit set"] by simp
+
+lemma (in information_space) conditional_mutual_information_generic_positive:
+  assumes "finite_random_variable MX X" and "finite_random_variable MY Y" and "finite_random_variable MZ Z"
+  shows "0 \<le> conditional_mutual_information b MX MY MZ X Y Z"
+proof (cases "space MX \<times> space MY \<times> space MZ = {}")
+  case True show ?thesis
+    unfolding conditional_mutual_information_generic_eq[OF assms] True
+    by simp
+next
+  case False
   let "?dXYZ A" = "real (distribution (\<lambda>x. (X x, Y x, Z x)) A)"
   let "?dXZ A" = "real (joint_distribution X Z A)"
   let "?dYZ A" = "real (joint_distribution Y Z A)"
   let "?dX A" = "real (distribution X A)"
   let "?dZ A" = "real (distribution Z A)"
-  let ?M = "X ` space M \<times> Y ` space M \<times> Z ` space M"
+  let ?M = "space MX \<times> space MY \<times> space MZ"
 
   have split_beta: "\<And>f. split f = (\<lambda>x. f (fst x) (snd x))" by (simp add: fun_eq_iff)
 
-  have "- (\<Sum>(x, y, z) \<in> ?M. ?dXYZ {(x, y, z)} *
-    log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))
-    \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
+  note space_simps = space_pair_algebra space_sigma algebra.simps
+
+  note finite_var = assms
+  note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
+  note XZ = finite_random_variable_pairI[OF finite_var(1,3)]
+  note ZX = finite_random_variable_pairI[OF finite_var(3,1)]
+  note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
+  note XYZ = finite_random_variable_pairI[OF finite_var(1) YZ]
+  note finite = finite_var(3) YZ XZ XYZ
+  note finite = finite[THEN finite_distribution_finite, simplified space_simps]
+
+  have order: "\<And>x y z. \<lbrakk>x \<in> space MX; y \<in> space MY; z \<in> space MZ; joint_distribution X Z {(x, z)} = 0\<rbrakk>
+          \<Longrightarrow> joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} = 0"
+    unfolding joint_distribution_commute_singleton[of X]
+    unfolding joint_distribution_assoc_singleton[symmetric]
+    using finite_distribution_order(6)[OF finite_var(2) ZX]
+    by (auto simp: space_simps)
+
+  note order = order
+    finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps]
+    finite_distribution_order(5,6)[OF finite_var(2,3), simplified space_simps]
+
+  have "- conditional_mutual_information b MX MY MZ X Y Z = - (\<Sum>(x, y, z) \<in> ?M. ?dXYZ {(x, y, z)} *
+    log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))"
+    unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal
+    by (intro setsum_cong) (auto intro!: arg_cong[where f="log b"] simp: real_of_pinfreal_mult[symmetric])
+  also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
     unfolding split_beta
   proof (rule log_setsum_divide)
-    show "?M \<noteq> {}" using not_empty by simp
+    show "?M \<noteq> {}" using False by simp
     show "1 < b" using b_gt_1 .
 
+    show "finite ?M" using assms
+      unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by auto
+
+    show "(\<Sum>x\<in>?M. ?dXYZ {(fst x, fst (snd x), snd (snd x))}) = 1"
+      unfolding setsum_cartesian_product'
+      unfolding setsum_commute[of _ "space MY"]
+      unfolding setsum_commute[of _ "space MZ"]
+      by (simp_all add: space_pair_algebra
+        setsum_real_joint_distribution_singleton[OF `finite_random_variable MX X` YZ]
+        setsum_real_joint_distribution_singleton[OF `finite_random_variable MY Y` finite_var(3)]
+        setsum_real_distribution[OF `finite_random_variable MZ Z`])
+
     fix x assume "x \<in> ?M"
     let ?x = "(fst x, fst (snd x), snd (snd x))"
 
@@ -663,120 +748,180 @@
      by (simp add: real_pinfreal_nonneg mult_nonneg_nonneg divide_nonneg_nonneg)
 
     assume *: "0 < ?dXYZ {?x}"
-    thus "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
-      apply (rule_tac divide_pos_pos mult_pos_pos)+
-      by (auto simp add: real_distribution_gt_0 intro: distribution_order(6) distribution_mono_gt_0)
-  qed (simp_all add: setsum_cartesian_product' sum_over_space_real_distribution setsum_real_distribution finite_space)
-  also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>Z`space M. ?dZ {z})"
+    with `x \<in> ?M` show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
+      using finite order
+      by (cases x)
+         (auto simp add: zero_less_real_of_pinfreal zero_less_mult_iff zero_less_divide_iff)
+  qed
+  also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>space MZ. ?dZ {z})"
     apply (simp add: setsum_cartesian_product')
     apply (subst setsum_commute)
     apply (subst (2) setsum_commute)
-    by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] setsum_real_distribution
+    by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric]
+                   setsum_real_joint_distribution_singleton[OF finite_var(1,3)]
+                   setsum_real_joint_distribution_singleton[OF finite_var(2,3)]
           intro!: setsum_cong)
-  finally show ?thesis
-    unfolding conditional_mutual_information_eq sum_over_space_real_distribution
-    by (simp add: real_of_pinfreal_mult[symmetric])
+  also have "log b (\<Sum>z\<in>space MZ. ?dZ {z}) = 0"
+    unfolding setsum_real_distribution[OF finite_var(3)] by simp
+  finally show ?thesis by simp
 qed
 
+lemma (in information_space) conditional_mutual_information_positive:
+  assumes "simple_function X" and "simple_function Y" and "simple_function Z"
+  shows "0 \<le> \<I>(X;Y|Z)"
+  using conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]]
+  by simp
+
 subsection {* Conditional Entropy *}
 
 definition (in prob_space)
   "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y"
 
-abbreviation (in finite_information_space)
-  finite_conditional_entropy ("\<H>'(_ | _')") where
+abbreviation (in information_space)
+  conditional_entropy_Pow ("\<H>'(_ | _')") where
   "\<H>(X | Y) \<equiv> conditional_entropy b
     \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr>
     \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y"
 
-lemma (in finite_information_space) conditional_entropy_positive:
-  "0 \<le> \<H>(X | Y)" unfolding conditional_entropy_def using conditional_mutual_information_positive .
+lemma (in information_space) conditional_entropy_positive:
+  "simple_function X \<Longrightarrow> simple_function Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
+  unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive)
 
-lemma (in finite_information_space) conditional_entropy_generic_eq:
-  assumes MX: "finite_measure_space MX (distribution X)"
-  assumes MY: "finite_measure_space MZ (distribution Z)"
+lemma (in measure_space) empty_measureI: "A = {} \<Longrightarrow> \<mu> A = 0" by simp
+
+lemma (in information_space) conditional_entropy_generic_eq:
+  assumes MX: "finite_random_variable MX X"
+  assumes MZ: "finite_random_variable MZ Z"
   shows "conditional_entropy b MX MZ X Z =
      - (\<Sum>(x, z)\<in>space MX \<times> space MZ.
          real (joint_distribution X Z {(x, z)}) *
          log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
-  unfolding conditional_entropy_def using assms
-  apply (simp add: conditional_mutual_information_generic_eq
-                   setsum_cartesian_product' setsum_commute[of _ "space MZ"]
-                   setsum_negf[symmetric] setsum_subtractf[symmetric])
-proof (safe intro!: setsum_cong, simp)
-  fix z x assume "z \<in> space MZ" "x \<in> space MX"
-  let "?XXZ x'" = "real (joint_distribution X (\<lambda>x. (X x, Z x)) {(x, x', z)})"
-  let "?XZ x'" = "real (joint_distribution X Z {(x', z)})"
-  let "?X" = "real (distribution X {x})"
-  interpret MX: finite_measure_space MX "distribution X" by fact
-  have *: "\<And>A. A = {} \<Longrightarrow> prob A = 0" by simp
-  have XXZ: "\<And>x'. ?XXZ x' = (if x' = x then ?XZ x else 0)"
-    by (auto simp: distribution_def intro!: arg_cong[where f=prob] *)
-  have "(\<Sum>x'\<in>space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) =
-    (\<Sum>x'\<in>{x}. ?XZ x' * log b (?XZ x') - (?XZ x' * log b ?X + ?XZ x' * log b (?XZ x')))"
-    using `x \<in> space MX` MX.finite_space
-    by (safe intro!: setsum_mono_zero_cong_right)
-       (auto split: split_if_asm simp: XXZ)
-  then show "(\<Sum>x'\<in>space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) +
-      ?XZ x * log b ?X = 0" by simp
+proof -
+  interpret MX: finite_sigma_algebra MX using MX by simp
+  interpret MZ: finite_sigma_algebra MZ using MZ by simp
+  let "?XXZ x y z" = "joint_distribution X (\<lambda>x. (X x, Z x)) {(x, y, z)}"
+  let "?XZ x z" = "joint_distribution X Z {(x, z)}"
+  let "?Z z" = "distribution Z {z}"
+  let "?f x y z" = "log b (real (?XXZ x y z) / (real (?XZ x z) * real (?XZ y z / ?Z z)))"
+  { fix x z have "?XXZ x x z = ?XZ x z"
+      unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>]) }
+  note this[simp]
+  { fix x x' :: 'b and z assume "x' \<noteq> x"
+    then have "?XXZ x x' z = 0"
+      by (auto simp: distribution_def intro!: arg_cong[where f=\<mu>] empty_measureI) }
+  note this[simp]
+  { fix x x' z assume *: "x \<in> space MX" "z \<in> space MZ"
+    then have "(\<Sum>x'\<in>space MX. real (?XXZ x x' z) * ?f x x' z)
+      = (\<Sum>x'\<in>space MX. if x = x' then real (?XZ x z) * ?f x x z else 0)"
+      by (auto intro!: setsum_cong)
+    also have "\<dots> = real (?XZ x z) * ?f x x z"
+      using `x \<in> space MX` by (simp add: setsum_cases[OF MX.finite_space])
+    also have "\<dots> = real (?XZ x z) * log b (real (?Z z) / real (?XZ x z))"
+      by (auto simp: real_of_pinfreal_mult[symmetric])
+    also have "\<dots> = - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))"
+      using assms[THEN finite_distribution_finite]
+      using finite_distribution_order(6)[OF MX MZ]
+      by (auto simp: log_simps field_simps zero_less_mult_iff zero_less_real_of_pinfreal real_of_pinfreal_eq_0)
+    finally have "(\<Sum>x'\<in>space MX. real (?XXZ x x' z) * ?f x x' z) =
+      - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))" . }
+  note * = this
+
+  show ?thesis
+    unfolding conditional_entropy_def
+    unfolding conditional_mutual_information_generic_eq[OF MX MX MZ]
+    by (auto simp: setsum_cartesian_product' setsum_negf[symmetric]
+                   setsum_commute[of _ "space MZ"] *   simp del: divide_pinfreal_def
+             intro!: setsum_cong)
 qed
 
-lemma (in finite_information_space) conditional_entropy_eq:
-  "\<H>(X | Z) =
+lemma (in information_space) conditional_entropy_eq:
+  assumes "simple_function X" "simple_function Z"
+  shows "\<H>(X | Z) =
      - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
          real (joint_distribution X Z {(x, z)}) *
          log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
-  by (simp add: finite_measure_space conditional_entropy_generic_eq)
+  using conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]
+  by simp
 
-lemma (in finite_information_space) conditional_entropy_eq_ce_with_hypothesis:
-  "\<H>(X | Y) =
+lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis:
+  assumes X: "simple_function X" and Y: "simple_function Y"
+  shows "\<H>(X | Y) =
     -(\<Sum>y\<in>Y`space M. real (distribution Y {y}) *
       (\<Sum>x\<in>X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) *
               log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))"
-  unfolding conditional_entropy_eq neg_equal_iff_equal
-  apply (simp add: setsum_commute[of _ "Y`space M"] setsum_cartesian_product' setsum_divide_distrib[symmetric])
-  apply (safe intro!: setsum_cong)
-  using real_distribution_order'[of Y _ X _]
-  by (auto simp add: setsum_subtractf[of _ _ "X`space M"])
+  unfolding conditional_entropy_eq[OF assms]
+  using finite_distribution_finite[OF finite_random_variable_pairI[OF assms[THEN simple_function_imp_finite_random_variable]]]
+  using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]]
+  using finite_distribution_finite[OF Y[THEN simple_function_imp_finite_random_variable]]
+  by (auto simp: setsum_cartesian_product'  setsum_commute[of _ "Y`space M"] setsum_right_distrib real_of_pinfreal_eq_0
+           intro!: setsum_cong)
 
-lemma (in finite_information_space) conditional_entropy_eq_cartesian_sum:
-  "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
+lemma (in information_space) conditional_entropy_eq_cartesian_product:
+  assumes "simple_function X" "simple_function Y"
+  shows "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
     real (joint_distribution X Y {(x,y)}) *
     log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))"
-proof -
-  { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M"
-    then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto
-    then have "joint_distribution X Y {x} = 0"
-      unfolding distribution_def by auto }
-  then show ?thesis using finite_space
-    unfolding conditional_entropy_eq neg_equal_iff_equal setsum_cartesian_product
-    by (auto intro!: setsum_mono_zero_cong_left)
-qed
+  unfolding conditional_entropy_eq[OF assms]
+  by (auto intro!: setsum_cong simp: setsum_cartesian_product')
 
 subsection {* Equalities *}
 
-lemma (in finite_information_space) mutual_information_eq_entropy_conditional_entropy:
-  "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)"
-  unfolding mutual_information_eq entropy_eq conditional_entropy_eq
-  using finite_space
-  by (auto simp add: setsum_addf setsum_subtractf setsum_cartesian_product'
-      setsum_left_distrib[symmetric] setsum_addf setsum_real_distribution)
+lemma (in information_space) mutual_information_eq_entropy_conditional_entropy:
+  assumes X: "simple_function X" and Z: "simple_function Z"
+  shows  "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)"
+proof -
+  let "?XZ x z" = "real (joint_distribution X Z {(x, z)})"
+  let "?Z z" = "real (distribution Z {z})"
+  let "?X x" = "real (distribution X {x})"
+  note fX = X[THEN simple_function_imp_finite_random_variable]
+  note fZ = Z[THEN simple_function_imp_finite_random_variable]
+  note fX[THEN finite_distribution_finite, simp] and fZ[THEN finite_distribution_finite, simp]
+  note finite_distribution_order[OF fX fZ, simp]
+  { fix x z assume "x \<in> X`space M" "z \<in> Z`space M"
+    have "?XZ x z * log b (?XZ x z / (?X x * ?Z z)) =
+          ?XZ x z * log b (?XZ x z / ?Z z) - ?XZ x z * log b (?X x)"
+      by (auto simp: log_simps real_of_pinfreal_mult[symmetric] zero_less_mult_iff
+                     zero_less_real_of_pinfreal field_simps real_of_pinfreal_eq_0 abs_mult) }
+  note * = this
+  show ?thesis
+    unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z]
+    using setsum_real_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]]
+    by (simp add: * setsum_cartesian_product' setsum_subtractf setsum_left_distrib[symmetric]
+                     setsum_real_distribution)
+qed
 
-lemma (in finite_information_space) conditional_entropy_less_eq_entropy:
-  "\<H>(X | Z) \<le> \<H>(X)"
+lemma (in information_space) conditional_entropy_less_eq_entropy:
+  assumes X: "simple_function X" and Z: "simple_function Z"
+  shows "\<H>(X | Z) \<le> \<H>(X)"
 proof -
-  have "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" using mutual_information_eq_entropy_conditional_entropy .
-  with mutual_information_positive[of X Z] entropy_positive[of X]
+  have "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] .
+  with mutual_information_positive[OF X Z] entropy_positive[OF X]
   show ?thesis by auto
 qed
 
-lemma (in finite_information_space) entropy_chain_rule:
-  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
-  unfolding entropy_eq[of X] entropy_eq_cartesian_sum conditional_entropy_eq_cartesian_sum
-  unfolding setsum_commute[of _ "X`space M"] setsum_negf[symmetric] setsum_addf[symmetric]
-  by (rule setsum_cong)
-     (simp_all add: setsum_negf setsum_addf setsum_subtractf setsum_real_distribution
-                    setsum_left_distrib[symmetric] joint_distribution_commute[of X Y])
+lemma (in information_space) entropy_chain_rule:
+  assumes X: "simple_function X" and Y: "simple_function Y"
+  shows  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
+proof -
+  let "?XY x y" = "real (joint_distribution X Y {(x, y)})"
+  let "?Y y" = "real (distribution Y {y})"
+  let "?X x" = "real (distribution X {x})"
+  note fX = X[THEN simple_function_imp_finite_random_variable]
+  note fY = Y[THEN simple_function_imp_finite_random_variable]
+  note fX[THEN finite_distribution_finite, simp] and fY[THEN finite_distribution_finite, simp]
+  note finite_distribution_order[OF fX fY, simp]
+  { fix x y assume "x \<in> X`space M" "y \<in> Y`space M"
+    have "?XY x y * log b (?XY x y / ?X x) =
+          ?XY x y * log b (?XY x y) - ?XY x y * log b (?X x)"
+      by (auto simp: log_simps real_of_pinfreal_mult[symmetric] zero_less_mult_iff
+                     zero_less_real_of_pinfreal field_simps real_of_pinfreal_eq_0 abs_mult) }
+  note * = this
+  show ?thesis
+    using setsum_real_joint_distribution_singleton[OF fY fX]
+    unfolding entropy_eq[OF X] conditional_entropy_eq_cartesian_product[OF Y X] entropy_eq_cartesian_product[OF X Y]
+    unfolding joint_distribution_commute_singleton[of Y X] setsum_commute[of _ "X`space M"]
+    by (simp add: * setsum_subtractf setsum_left_distrib[symmetric])
+qed
 
 section {* Partitioning *}
 
@@ -893,15 +1038,26 @@
   finally show ?thesis .
 qed
 
-lemma (in finite_information_space) entropy_partition:
+lemma (in information_space) entropy_partition:
+  assumes sf: "simple_function X" "simple_function P"
   assumes svi: "subvimage (space M) X P"
   shows "\<H>(X) = \<H>(P) + \<H>(X|P)"
 proof -
+  let "?XP x p" = "real (joint_distribution X P {(x, p)})"
+  let "?X x" = "real (distribution X {x})"
+  let "?P p" = "real (distribution P {p})"
+  note fX = sf(1)[THEN simple_function_imp_finite_random_variable]
+  note fP = sf(2)[THEN simple_function_imp_finite_random_variable]
+  note fX[THEN finite_distribution_finite, simp] and fP[THEN finite_distribution_finite, simp]
+  note finite_distribution_order[OF fX fP, simp]
   have "(\<Sum>x\<in>X ` space M. real (distribution X {x}) * log b (real (distribution X {x}))) =
     (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M.
     real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})))"
   proof (subst setsum_image_split[OF svi],
-      safe intro!: finite_imageI finite_space setsum_mono_zero_cong_left imageI)
+      safe intro!: setsum_mono_zero_cong_left imageI)
+    show "finite (X ` space M)" "finite (X ` space M)" "finite (P ` space M)"
+      using sf unfolding simple_function_def by auto
+  next
     fix p x assume in_space: "p \<in> space M" "x \<in> space M"
     assume "real (joint_distribution X P {(X x, P p)}) * log b (real (joint_distribution X P {(X x, P p)})) \<noteq> 0"
     hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M \<noteq> {}" by (auto simp: distribution_def)
@@ -920,26 +1076,41 @@
           log b (real (joint_distribution X P {(X x, P p)}))"
       by (auto simp: distribution_def)
   qed
-  thus ?thesis
-  unfolding entropy_eq conditional_entropy_eq
+  moreover have "\<And>x y. real (joint_distribution X P {(x, y)}) *
+      log b (real (joint_distribution X P {(x, y)}) / real (distribution P {y})) =
+      real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})) -
+      real (joint_distribution X P {(x, y)}) * log b (real (distribution P {y}))"
+    by (auto simp add: log_simps zero_less_mult_iff field_simps)
+  ultimately show ?thesis
+    unfolding sf[THEN entropy_eq] conditional_entropy_eq[OF sf]
+    using setsum_real_joint_distribution_singleton[OF fX fP]
     by (simp add: setsum_cartesian_product' setsum_subtractf setsum_real_distribution
       setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"])
 qed
 
-corollary (in finite_information_space) entropy_data_processing:
-  "\<H>(f \<circ> X) \<le> \<H>(X)"
-  by (subst (2) entropy_partition[of _ "f \<circ> X"]) (auto intro: conditional_entropy_positive)
+corollary (in information_space) entropy_data_processing:
+  assumes X: "simple_function X" shows "\<H>(f \<circ> X) \<le> \<H>(X)"
+proof -
+  note X
+  moreover have fX: "simple_function (f \<circ> X)" using X by auto
+  moreover have "subvimage (space M) X (f \<circ> X)" by auto
+  ultimately have "\<H>(X) = \<H>(f\<circ>X) + \<H>(X|f\<circ>X)" by (rule entropy_partition)
+  then show "\<H>(f \<circ> X) \<le> \<H>(X)"
+    by (auto intro: conditional_entropy_positive[OF X fX])
+qed
 
-corollary (in finite_information_space) entropy_of_inj:
-  assumes "inj_on f (X`space M)"
+corollary (in information_space) entropy_of_inj:
+  assumes X: "simple_function X" and inj: "inj_on f (X`space M)"
   shows "\<H>(f \<circ> X) = \<H>(X)"
 proof (rule antisym)
-  show "\<H>(f \<circ> X) \<le> \<H>(X)" using entropy_data_processing .
+  show "\<H>(f \<circ> X) \<le> \<H>(X)" using entropy_data_processing[OF X] .
 next
+  have sf: "simple_function (f \<circ> X)"
+    using X by auto
   have "\<H>(X) = \<H>(the_inv_into (X`space M) f \<circ> (f \<circ> X))"
-    by (auto intro!: mutual_information_cong simp: entropy_def the_inv_into_f_f[OF assms])
+    by (auto intro!: mutual_information_cong simp: entropy_def the_inv_into_f_f[OF inj])
   also have "... \<le> \<H>(f \<circ> X)"
-    using entropy_data_processing .
+    using entropy_data_processing[OF sf] .
   finally show "\<H>(X) \<le> \<H>(f \<circ> X)" .
 qed
 
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -3,12 +3,22 @@
 header {*Lebesgue Integration*}
 
 theory Lebesgue_Integration
-imports Measure Borel
+imports Measure Borel_Space
 begin
 
-section "@{text \<mu>}-null sets"
-
-abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}"
+lemma image_set_cong:
+  assumes A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y\<in>B. f x = g y"
+  assumes B: "\<And>y. y \<in> B \<Longrightarrow> \<exists>x\<in>A. g y = f x"
+  shows "f ` A = g ` B"
+proof safe
+  fix x assume "x \<in> A"
+  with A obtain y where "f x = g y" "y \<in> B" by auto
+  then show "f x \<in> g ` B" by auto
+next
+  fix y assume "y \<in> B"
+  with B obtain x where "g y = f x" "x \<in> A" by auto
+  then show "g y \<in> f ` A" by auto
+qed
 
 lemma sums_If_finite:
   assumes finite: "finite {r. P r}"
@@ -469,6 +479,22 @@
   unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)]
   by auto
 
+lemma (in sigma_algebra) simple_function_vimage:
+  fixes g :: "'a \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  assumes g: "simple_function g" and f: "f \<in> S \<rightarrow> space M"
+  shows "sigma_algebra.simple_function (vimage_algebra S f) (\<lambda>x. g (f x))"
+proof -
+  have subset: "(\<lambda>x. g (f x)) ` S \<subseteq> g ` space M"
+    using f by auto
+  interpret V: sigma_algebra "vimage_algebra S f"
+    using f by (rule sigma_algebra_vimage)
+  show ?thesis using g
+    unfolding simple_function_eq_borel_measurable
+    unfolding V.simple_function_eq_borel_measurable
+    using measurable_vimage[OF _ f, of g borel]
+    using finite_subset[OF subset] by auto
+qed
+
 section "Simple integral"
 
 definition (in measure_space)
@@ -484,6 +510,19 @@
   thus ?thesis unfolding simple_integral_def by simp
 qed
 
+lemma (in measure_space) simple_integral_cong_measure:
+  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A" and "simple_function f"
+  shows "measure_space.simple_integral M \<nu> f = simple_integral f"
+proof -
+  interpret v: measure_space M \<nu>
+    by (rule measure_space_cong) fact
+  have "\<And>x. x \<in> space M \<Longrightarrow> f -` {f x} \<inter> space M \<in> sets M"
+    using `simple_function f`[THEN simple_functionD(2)] by auto
+  with assms show ?thesis
+    unfolding simple_integral_def v.simple_integral_def
+    by (auto intro!: setsum_cong)
+qed
+
 lemma (in measure_space) simple_integral_const[simp]:
   "simple_integral (\<lambda>x. c) = c * \<mu> (space M)"
 proof (cases "space M = {}")
@@ -590,22 +629,62 @@
     by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong)
 qed
 
+lemma (in measure_space) simple_integral_mono_AE:
+  assumes "simple_function f" and "simple_function g"
+  and mono: "AE x. f x \<le> g x"
+  shows "simple_integral f \<le> simple_integral g"
+proof -
+  let "?S x" = "(g -` {g x} \<inter> space M) \<inter> (f -` {f x} \<inter> space M)"
+  have *: "\<And>x. g -` {g x} \<inter> f -` {f x} \<inter> space M = ?S x"
+    "\<And>x. f -` {f x} \<inter> g -` {g x} \<inter> space M = ?S x" by auto
+  show ?thesis
+    unfolding *
+      simple_function_partition[OF `simple_function f` `simple_function g`]
+      simple_function_partition[OF `simple_function g` `simple_function f`]
+  proof (safe intro!: setsum_mono)
+    fix x assume "x \<in> space M"
+    then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto
+    show "the_elem (f`?S x) * \<mu> (?S x) \<le> the_elem (g`?S x) * \<mu> (?S x)"
+    proof (cases "f x \<le> g x")
+      case True then show ?thesis using * by (auto intro!: mult_right_mono)
+    next
+      case False
+      obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
+        using mono by (auto elim!: AE_E)
+      have "?S x \<subseteq> N" using N `x \<in> space M` False by auto
+      moreover have "?S x \<in> sets M" using assms `x \<in> space M`
+        by (rule_tac Int) (auto intro!: simple_functionD(2))
+      ultimately have "\<mu> (?S x) \<le> \<mu> N"
+        using `N \<in> sets M` by (auto intro!: measure_mono)
+      then show ?thesis using `\<mu> N = 0` by auto
+    qed
+  qed
+qed
+
 lemma (in measure_space) simple_integral_mono:
   assumes "simple_function f" and "simple_function g"
   and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
   shows "simple_integral f \<le> simple_integral g"
-  unfolding
-    simple_function_partition[OF `simple_function f` `simple_function g`]
-    simple_function_partition[OF `simple_function g` `simple_function f`]
-  apply (subst Int_commute)
-proof (safe intro!: setsum_mono)
-  fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M"
-  assume "x \<in> space M"
-  hence "f ` ?S = {f x}" "g ` ?S = {g x}" by auto
-  thus "the_elem (f`?S) * \<mu> ?S \<le> the_elem (g`?S) * \<mu> ?S"
-    using mono[OF `x \<in> space M`] by (auto intro!: mult_right_mono)
+proof (rule simple_integral_mono_AE[OF assms(1, 2)])
+  show "AE x. f x \<le> g x"
+    using mono by (rule AE_cong) auto
 qed
 
+lemma (in measure_space) simple_integral_cong_AE:
+  assumes "simple_function f" "simple_function g" and "AE x. f x = g x"
+  shows "simple_integral f = simple_integral g"
+  using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
+
+lemma (in measure_space) simple_integral_cong':
+  assumes sf: "simple_function f" "simple_function g"
+  and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
+  shows "simple_integral f = simple_integral g"
+proof (intro simple_integral_cong_AE sf AE_I)
+  show "\<mu> {x\<in>space M. f x \<noteq> g x} = 0" by fact
+  show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
+    using sf[THEN borel_measurable_simple_function] by auto
+qed simp
+
 lemma (in measure_space) simple_integral_indicator:
   assumes "A \<in> sets M"
   assumes "simple_function f"
@@ -637,7 +716,8 @@
       using assms(2) unfolding simple_function_def by auto
     show "f ` A \<union> {0} \<subseteq> f`space M \<union> {0}"
       using sets_into_space[OF assms(1)] by auto
-    have "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}" by (auto simp: image_iff)
+    have "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
+      by (auto simp: image_iff)
     thus "\<forall>i\<in>f ` space M \<union> {0} - (f ` A \<union> {0}).
       i * \<mu> (f -` {i} \<inter> space M \<inter> A) = 0" by auto
   next
@@ -670,45 +750,22 @@
   assumes "simple_function u" "N \<in> null_sets"
   shows "simple_integral (\<lambda>x. u x * indicator N x) = 0"
 proof -
-  have "simple_integral (\<lambda>x. u x * indicator N x) \<le>
-    simple_integral (\<lambda>x. \<omega> * indicator N x)"
-    using assms
-    by (safe intro!: simple_integral_mono simple_function_mult simple_function_indicator simple_function_const) simp
-  also have "... = 0" apply(subst simple_integral_mult)
-    using assms(2) by auto
-  finally show ?thesis by auto
+  have "AE x. indicator N x = (0 :: pinfreal)"
+    using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
+  then have "simple_integral (\<lambda>x. u x * indicator N x) = simple_integral (\<lambda>x. 0)"
+    using assms by (intro simple_integral_cong_AE) (auto intro!: AE_disjI2)
+  then show ?thesis by simp
 qed
 
-lemma (in measure_space) simple_integral_cong':
-  assumes f: "simple_function f" and g: "simple_function g"
-  and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
-  shows "simple_integral f = simple_integral g"
-proof -
-  let ?h = "\<lambda>h. \<lambda>x. (h x * indicator {x\<in>space M. f x = g x} x
-    + h x * indicator {x\<in>space M. f x \<noteq> g x} x
-    + h x * indicator (-space M) x::pinfreal)"
-  have *:"\<And>h. h = ?h h" unfolding indicator_def apply rule by auto
-  have mea_neq:"{x \<in> space M. f x \<noteq> g x} \<in> sets M" using f g by (auto simp: borel_measurable_simple_function)
-  then have mea_nullset: "{x \<in> space M. f x \<noteq> g x} \<in> null_sets" using mea by auto
-  have h1:"\<And>h::_=>pinfreal. simple_function h \<Longrightarrow>
-    simple_function (\<lambda>x. h x * indicator {x\<in>space M. f x = g x} x)"
-    apply(safe intro!: simple_function_add simple_function_mult simple_function_indicator)
-    using f g by (auto simp: borel_measurable_simple_function)
-  have h2:"\<And>h::_\<Rightarrow>pinfreal. simple_function h \<Longrightarrow>
-    simple_function (\<lambda>x. h x * indicator {x\<in>space M. f x \<noteq> g x} x)"
-    apply(safe intro!: simple_function_add simple_function_mult simple_function_indicator)
-    by(rule mea_neq)
-  have **:"\<And>a b c d e f. a = b \<Longrightarrow> c = d \<Longrightarrow> e = f \<Longrightarrow> a+c+e = b+d+f" by auto
-  note *** = simple_integral_add[OF simple_function_add[OF h1 h2] simple_function_notspace]
-    simple_integral_add[OF h1 h2]
-  show ?thesis apply(subst *[of g]) apply(subst *[of f])
-    unfolding ***[OF f f] ***[OF g g]
-  proof(rule **) case goal1 show ?case apply(rule arg_cong[where f=simple_integral]) apply rule 
-      unfolding indicator_def by auto
-  next note * = simple_integral_null_set[OF _ mea_nullset]
-    case goal2 show ?case unfolding *[OF f] *[OF g] ..
-  next case goal3 show ?case apply(rule simple_integral_cong) by auto
-  qed
+lemma (in measure_space) simple_integral_cong_AE_mult_indicator:
+  assumes sf: "simple_function f" and eq: "AE x. x \<in> S" and "S \<in> sets M"
+  shows "simple_integral f = simple_integral (\<lambda>x. f x * indicator S x)"
+proof (rule simple_integral_cong_AE)
+  show "simple_function f" by fact
+  show "simple_function (\<lambda>x. f x * indicator S x)"
+    using sf `S \<in> sets M` by auto
+  from eq show "AE x. f x = f x * indicator S x"
+    by (rule AE_mp) simp
 qed
 
 lemma (in measure_space) simple_integral_restricted:
@@ -746,12 +803,49 @@
   unfolding simple_integral_def_raw
   unfolding measure_space.simple_integral_def_raw[OF assms] by simp
 
+lemma (in measure_space) simple_integral_vimage:
+  fixes g :: "'a \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  assumes f: "bij_betw f S (space M)"
+  shows "simple_integral g =
+         measure_space.simple_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g (f x))"
+    (is "_ = measure_space.simple_integral ?T ?\<mu> _")
+proof -
+  from f interpret T: measure_space ?T ?\<mu> by (rule measure_space_isomorphic)
+  have surj: "f`S = space M"
+    using f unfolding bij_betw_def by simp
+  have *: "(\<lambda>x. g (f x)) ` S = g ` f ` S" by auto
+  have **: "f`S = space M" using f unfolding bij_betw_def by auto
+  { fix x assume "x \<in> space M"
+    have "(f ` ((\<lambda>x. g (f x)) -` {g x} \<inter> S)) =
+      (f ` (f -` (g -` {g x}) \<inter> S))" by auto
+    also have "f -` (g -` {g x}) \<inter> S = f -` (g -` {g x} \<inter> space M) \<inter> S"
+      using f unfolding bij_betw_def by auto
+    also have "(f ` (f -` (g -` {g x} \<inter> space M) \<inter> S)) = g -` {g x} \<inter> space M"
+      using ** by (intro image_vimage_inter_eq) auto
+    finally have "(f ` ((\<lambda>x. g (f x)) -` {g x} \<inter> S)) = g -` {g x} \<inter> space M" by auto }
+  then show ?thesis using assms
+    unfolding simple_integral_def T.simple_integral_def bij_betw_def
+    by (auto simp add: * intro!: setsum_cong)
+qed
+
 section "Continuous posititve integration"
 
 definition (in measure_space)
   "positive_integral f =
     (SUP g : {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}. simple_integral g)"
 
+lemma (in measure_space) positive_integral_cong_measure:
+  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A"
+  shows "measure_space.positive_integral M \<nu> f = positive_integral f"
+proof -
+  interpret v: measure_space M \<nu>
+    by (rule measure_space_cong) fact
+  with assms show ?thesis
+    unfolding positive_integral_def v.positive_integral_def SUPR_def
+    by (auto intro!: arg_cong[where f=Sup] image_cong
+             simp: simple_integral_cong_measure[of \<nu>])
+qed
+
 lemma (in measure_space) positive_integral_alt1:
   "positive_integral f =
     (SUP g : {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. simple_integral g)"
@@ -863,15 +957,122 @@
   with assms show "simple_integral f \<le> y" by auto
 qed
 
-lemma (in measure_space) positive_integral_mono:
-  assumes mono: "\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x"
+lemma (in measure_space) positive_integral_mono_AE:
+  assumes ae: "AE x. u x \<le> v x"
   shows "positive_integral u \<le> positive_integral v"
   unfolding positive_integral_alt1
 proof (safe intro!: SUPR_mono)
-  fix a assume a: "simple_function a" and "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
-  with mono have "\<forall>x\<in>space M. a x \<le> v x \<and> a x \<noteq> \<omega>" by fastsimp
-  with a show "\<exists>b\<in>{g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}. simple_integral a \<le> simple_integral b"
-    by (auto intro!: bexI[of _ a])
+  fix a assume a: "simple_function a" and mono: "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>"
+  from ae obtain N where N: "{x\<in>space M. \<not> u x \<le> v x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0"
+    by (auto elim!: AE_E)
+  have "simple_function (\<lambda>x. a x * indicator (space M - N) x)"
+    using `N \<in> sets M` a by auto
+  with a show "\<exists>b\<in>{g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}.
+    simple_integral a \<le> simple_integral b"
+  proof (safe intro!: bexI[of _ "\<lambda>x. a x * indicator (space M - N) x"]
+                      simple_integral_mono_AE)
+    show "AE x. a x \<le> a x * indicator (space M - N) x"
+    proof (rule AE_I, rule subset_refl)
+      have *: "{x \<in> space M. \<not> a x \<le> a x * indicator (space M - N) x} =
+        N \<inter> {x \<in> space M. a x \<noteq> 0}" (is "?N = _")
+        using `N \<in> sets M`[THEN sets_into_space] by (auto simp: indicator_def)
+      then show "?N \<in> sets M" 
+        using `N \<in> sets M` `simple_function a`[THEN borel_measurable_simple_function]
+        by (auto intro!: measure_mono Int)
+      then have "\<mu> ?N \<le> \<mu> N"
+        unfolding * using `N \<in> sets M` by (auto intro!: measure_mono)
+      then show "\<mu> ?N = 0" using `\<mu> N = 0` by auto
+    qed
+  next
+    fix x assume "x \<in> space M"
+    show "a x * indicator (space M - N) x \<le> v x"
+    proof (cases "x \<in> N")
+      case True then show ?thesis by simp
+    next
+      case False
+      with N mono have "a x \<le> u x" "u x \<le> v x" using `x \<in> space M` by auto
+      with False `x \<in> space M` show "a x * indicator (space M - N) x \<le> v x" by auto
+    qed
+    assume "a x * indicator (space M - N) x = \<omega>"
+    with mono `x \<in> space M` show False
+      by (simp split: split_if_asm add: indicator_def)
+  qed
+qed
+
+lemma (in measure_space) positive_integral_cong_AE:
+  "AE x. u x = v x \<Longrightarrow> positive_integral u = positive_integral v"
+  by (auto simp: eq_iff intro!: positive_integral_mono_AE)
+
+lemma (in measure_space) positive_integral_mono:
+  assumes mono: "\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x"
+  shows "positive_integral u \<le> positive_integral v"
+  using mono by (auto intro!: AE_cong positive_integral_mono_AE)
+
+lemma (in measure_space) positive_integral_vimage:
+  fixes g :: "'a \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  assumes f: "bij_betw f S (space M)"
+  shows "positive_integral g =
+         measure_space.positive_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g (f x))"
+    (is "_ = measure_space.positive_integral ?T ?\<mu> _")
+proof -
+  from f interpret T: measure_space ?T ?\<mu> by (rule measure_space_isomorphic)
+  have f_fun: "f \<in> S \<rightarrow> space M" using assms unfolding bij_betw_def by auto
+  from assms have inv: "bij_betw (the_inv_into S f) (space M) S"
+    by (rule bij_betw_the_inv_into)
+  then have inv_fun: "the_inv_into S f \<in> space M \<rightarrow> S" unfolding bij_betw_def by auto
+
+  have surj: "f`S = space M"
+    using f unfolding bij_betw_def by simp
+  have inj: "inj_on f S"
+    using f unfolding bij_betw_def by simp
+  have inv_f: "\<And>x. x \<in> space M \<Longrightarrow> f (the_inv_into S f x) = x"
+    using f_the_inv_into_f[of f S] f unfolding bij_betw_def by auto
+
+  from simple_integral_vimage[OF assms, symmetric]
+  have *: "simple_integral = T.simple_integral \<circ> (\<lambda>g. g \<circ> f)" by (simp add: comp_def)
+  show ?thesis
+    unfolding positive_integral_alt1 T.positive_integral_alt1 SUPR_def * image_compose
+  proof (safe intro!: arg_cong[where f=Sup] image_set_cong, simp_all add: comp_def)
+    fix g' :: "'a \<Rightarrow> pinfreal" assume "simple_function g'" "\<forall>x\<in>space M. g' x \<le> g x \<and> g' x \<noteq> \<omega>"
+    then show "\<exists>h. T.simple_function h \<and> (\<forall>x\<in>S. h x \<le> g (f x) \<and> h x \<noteq> \<omega>) \<and>
+                   T.simple_integral (\<lambda>x. g' (f x)) = T.simple_integral h"
+      using f unfolding bij_betw_def
+      by (auto intro!: exI[of _ "\<lambda>x. g' (f x)"]
+               simp add: le_fun_def simple_function_vimage[OF _ f_fun])
+  next
+    fix g' :: "'d \<Rightarrow> pinfreal" assume g': "T.simple_function g'" "\<forall>x\<in>S. g' x \<le> g (f x) \<and> g' x \<noteq> \<omega>"
+    let ?g = "\<lambda>x. g' (the_inv_into S f x)"
+    show "\<exists>h. simple_function h \<and> (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>) \<and>
+              T.simple_integral g' = T.simple_integral (\<lambda>x. h (f x))"
+    proof (intro exI[of _ ?g] conjI ballI)
+      { fix x assume x: "x \<in> space M"
+        then have "the_inv_into S f x \<in> S" using inv_fun by auto
+        with g' have "g' (the_inv_into S f x) \<le> g (f (the_inv_into S f x)) \<and> g' (the_inv_into S f x) \<noteq> \<omega>"
+          by auto
+        then show "g' (the_inv_into S f x) \<le> g x" "g' (the_inv_into S f x) \<noteq> \<omega>"
+          using f_the_inv_into_f[of f S x] x f unfolding bij_betw_def by auto }
+      note vimage_vimage_inv[OF f inv_f inv_fun, simp]
+      from T.simple_function_vimage[OF g'(1), unfolded space_vimage_algebra, OF inv_fun]
+      show "simple_function (\<lambda>x. g' (the_inv_into S f x))"
+        unfolding simple_function_def by (simp add: simple_function_def)
+      show "T.simple_integral g' = T.simple_integral (\<lambda>x. ?g (f x))"
+        using the_inv_into_f_f[OF inj] by (auto intro!: T.simple_integral_cong)
+    qed
+  qed
+qed
+
+lemma (in measure_space) positive_integral_vimage_inv:
+  fixes g :: "'d \<Rightarrow> pinfreal" and f :: "'d \<Rightarrow> 'a"
+  assumes f: "bij_betw f S (space M)"
+  shows "measure_space.positive_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) g =
+      positive_integral (\<lambda>x. g (the_inv_into S f x))"
+proof -
+  interpret v: measure_space "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)"
+    using f by (rule measure_space_isomorphic)
+  show ?thesis
+    unfolding positive_integral_vimage[OF f, of "\<lambda>x. g (the_inv_into S f x)"]
+    using f[unfolded bij_betw_def]
+    by (auto intro!: v.positive_integral_cong simp: the_inv_into_f_f)
 qed
 
 lemma (in measure_space) positive_integral_SUP_approx:
@@ -901,26 +1102,25 @@
   have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M"
     using `simple_function u` by (auto simp add: simple_function_def)
 
-  { fix i
-    have "(\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
-    proof safe
-      fix x assume "x \<in> space M"
-      show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
-      proof cases
-        assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
-      next
-        assume "u x \<noteq> 0"
-        with `a < 1` real `x \<in> space M`
-        have "a * u x < 1 * u x" by (rule_tac pinfreal_mult_strict_right_mono) (auto simp: image_iff)
-        also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
-          unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
-        finally obtain i where "a * u x < f i x" unfolding SUPR_def
-          by (auto simp add: less_Sup_iff)
-        hence "a * u x \<le> f i x" by auto
-        thus ?thesis using `x \<in> space M` by auto
-      qed
-    qed auto }
-  note measure_conv = measure_up[OF u Int[OF u B] this]
+  have "\<And>i. (\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def
+  proof safe
+    fix x i assume "x \<in> space M"
+    show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)"
+    proof cases
+      assume "u x = 0" thus ?thesis using `x \<in> space M` by simp
+    next
+      assume "u x \<noteq> 0"
+      with `a < 1` real `x \<in> space M`
+      have "a * u x < 1 * u x" by (rule_tac pinfreal_mult_strict_right_mono) (auto simp: image_iff)
+      also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s`
+        unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def)
+      finally obtain i where "a * u x < f i x" unfolding SUPR_def
+        by (auto simp add: less_Sup_iff)
+      hence "a * u x \<le> f i x" by auto
+      thus ?thesis using `x \<in> space M` by auto
+    qed
+  qed auto
+  note measure_conv = measure_up[OF Int[OF u B] this]
 
   have "simple_integral u = (SUP i. simple_integral (?uB i))"
     unfolding simple_integral_indicator[OF B `simple_function u`]
@@ -986,6 +1186,34 @@
   qed
 qed
 
+lemma (in measure_space) positive_integral_monotone_convergence_SUP:
+  assumes "\<And>i x. x \<in> space M \<Longrightarrow> f i x \<le> f (Suc i) x"
+  assumes "\<And>i. f i \<in> borel_measurable M"
+  shows "(SUP i. positive_integral (f i)) = positive_integral (\<lambda>x. SUP i. f i x)"
+    (is "_ = positive_integral ?u")
+proof -
+  have "?u \<in> borel_measurable M"
+    using borel_measurable_SUP[of _ f] assms by (simp add: SUPR_fun_expand)
+
+  show ?thesis
+  proof (rule antisym)
+    show "(SUP j. positive_integral (f j)) \<le> positive_integral ?u"
+      by (auto intro!: SUP_leI positive_integral_mono le_SUPI)
+  next
+    def rf \<equiv> "\<lambda>i. \<lambda>x\<in>space M. f i x" and ru \<equiv> "\<lambda>x\<in>space M. ?u x"
+    have "\<And>i. rf i \<in> borel_measurable M" unfolding rf_def
+      using assms by (simp cong: measurable_cong)
+    moreover have iso: "rf \<up> ru" using assms unfolding rf_def ru_def
+      unfolding isoton_def SUPR_fun_expand le_fun_def fun_eq_iff
+      by (auto simp: restrict_def le_fun_def SUPR_fun_expand fun_eq_iff)
+    ultimately have "positive_integral ru \<le> (SUP i. positive_integral (rf i))"
+      unfolding positive_integral_def[of ru]
+      by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx)
+    then show "positive_integral ?u \<le> (SUP i. positive_integral (f i))"
+      unfolding ru_def rf_def by (simp cong: positive_integral_cong)
+  qed
+qed
+
 lemma (in measure_space) SUP_simple_integral_sequences:
   assumes f: "f \<up> u" "\<And>i. simple_function (f i)"
   and g: "g \<up> u" "\<And>i. simple_function (g i)"
@@ -1193,47 +1421,16 @@
 qed
 
 lemma (in measure_space) positive_integral_null_set:
-  assumes borel: "u \<in> borel_measurable M" and "N \<in> null_sets"
-  shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" (is "?I = 0")
+  assumes "N \<in> null_sets" shows "positive_integral (\<lambda>x. u x * indicator N x) = 0"
 proof -
-  have "N \<in> sets M" using `N \<in> null_sets` by auto
-  have "(\<lambda>i x. min (of_nat i) (u x) * indicator N x) \<up> (\<lambda>x. u x * indicator N x)"
-    unfolding isoton_fun_expand
-  proof (safe intro!: isoton_cmult_left, unfold isoton_def, safe)
-    fix j i show "min (of_nat j) (u i) \<le> min (of_nat (Suc j)) (u i)"
-      by (rule min_max.inf_mono) auto
-  next
-    fix i show "(SUP j. min (of_nat j) (u i)) = u i"
-    proof (cases "u i")
-      case infinite
-      moreover hence "\<And>j. min (of_nat j) (u i) = of_nat j"
-        by (auto simp: min_def)
-      ultimately show ?thesis by (simp add: Sup_\<omega>)
-    next
-      case (preal r)
-      obtain j where "r \<le> of_nat j" using ex_le_of_nat ..
-      hence "u i \<le> of_nat j" using preal by (auto simp: real_of_nat_def)
-      show ?thesis
-      proof (rule pinfreal_SUPI)
-        fix y assume "\<And>j. j \<in> UNIV \<Longrightarrow> min (of_nat j) (u i) \<le> y"
-        note this[of j]
-        moreover have "min (of_nat j) (u i) = u i"
-          using `u i \<le> of_nat j` by (auto simp: min_def)
-        ultimately show "u i \<le> y" by simp
-      qed simp
-    qed
+  have "positive_integral (\<lambda>x. u x * indicator N x) = positive_integral (\<lambda>x. 0)"
+  proof (intro positive_integral_cong_AE AE_I)
+    show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
+      by (auto simp: indicator_def)
+    show "\<mu> N = 0" "N \<in> sets M"
+      using assms by auto
   qed
-  from positive_integral_isoton[OF this]
-  have "?I = (SUP i. positive_integral (\<lambda>x. min (of_nat i) (u x) * indicator N x))"
-    unfolding isoton_def using borel `N \<in> sets M` by (simp add: borel_measurable_indicator)
-  also have "\<dots> \<le> (SUP i. positive_integral (\<lambda>x. of_nat i * indicator N x))"
-  proof (rule SUP_mono, rule bexI, rule positive_integral_mono)
-    fix x i show "min (of_nat i) (u x) * indicator N x \<le> of_nat i * indicator N x"
-      by (cases "x \<in> N") auto
-  qed simp
-  also have "\<dots> = 0"
-    using `N \<in> null_sets` by (simp add: positive_integral_cmult_indicator)
-  finally show ?thesis by simp
+  then show ?thesis by simp
 qed
 
 lemma (in measure_space) positive_integral_Markov_inequality:
@@ -1270,7 +1467,7 @@
   proof
     assume "\<mu> ?A = 0"
     hence "?A \<in> null_sets" using `?A \<in> sets M` by auto
-    from positive_integral_null_set[OF borel this]
+    from positive_integral_null_set[OF this]
     have "0 = positive_integral (\<lambda>x. u x * indicator ?A x)" by simp
     thus "positive_integral u = 0" unfolding u by simp
   next
@@ -1309,34 +1506,6 @@
   qed
 qed
 
-lemma (in measure_space) positive_integral_cong_on_null_sets:
-  assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M"
-  and measure: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
-  shows "positive_integral f = positive_integral g"
-proof -
-  let ?N = "{x\<in>space M. f x \<noteq> g x}" and ?E = "{x\<in>space M. f x = g x}"
-  let "?A h x" = "h x * indicator ?E x :: pinfreal"
-  let "?B h x" = "h x * indicator ?N x :: pinfreal"
-
-  have A: "positive_integral (?A f) = positive_integral (?A g)"
-    by (auto intro!: positive_integral_cong simp: indicator_def)
-
-  have [intro]: "?N \<in> sets M" "?E \<in> sets M" using f g by auto
-  hence "?N \<in> null_sets" using measure by auto
-  hence B: "positive_integral (?B f) = positive_integral (?B g)"
-    using f g by (simp add: positive_integral_null_set)
-
-  have "positive_integral f = positive_integral (\<lambda>x. ?A f x + ?B f x)"
-    by (auto intro!: positive_integral_cong simp: indicator_def)
-  also have "\<dots> = positive_integral (?A f) + positive_integral (?B f)"
-    using f g by (auto intro!: positive_integral_add borel_measurable_indicator)
-  also have "\<dots> = positive_integral (\<lambda>x. ?A g x + ?B g x)"
-    unfolding A B using f g by (auto intro!: positive_integral_add[symmetric] borel_measurable_indicator)
-  also have "\<dots> = positive_integral g"
-    by (auto intro!: positive_integral_cong simp: indicator_def)
-  finally show ?thesis by simp
-qed
-
 lemma (in measure_space) positive_integral_restricted:
   assumes "A \<in> sets M"
   shows "measure_space.positive_integral (restricted_space A) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)"
@@ -1346,13 +1515,13 @@
   then interpret R: measure_space ?R \<mu> .
   have saR: "sigma_algebra ?R" by fact
   have *: "R.positive_integral f = R.positive_integral ?f"
-    by (auto intro!: R.positive_integral_cong)
+    by (intro R.positive_integral_cong) auto
   show ?thesis
     unfolding * R.positive_integral_def positive_integral_def
     unfolding simple_function_restricted[OF `A \<in> sets M`]
     apply (simp add: SUPR_def)
     apply (rule arg_cong[where f=Sup])
-  proof (auto simp: image_iff simple_integral_restricted[OF `A \<in> sets M`])
+  proof (auto simp add: image_iff simple_integral_restricted[OF `A \<in> sets M`])
     fix g assume "simple_function (\<lambda>x. g x * indicator A x)"
       "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x"
     then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and>
@@ -1413,6 +1582,29 @@
   shows "integral f = integral g"
   using assms by (simp cong: positive_integral_cong add: integral_def)
 
+lemma (in measure_space) integral_cong_measure:
+  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A"
+  shows "measure_space.integral M \<nu> f = integral f"
+proof -
+  interpret v: measure_space M \<nu>
+    by (rule measure_space_cong) fact
+  show ?thesis
+    unfolding integral_def v.integral_def
+    by (simp add: positive_integral_cong_measure[OF assms])
+qed
+
+lemma (in measure_space) integral_cong_AE:
+  assumes cong: "AE x. f x = g x"
+  shows "integral f = integral g"
+proof -
+  have "AE x. Real (f x) = Real (g x)"
+    using cong by (rule AE_mp) simp
+  moreover have "AE x. Real (- f x) = Real (- g x)"
+    using cong by (rule AE_mp) simp
+  ultimately show ?thesis
+    by (simp cong: positive_integral_cong_AE add: integral_def)
+qed
+
 lemma (in measure_space) integrable_cong:
   "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable f \<longleftrightarrow> integrable g"
   by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
@@ -1425,6 +1617,23 @@
   thus ?thesis by (simp del: Real_eq_0 add: integral_def)
 qed
 
+lemma (in measure_space) integral_vimage_inv:
+  assumes f: "bij_betw f S (space M)"
+  shows "measure_space.integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g x) = integral (\<lambda>x. g (the_inv_into S f x))"
+proof -
+  interpret v: measure_space "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)"
+    using f by (rule measure_space_isomorphic)
+  have "\<And>x. x \<in> space (vimage_algebra S f) \<Longrightarrow> the_inv_into S f (f x) = x"
+    using f[unfolded bij_betw_def] by (simp add: the_inv_into_f_f)
+  then have *: "v.positive_integral (\<lambda>x. Real (g (the_inv_into S f (f x)))) = v.positive_integral (\<lambda>x. Real (g x))"
+     "v.positive_integral (\<lambda>x. Real (- g (the_inv_into S f (f x)))) = v.positive_integral (\<lambda>x. Real (- g x))"
+    by (auto intro!: v.positive_integral_cong)
+  show ?thesis
+    unfolding integral_def v.integral_def
+    unfolding positive_integral_vimage[OF f]
+    by (simp add: *)
+qed
+
 lemma (in measure_space) integral_minus[intro, simp]:
   assumes "integrable f"
   shows "integrable (\<lambda>x. - f x)" "integral (\<lambda>x. - f x) = - integral f"
@@ -1549,15 +1758,26 @@
   thus ?P ?I by auto
 qed
 
+lemma (in measure_space) integral_mono_AE:
+  assumes fg: "integrable f" "integrable g"
+  and mono: "AE t. f t \<le> g t"
+  shows "integral f \<le> integral g"
+proof -
+  have "AE x. Real (f x) \<le> Real (g x)"
+    using mono by (rule AE_mp) (auto intro!: AE_cong)
+  moreover have "AE x. Real (- g x) \<le> Real (- f x)" 
+    using mono by (rule AE_mp) (auto intro!: AE_cong)
+  ultimately show ?thesis using fg
+    by (auto simp: integral_def integrable_def diff_minus
+             intro!: add_mono real_of_pinfreal_mono positive_integral_mono_AE)
+qed
+
 lemma (in measure_space) integral_mono:
   assumes fg: "integrable f" "integrable g"
   and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
   shows "integral f \<le> integral g"
-  using fg unfolding integral_def integrable_def diff_minus
-proof (safe intro!: add_mono real_of_pinfreal_mono le_imp_neg_le positive_integral_mono)
-  fix x assume "x \<in> space M" from mono[OF this]
-  show "Real (f x) \<le> Real (g x)" "Real (- g x) \<le> Real (- f x)" by auto
-qed
+  apply (rule integral_mono_AE[OF fg])
+  using mono by (rule AE_cong) auto
 
 lemma (in measure_space) integral_diff[simp, intro]:
   assumes f: "integrable f" and g: "integrable g"
@@ -1796,6 +2016,29 @@
     by (simp add: real_of_pinfreal_eq_0)
 qed
 
+lemma (in measure_space) positive_integral_omega:
+  assumes "f \<in> borel_measurable M"
+  and "positive_integral f \<noteq> \<omega>"
+  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
+proof -
+  have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = positive_integral (\<lambda>x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x)"
+    using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp
+  also have "\<dots> \<le> positive_integral f"
+    by (auto intro!: positive_integral_mono simp: indicator_def)
+  finally show ?thesis
+    using assms(2) by (cases ?thesis) auto
+qed
+
+lemma (in measure_space) simple_integral_omega:
+  assumes "simple_function f"
+  and "simple_integral f \<noteq> \<omega>"
+  shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0"
+proof (rule positive_integral_omega)
+  show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
+  show "positive_integral f \<noteq> \<omega>"
+    using assms by (simp add: positive_integral_eq_simple_integral)
+qed
+
 lemma (in measure_space) integral_dominated_convergence:
   assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
   and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x"
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -1,38 +1,115 @@
-
+(*  Author: Robert Himmelmann, TU Muenchen *)
 header {* Lebsegue measure *}
-(*  Author:                     Robert Himmelmann, TU Muenchen *)
-
 theory Lebesgue_Measure
-  imports Gauge_Measure Measure Lebesgue_Integration
+  imports Product_Measure Gauge_Measure Complete_Measure
 begin
 
-subsection {* Various *}
+lemma (in complete_lattice) SUP_pair:
+  "(SUP i:A. SUP j:B. f i j) = (SUP p:A\<times>B. (\<lambda> (i, j). f i j) p)" (is "?l = ?r")
+proof (intro antisym SUP_leI)
+  fix i j assume "i \<in> A" "j \<in> B"
+  then have "(case (i,j) of (i,j) \<Rightarrow> f i j) \<le> ?r"
+    by (intro SUPR_upper) auto
+  then show "f i j \<le> ?r" by auto
+next
+  fix p assume "p \<in> A \<times> B"
+  then obtain i j where "p = (i,j)" "i \<in> A" "j \<in> B" by auto
+  have "f i j \<le> (SUP j:B. f i j)" using `j \<in> B` by (intro SUPR_upper)
+  also have "(SUP j:B. f i j) \<le> ?l" using `i \<in> A` by (intro SUPR_upper)
+  finally show "(case p of (i, j) \<Rightarrow> f i j) \<le> ?l" using `p = (i,j)` by simp
+qed
 
-lemma seq_offset_iff:"f ----> l \<longleftrightarrow> (\<lambda>i. f (i + k)) ----> l"
-  using seq_offset_rev seq_offset[of f l k] by auto
+lemma (in complete_lattice) SUP_surj_compose:
+  assumes *: "f`A = B" shows "SUPR A (g \<circ> f) = SUPR B g"
+  unfolding SUPR_def unfolding *[symmetric]
+  by (simp add: image_compose)
+
+lemma (in complete_lattice) SUP_swap:
+  "(SUP i:A. SUP j:B. f i j) = (SUP j:B. SUP i:A. f i j)"
+proof -
+  have *: "(\<lambda>(i,j). (j,i)) ` (B \<times> A) = A \<times> B" by auto
+  show ?thesis
+    unfolding SUP_pair SUP_surj_compose[symmetric, OF *]
+    by (auto intro!: arg_cong[where f=Sup] image_eqI simp: comp_def SUPR_def)
+qed
 
-lemma has_integral_disjoint_union: fixes f::"'n::ordered_euclidean_space \<Rightarrow> 'a::banach"
-  assumes "(f has_integral i) s" "(f has_integral j) t" "s \<inter> t = {}"
-  shows "(f has_integral (i + j)) (s \<union> t)"
-  apply(rule has_integral_union[OF assms(1-2)]) unfolding assms by auto
+lemma SUP_\<omega>: "(SUP i:A. f i) = \<omega> \<longleftrightarrow> (\<forall>x<\<omega>. \<exists>i\<in>A. x < f i)"
+proof
+  assume *: "(SUP i:A. f i) = \<omega>"
+  show "(\<forall>x<\<omega>. \<exists>i\<in>A. x < f i)" unfolding *[symmetric]
+  proof (intro allI impI)
+    fix x assume "x < SUPR A f" then show "\<exists>i\<in>A. x < f i"
+      unfolding less_SUP_iff by auto
+  qed
+next
+  assume *: "\<forall>x<\<omega>. \<exists>i\<in>A. x < f i"
+  show "(SUP i:A. f i) = \<omega>"
+  proof (rule pinfreal_SUPI)
+    fix y assume **: "\<And>i. i \<in> A \<Longrightarrow> f i \<le> y"
+    show "\<omega> \<le> y"
+    proof cases
+      assume "y < \<omega>"
+      from *[THEN spec, THEN mp, OF this]
+      obtain i where "i \<in> A" "\<not> (f i \<le> y)" by auto
+      with ** show ?thesis by auto
+    qed auto
+  qed auto
+qed
 
-lemma lim_eq: assumes "\<forall>n>N. f n = g n" shows "(f ----> l) \<longleftrightarrow> (g ----> l)" using assms 
-proof(induct N arbitrary: f g) case 0
-  hence *:"\<And>n. f (Suc n) = g (Suc n)" by auto
-  show ?case apply(subst LIMSEQ_Suc_iff[THEN sym]) apply(subst(2) LIMSEQ_Suc_iff[THEN sym])
-    unfolding * ..
-next case (Suc n)
-  show ?case apply(subst LIMSEQ_Suc_iff[THEN sym]) apply(subst(2) LIMSEQ_Suc_iff[THEN sym])
-    apply(rule Suc(1)) using Suc(2) by auto
+lemma psuminf_commute:
+  shows "(\<Sum>\<^isub>\<infinity> i j. f i j) = (\<Sum>\<^isub>\<infinity> j i. f i j)"
+proof -
+  have "(SUP n. \<Sum> i < n. SUP m. \<Sum> j < m. f i j) = (SUP n. SUP m. \<Sum> i < n. \<Sum> j < m. f i j)"
+    apply (subst SUPR_pinfreal_setsum)
+    by auto
+  also have "\<dots> = (SUP m n. \<Sum> j < m. \<Sum> i < n. f i j)"
+    apply (subst SUP_swap)
+    apply (subst setsum_commute)
+    by auto
+  also have "\<dots> = (SUP m. \<Sum> j < m. SUP n. \<Sum> i < n. f i j)"
+    apply (subst SUPR_pinfreal_setsum)
+    by auto
+  finally show ?thesis
+    unfolding psuminf_def by auto
+qed
+
+lemma psuminf_SUP_eq:
+  assumes "\<And>n i. f n i \<le> f (Suc n) i"
+  shows "(\<Sum>\<^isub>\<infinity> i. SUP n::nat. f n i) = (SUP n::nat. \<Sum>\<^isub>\<infinity> i. f n i)"
+proof -
+  { fix n :: nat
+    have "(\<Sum>i<n. SUP k. f k i) = (SUP k. \<Sum>i<n. f k i)"
+      using assms by (auto intro!: SUPR_pinfreal_setsum[symmetric]) }
+  note * = this
+  show ?thesis
+    unfolding psuminf_def
+    unfolding *
+    apply (subst SUP_swap) ..
 qed
 
 subsection {* Standard Cubes *}
 
-definition cube where
-  "cube (n::nat) \<equiv> {\<chi>\<chi> i. - real n .. (\<chi>\<chi> i. real n)::_::ordered_euclidean_space}"
+definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
+  "cube n \<equiv> {\<chi>\<chi> i. - real n .. \<chi>\<chi> i. real n}"
+
+lemma cube_closed[intro]: "closed (cube n)"
+  unfolding cube_def by auto
+
+lemma cube_subset[intro]: "n \<le> N \<Longrightarrow> cube n \<subseteq> cube N"
+  by (fastsimp simp: eucl_le[where 'a='a] cube_def)
 
-lemma cube_subset[intro]:"n\<le>N \<Longrightarrow> cube n \<subseteq> (cube N::'a::ordered_euclidean_space set)"
-  apply(auto simp: eucl_le[where 'a='a] cube_def) apply(erule_tac[!] x=i in allE)+ by auto
+lemma cube_subset_iff:
+  "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N"
+proof
+  assume subset: "cube n \<subseteq> (cube N::'a set)"
+  then have "((\<chi>\<chi> i. real n)::'a) \<in> cube N"
+    using DIM_positive[where 'a='a]
+    by (fastsimp simp: cube_def eucl_le[where 'a='a])
+  then show "n \<le> N"
+    by (fastsimp simp: cube_def eucl_le[where 'a='a])
+next
+  assume "n \<le> N" then show "cube n \<subseteq> (cube N::'a set)" by (rule cube_subset)
+qed
 
 lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n"
   unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta'
@@ -63,202 +140,277 @@
   apply(rule has_gmeasure_subset[of "s\<inter>cube n" _ "cube n"])
   unfolding has_gmeasure_measure[THEN sym] using assms by auto
 
+lemma has_gmeasure_cube[intro]: "(cube n::('a::ordered_euclidean_space) set)
+  has_gmeasure ((2 * real n) ^ (DIM('a)))"
+proof-
+  have "content {\<chi>\<chi> i. - real n..(\<chi>\<chi> i. real n)::'a} = (2 * real n) ^ (DIM('a))"
+    apply(subst content_closed_interval) defer
+    by (auto simp add:setprod_constant)
+  thus ?thesis unfolding cube_def
+    using has_gmeasure_interval(1)[of "(\<chi>\<chi> i. - real n)::'a" "(\<chi>\<chi> i. real n)::'a"]
+    by auto
+qed
+
+lemma gmeasure_cube_eq[simp]:
+  "gmeasure (cube n::('a::ordered_euclidean_space) set) = (2 * real n) ^ DIM('a)"
+  by (intro measure_unique) auto
+
+lemma gmeasure_cube_ge_n: "gmeasure (cube n::('a::ordered_euclidean_space) set) \<ge> real n"
+proof cases
+  assume "n = 0" then show ?thesis by simp
+next
+  assume "n \<noteq> 0"
+  have "real n \<le> (2 * real n)^1" by simp
+  also have "\<dots> \<le> (2 * real n)^DIM('a)"
+    using DIM_positive[where 'a='a] `n \<noteq> 0`
+    by (intro power_increasing) auto
+  also have "\<dots> = gmeasure (cube n::'a set)" by simp
+  finally show ?thesis .
+qed
+
+lemma gmeasure_setsum:
+  assumes "finite A" and "\<And>s t. s \<in> A \<Longrightarrow> t \<in> A \<Longrightarrow> s \<noteq> t \<Longrightarrow> f s \<inter> f t = {}"
+    and "\<And>i. i \<in> A \<Longrightarrow> gmeasurable (f i)"
+  shows "gmeasure (\<Union>i\<in>A. f i) = (\<Sum>i\<in>A. gmeasure (f i))"
+proof -
+  have "gmeasure (\<Union>i\<in>A. f i) = gmeasure (\<Union>f ` A)" by auto
+  also have "\<dots> = setsum gmeasure (f ` A)" using assms
+  proof (intro measure_negligible_unions)
+    fix X Y assume "X \<in> f`A" "Y \<in> f`A" "X \<noteq> Y"
+    then have "X \<inter> Y = {}" using assms by auto
+    then show "negligible (X \<inter> Y)" by auto
+  qed auto
+  also have "\<dots> = setsum gmeasure (f ` A - {{}})"
+    using assms by (intro setsum_mono_zero_cong_right) auto
+  also have "\<dots> = (\<Sum>i\<in>A - {i. f i = {}}. gmeasure (f i))"
+  proof (intro setsum_reindex_cong inj_onI)
+    fix s t assume *: "s \<in> A - {i. f i = {}}" "t \<in> A - {i. f i = {}}" "f s = f t"
+    show "s = t"
+    proof (rule ccontr)
+      assume "s \<noteq> t" with assms(2)[of s t] * show False by auto
+    qed
+  qed auto
+  also have "\<dots> = (\<Sum>i\<in>A. gmeasure (f i))"
+    using assms by (intro setsum_mono_zero_cong_left) auto
+  finally show ?thesis .
+qed
+
+lemma gmeasurable_finite_UNION[intro]:
+  assumes "\<And>i. i \<in> S \<Longrightarrow> gmeasurable (A i)" "finite S"
+  shows "gmeasurable (\<Union>i\<in>S. A i)"
+  unfolding UNION_eq_Union_image using assms
+  by (intro gmeasurable_finite_unions) auto
+
+lemma gmeasurable_countable_UNION[intro]:
+  fixes A :: "nat \<Rightarrow> ('a::ordered_euclidean_space) set"
+  assumes measurable: "\<And>i. gmeasurable (A i)"
+    and finite: "\<And>n. gmeasure (UNION {.. n} A) \<le> B"
+  shows "gmeasurable (\<Union>i. A i)"
+proof -
+  have *: "\<And>n. \<Union>{A k |k. k \<le> n} = (\<Union>i\<le>n. A i)"
+    "(\<Union>{A n |n. n \<in> UNIV}) = (\<Union>i. A i)" by auto
+  show ?thesis
+    by (rule gmeasurable_countable_unions_strong[of A B, unfolded *, OF assms])
+qed
 
 subsection {* Measurability *}
 
-definition lmeasurable :: "('a::ordered_euclidean_space) set => bool" where
-  "lmeasurable s \<equiv> (\<forall>n::nat. gmeasurable (s \<inter> cube n))"
+definition lebesgue :: "'a::ordered_euclidean_space algebra" where
+  "lebesgue = \<lparr> space = UNIV, sets = {A. \<forall>n. gmeasurable (A \<inter> cube n)} \<rparr>"
+
+lemma space_lebesgue[simp]:"space lebesgue = UNIV"
+  unfolding lebesgue_def by auto
 
-lemma lmeasurableD[dest]:assumes "lmeasurable s"
-  shows "\<And>n. gmeasurable (s \<inter> cube n)"
-  using assms unfolding lmeasurable_def by auto
+lemma lebesgueD[dest]: assumes "S \<in> sets lebesgue"
+  shows "\<And>n. gmeasurable (S \<inter> cube n)"
+  using assms unfolding lebesgue_def by auto
 
-lemma measurable_imp_lmeasurable: assumes"gmeasurable s"
-  shows "lmeasurable s" unfolding lmeasurable_def cube_def 
+lemma lebesgueI[intro]: assumes "gmeasurable S"
+  shows "S \<in> sets lebesgue" unfolding lebesgue_def cube_def
   using assms gmeasurable_interval by auto
 
-lemma lmeasurable_empty[intro]: "lmeasurable {}"
-  using gmeasurable_empty apply- apply(drule_tac measurable_imp_lmeasurable) .
-
-lemma lmeasurable_union[intro]: assumes "lmeasurable s" "lmeasurable t"
-  shows "lmeasurable (s \<union> t)"
-  using assms unfolding lmeasurable_def Int_Un_distrib2 
-  by(auto intro:gmeasurable_union)
+lemma lebesgueI2: "(\<And>n. gmeasurable (S \<inter> cube n)) \<Longrightarrow> S \<in> sets lebesgue"
+  using assms unfolding lebesgue_def by auto
 
-lemma lmeasurable_countable_unions_strong:
-  fixes s::"nat => 'a::ordered_euclidean_space set"
-  assumes "\<And>n::nat. lmeasurable(s n)"
-  shows "lmeasurable(\<Union>{ s(n) |n. n \<in> UNIV })"
-  unfolding lmeasurable_def
-proof fix n::nat
-  have *:"\<Union>{s n |n. n \<in> UNIV} \<inter> cube n = \<Union>{s k \<inter> cube n |k. k \<in> UNIV}" by auto
-  show "gmeasurable (\<Union>{s n |n. n \<in> UNIV} \<inter> cube n)" unfolding *
-    apply(rule gmeasurable_countable_unions_strong)
-    apply(rule assms[unfolded lmeasurable_def,rule_format])
-  proof- fix k::nat
-    show "gmeasure (\<Union>{s ka \<inter> cube n |ka. ka \<le> k}) \<le> gmeasure (cube n::'a set)"
-      apply(rule measure_subset) apply(rule gmeasurable_finite_unions)
-      using assms(1)[unfolded lmeasurable_def] by auto
-  qed
+interpretation lebesgue: sigma_algebra lebesgue
+proof
+  show "sets lebesgue \<subseteq> Pow (space lebesgue)"
+    unfolding lebesgue_def by auto
+  show "{} \<in> sets lebesgue"
+    using gmeasurable_empty by auto
+  { fix A B :: "'a set" assume "A \<in> sets lebesgue" "B \<in> sets lebesgue"
+    then show "A \<union> B \<in> sets lebesgue"
+      by (auto intro: gmeasurable_union simp: lebesgue_def Int_Un_distrib2) }
+  { fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets lebesgue"
+    show "(\<Union>i. A i) \<in> sets lebesgue"
+    proof (rule lebesgueI2)
+      fix n show "gmeasurable ((\<Union>i. A i) \<inter> cube n)" unfolding UN_extend_simps
+        using A
+        by (intro gmeasurable_countable_UNION[where B="gmeasure (cube n::'a set)"])
+           (auto intro!: measure_subset gmeasure_setsum simp: UN_extend_simps simp del: gmeasure_cube_eq UN_simps)
+    qed }
+  { fix A assume A: "A \<in> sets lebesgue" show "space lebesgue - A \<in> sets lebesgue"
+    proof (rule lebesgueI2)
+      fix n
+      have *: "(space lebesgue - A) \<inter> cube n = cube n - (A \<inter> cube n)"
+        unfolding lebesgue_def by auto
+      show "gmeasurable ((space lebesgue - A) \<inter> cube n)" unfolding *
+        using A by (auto intro!: gmeasurable_diff)
+    qed }
 qed
 
-lemma lmeasurable_inter[intro]: fixes s::"'a :: ordered_euclidean_space set"
-  assumes "lmeasurable s" "lmeasurable t" shows "lmeasurable (s \<inter> t)"
-  unfolding lmeasurable_def
-proof fix n::nat
-  have *:"s \<inter> t \<inter> cube n = (s \<inter> cube n) \<inter> (t \<inter> cube n)" by auto
-  show "gmeasurable (s \<inter> t \<inter> cube n)"
-    using assms unfolding lmeasurable_def *
-    using gmeasurable_inter[of "s \<inter> cube n" "t \<inter> cube n"] by auto
+lemma lebesgueI_borel[intro, simp]: fixes s::"'a::ordered_euclidean_space set"
+  assumes "s \<in> sets borel" shows "s \<in> sets lebesgue"
+proof- let ?S = "range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)})"
+  have *:"?S \<subseteq> sets lebesgue" by auto
+  have "s \<in> sigma_sets UNIV ?S" using assms
+    unfolding borel_eq_atLeastAtMost by (simp add: sigma_def)
+  thus ?thesis
+    using lebesgue.sigma_subset[of "\<lparr> space = UNIV, sets = ?S\<rparr>", simplified, OF *]
+    by (auto simp: sigma_def)
 qed
 
-lemma lmeasurable_complement[intro]: assumes "lmeasurable s"
-  shows "lmeasurable (UNIV - s)"
-  unfolding lmeasurable_def
-proof fix n::nat
-  have *:"(UNIV - s) \<inter> cube n = cube n - (s \<inter> cube n)" by auto
-  show "gmeasurable ((UNIV - s) \<inter> cube n)" unfolding * 
-    apply(rule gmeasurable_diff) using assms unfolding lmeasurable_def by auto
-qed
-
-lemma lmeasurable_finite_unions:
-  assumes "finite f" "\<And>s. s \<in> f \<Longrightarrow> lmeasurable s"
-  shows "lmeasurable (\<Union> f)" unfolding lmeasurable_def
-proof fix n::nat have *:"(\<Union>f \<inter> cube n) = \<Union>{x \<inter> cube n | x . x\<in>f}" by auto
-  show "gmeasurable (\<Union>f \<inter> cube n)" unfolding *
-    apply(rule gmeasurable_finite_unions) unfolding simple_image 
-    using assms unfolding lmeasurable_def by auto
-qed
-
-lemma negligible_imp_lmeasurable[dest]: fixes s::"'a::ordered_euclidean_space set"
-  assumes "negligible s" shows "lmeasurable s"
-  unfolding lmeasurable_def
-proof case goal1
+lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set"
+  assumes "negligible s" shows "s \<in> sets lebesgue"
+proof (rule lebesgueI2)
+  fix n
   have *:"\<And>x. (if x \<in> cube n then indicator s x else 0) = (if x \<in> s \<inter> cube n then 1 else 0)"
     unfolding indicator_def_raw by auto
   note assms[unfolded negligible_def,rule_format,of "(\<chi>\<chi> i. - real n)::'a" "\<chi>\<chi> i. real n"]
-  thus ?case apply-apply(rule gmeasurableI[of _ 0]) unfolding has_gmeasure_def
+  thus "gmeasurable (s \<inter> cube n)" apply-apply(rule gmeasurableI[of _ 0]) unfolding has_gmeasure_def
     apply(subst(asm) has_integral_restrict_univ[THEN sym]) unfolding cube_def[symmetric]
     apply(subst has_integral_restrict_univ[THEN sym]) unfolding * .
 qed
 
-
 section {* The Lebesgue Measure *}
 
-definition has_lmeasure (infixr "has'_lmeasure" 80) where
-  "s has_lmeasure m \<equiv> lmeasurable s \<and> ((\<lambda>n. Real (gmeasure (s \<inter> cube n))) ---> m) sequentially"
+definition "lmeasure A = (SUP n. Real (gmeasure (A \<inter> cube n)))"
 
-lemma has_lmeasureD: assumes "s has_lmeasure m"
-  shows "lmeasurable s" "gmeasurable (s \<inter> cube n)"
-  "((\<lambda>n. Real (gmeasure (s \<inter> cube n))) ---> m) sequentially"
-  using assms unfolding has_lmeasure_def lmeasurable_def by auto
+lemma lmeasure_eq_0: assumes "negligible S" shows "lmeasure S = 0"
+proof -
+  from lebesgueI_negligible[OF assms]
+  have "\<And>n. gmeasurable (S \<inter> cube n)" by auto
+  from gmeasurable_measure_eq_0[OF this]
+  have "\<And>n. gmeasure (S \<inter> cube n) = 0" using assms by auto
+  then show ?thesis unfolding lmeasure_def by simp
+qed
+
+lemma lmeasure_iff_LIMSEQ:
+  assumes "A \<in> sets lebesgue" "0 \<le> m"
+  shows "lmeasure A = Real m \<longleftrightarrow> (\<lambda>n. (gmeasure (A \<inter> cube n))) ----> m"
+  unfolding lmeasure_def using assms cube_subset[where 'a='a]
+  by (intro SUP_eq_LIMSEQ monoI measure_subset) force+
 
-lemma has_lmeasureI: assumes "lmeasurable s" "((\<lambda>n. Real (gmeasure (s \<inter> cube n))) ---> m) sequentially"
-  shows "s has_lmeasure m" using assms unfolding has_lmeasure_def by auto
-
-definition lmeasure where
-  "lmeasure s \<equiv> SOME m. s has_lmeasure m"
+interpretation lebesgue: measure_space lebesgue lmeasure
+proof
+  show "lmeasure {} = 0"
+    by (auto intro!: lmeasure_eq_0)
+  show "countably_additive lebesgue lmeasure"
+  proof (unfold countably_additive_def, intro allI impI conjI)
+    fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets lebesgue" "disjoint_family A"
+    then have A: "\<And>i. A i \<in> sets lebesgue" by auto
+    show "(\<Sum>\<^isub>\<infinity>n. lmeasure (A n)) = lmeasure (\<Union>i. A i)" unfolding lmeasure_def
+    proof (subst psuminf_SUP_eq)
+      { fix i n
+        have "gmeasure (A i \<inter> cube n) \<le> gmeasure (A i \<inter> cube (Suc n))"
+          using A cube_subset[of n "Suc n"] by (auto intro!: measure_subset)
+        then show "Real (gmeasure (A i \<inter> cube n)) \<le> Real (gmeasure (A i \<inter> cube (Suc n)))"
+          by auto }
+      show "(SUP n. \<Sum>\<^isub>\<infinity>i. Real (gmeasure (A i \<inter> cube n))) = (SUP n. Real (gmeasure ((\<Union>i. A i) \<inter> cube n)))"
+      proof (intro arg_cong[where f="SUPR UNIV"] ext)
+        fix n
+        have sums: "(\<lambda>i. gmeasure (A i \<inter> cube n)) sums gmeasure (\<Union>{A i \<inter> cube n |i. i \<in> UNIV})"
+        proof (rule has_gmeasure_countable_negligible_unions(2))
+          fix i show "gmeasurable (A i \<inter> cube n)" using A by auto
+        next
+          fix i m :: nat assume "m \<noteq> i"
+          then have "A m \<inter> cube n \<inter> (A i \<inter> cube n) = {}"
+            using `disjoint_family A` unfolding disjoint_family_on_def by auto
+          then show "negligible (A m \<inter> cube n \<inter> (A i \<inter> cube n))" by auto
+        next
+          fix i
+          have "(\<Sum>k = 0..i. gmeasure (A k \<inter> cube n)) = gmeasure (\<Union>k\<le>i . A k \<inter> cube n)"
+            unfolding atLeast0AtMost using A
+          proof (intro gmeasure_setsum[symmetric])
+            fix s t :: nat assume "s \<noteq> t" then have "A t \<inter> A s = {}"
+              using `disjoint_family A` unfolding disjoint_family_on_def by auto
+            then show "A s \<inter> cube n \<inter> (A t \<inter> cube n) = {}" by auto
+          qed auto
+          also have "\<dots> \<le> gmeasure (cube n :: 'b set)" using A
+            by (intro measure_subset gmeasurable_finite_UNION) auto
+          finally show "(\<Sum>k = 0..i. gmeasure (A k \<inter> cube n)) \<le> gmeasure (cube n :: 'b set)" .
+        qed
+        show "(\<Sum>\<^isub>\<infinity>i. Real (gmeasure (A i \<inter> cube n))) = Real (gmeasure ((\<Union>i. A i) \<inter> cube n))"
+          unfolding psuminf_def
+          apply (subst setsum_Real)
+          apply (simp add: measure_pos_le)
+        proof (rule SUP_eq_LIMSEQ[THEN iffD2])
+          have "(\<Union>{A i \<inter> cube n |i. i \<in> UNIV}) = (\<Union>i. A i) \<inter> cube n" by auto
+          with sums show "(\<lambda>i. \<Sum>k<i. gmeasure (A k \<inter> cube n)) ----> gmeasure ((\<Union>i. A i) \<inter> cube n)"
+            unfolding sums_def atLeast0LessThan by simp
+        qed (auto intro!: monoI setsum_nonneg setsum_mono2)
+      qed
+    qed
+  qed
+qed
 
-lemma has_lmeasure_has_gmeasure: assumes "s has_lmeasure (Real m)" "m\<ge>0"
+lemma lmeasure_finite_has_gmeasure: assumes "s \<in> sets lebesgue" "lmeasure s = Real m" "0 \<le> m"
   shows "s has_gmeasure m"
-proof- note s = has_lmeasureD[OF assms(1)]
+proof-
   have *:"(\<lambda>n. (gmeasure (s \<inter> cube n))) ----> m"
-    using s(3) apply(subst (asm) lim_Real) using s(2) assms(2) by auto
-
+    using `lmeasure s = Real m` unfolding lmeasure_iff_LIMSEQ[OF `s \<in> sets lebesgue` `0 \<le> m`] .
+  have s: "\<And>n. gmeasurable (s \<inter> cube n)" using assms by auto
   have "(\<lambda>x. if x \<in> s then 1 else (0::real)) integrable_on UNIV \<and>
     (\<lambda>k. integral UNIV (\<lambda>x. if x \<in> s \<inter> cube k then 1 else (0::real)))
     ----> integral UNIV (\<lambda>x. if x \<in> s then 1 else 0)"
   proof(rule monotone_convergence_increasing)
-    have "\<forall>n. gmeasure (s \<inter> cube n) \<le> m" apply(rule ccontr) unfolding not_all not_le
-    proof(erule exE) fix k assume k:"m < gmeasure (s \<inter> cube k)"
-      hence "gmeasure (s \<inter> cube k) - m > 0" by auto
-      from *[unfolded Lim_sequentially,rule_format,OF this] guess N ..
-      note this[unfolded dist_real_def,rule_format,of "N + k"]
-      moreover have "gmeasure (s \<inter> cube (N + k)) \<ge> gmeasure (s \<inter> cube k)" apply-
-        apply(rule measure_subset) prefer 3 using s(2) 
-        using cube_subset[of k "N + k"] by auto
-      ultimately show False by auto
-    qed
-    thus *:"bounded {integral UNIV (\<lambda>x. if x \<in> s \<inter> cube k then 1 else (0::real)) |k. True}" 
-      unfolding integral_measure_univ[OF s(2)] bounded_def apply-
+    have "lmeasure s \<le> Real m" using `lmeasure s = Real m` by simp
+    then have "\<forall>n. gmeasure (s \<inter> cube n) \<le> m"
+      unfolding lmeasure_def complete_lattice_class.SUP_le_iff
+      using `0 \<le> m` by (auto simp: measure_pos_le)
+    thus *:"bounded {integral UNIV (\<lambda>x. if x \<in> s \<inter> cube k then 1 else (0::real)) |k. True}"
+      unfolding integral_measure_univ[OF s] bounded_def apply-
       apply(rule_tac x=0 in exI,rule_tac x=m in exI) unfolding dist_real_def
       by (auto simp: measure_pos_le)
-
     show "\<forall>k. (\<lambda>x. if x \<in> s \<inter> cube k then (1::real) else 0) integrable_on UNIV"
       unfolding integrable_restrict_univ
-      using s(2) unfolding gmeasurable_def has_gmeasure_def by auto
+      using s unfolding gmeasurable_def has_gmeasure_def by auto
     have *:"\<And>n. n \<le> Suc n" by auto
     show "\<forall>k. \<forall>x\<in>UNIV. (if x \<in> s \<inter> cube k then 1 else 0) \<le> (if x \<in> s \<inter> cube (Suc k) then 1 else (0::real))"
       using cube_subset[OF *] by fastsimp
     show "\<forall>x\<in>UNIV. (\<lambda>k. if x \<in> s \<inter> cube k then 1 else 0) ----> (if x \<in> s then 1 else (0::real))"
-      unfolding Lim_sequentially 
+      unfolding Lim_sequentially
     proof safe case goal1 from real_arch_lt[of "norm x"] guess N .. note N = this
       show ?case apply(rule_tac x=N in exI)
       proof safe case goal1
         have "x \<in> cube n" using cube_subset[OF goal1] N
-          using ball_subset_cube[of N] by(auto simp: dist_norm) 
+          using ball_subset_cube[of N] by(auto simp: dist_norm)
         thus ?case using `e>0` by auto
       qed
     qed
   qed note ** = conjunctD2[OF this]
   hence *:"m = integral UNIV (\<lambda>x. if x \<in> s then 1 else 0)" apply-
-    apply(rule LIMSEQ_unique[OF _ **(2)]) unfolding measure_integral_univ[THEN sym,OF s(2)] using * .
+    apply(rule LIMSEQ_unique[OF _ **(2)]) unfolding measure_integral_univ[THEN sym,OF s] using * .
   show ?thesis unfolding has_gmeasure * apply(rule integrable_integral) using ** by auto
 qed
 
-lemma has_lmeasure_unique: "s has_lmeasure m1 \<Longrightarrow> s has_lmeasure m2 \<Longrightarrow> m1 = m2"
-  unfolding has_lmeasure_def apply(rule Lim_unique) using trivial_limit_sequentially by auto
-
-lemma lmeasure_unique[intro]: assumes "A has_lmeasure m" shows "lmeasure A = m"
-  using assms unfolding lmeasure_def lmeasurable_def apply-
-  apply(rule some_equality) defer apply(rule has_lmeasure_unique) by auto
-
-lemma glmeasurable_finite: assumes "lmeasurable s" "lmeasure s \<noteq> \<omega>" 
+lemma lmeasure_finite_gmeasurable: assumes "s \<in> sets lebesgue" "lmeasure s \<noteq> \<omega>"
   shows "gmeasurable s"
-proof-  have "\<exists>B. \<forall>n. gmeasure (s \<inter> cube n) \<le> B"
-  proof(rule ccontr) case goal1
-    note as = this[unfolded not_ex not_all not_le]
-    have "s has_lmeasure \<omega>" apply- apply(rule has_lmeasureI[OF assms(1)])
-      unfolding Lim_omega
-    proof fix B::real
-      from as[rule_format,of B] guess N .. note N = this
-      have "\<And>n. N \<le> n \<Longrightarrow> B \<le> gmeasure (s \<inter> cube n)"
-        apply(rule order_trans[where y="gmeasure (s \<inter> cube N)"]) defer
-        apply(rule measure_subset) prefer 3
-        using cube_subset N assms(1)[unfolded lmeasurable_def] by auto
-      thus "\<exists>N. \<forall>n\<ge>N. Real B \<le> Real (gmeasure (s \<inter> cube n))" apply-
-        apply(subst Real_max') apply(rule_tac x=N in exI,safe)
-        unfolding pinfreal_less_eq apply(subst if_P) by auto
-    qed note lmeasure_unique[OF this]
-    thus False using assms(2) by auto
-  qed then guess B .. note B=this
+proof (cases "lmeasure s")
+  case (preal m) from lmeasure_finite_has_gmeasure[OF `s \<in> sets lebesgue` this]
+  show ?thesis unfolding gmeasurable_def by auto
+qed (insert assms, auto)
 
-  show ?thesis apply(rule gmeasurable_nested_unions[of "\<lambda>n. s \<inter> cube n",
-    unfolded Union_inter_cube,THEN conjunct1, where B1=B])
-  proof- fix n::nat
-    show " gmeasurable (s \<inter> cube n)" using assms by auto
-    show "gmeasure (s \<inter> cube n) \<le> B" using B by auto
-    show "s \<inter> cube n \<subseteq> s \<inter> cube (Suc n)"
-      by (rule Int_mono) (simp_all add: cube_subset)
-  qed
-qed
-
-lemma lmeasure_empty[intro]:"lmeasure {} = 0"
-  apply(rule lmeasure_unique)
-  unfolding has_lmeasure_def by auto
-
-lemma lmeasurableI[dest]:"s has_lmeasure m \<Longrightarrow> lmeasurable s"
-  unfolding has_lmeasure_def by auto
-
-lemma has_gmeasure_has_lmeasure: assumes "s has_gmeasure m"
-  shows "s has_lmeasure (Real m)"
-proof- have gmea:"gmeasurable s" using assms by auto
+lemma has_gmeasure_lmeasure: assumes "s has_gmeasure m"
+  shows "lmeasure s = Real m"
+proof-
+  have gmea:"gmeasurable s" using assms by auto
+  then have s: "s \<in> sets lebesgue" by auto
   have m:"m \<ge> 0" using assms by auto
   have *:"m = gmeasure (\<Union>{s \<inter> cube n |n. n \<in> UNIV})" unfolding Union_inter_cube
     using assms by(rule measure_unique[THEN sym])
-  show ?thesis unfolding has_lmeasure_def
-    apply(rule,rule measurable_imp_lmeasurable[OF gmea])
-    apply(subst lim_Real) apply(rule,rule,rule m) unfolding *
+  show ?thesis
+    unfolding lmeasure_iff_LIMSEQ[OF s `0 \<le> m`] unfolding *
     apply(rule gmeasurable_nested_unions[THEN conjunct2, where B1="gmeasure s"])
   proof- fix n::nat show *:"gmeasurable (s \<inter> cube n)"
       using gmeasurable_inter[OF gmea gmeasurable_cube] .
@@ -266,287 +418,26 @@
       apply(rule * gmea)+ by auto
     show "s \<inter> cube n \<subseteq> s \<inter> cube (Suc n)" using cube_subset[of n "Suc n"] by auto
   qed
-qed    
-    
-lemma gmeasure_lmeasure: assumes "gmeasurable s" shows "lmeasure s = Real (gmeasure s)"
-proof- note has_gmeasure_measureI[OF assms]
-  note has_gmeasure_has_lmeasure[OF this]
-  thus ?thesis by(rule lmeasure_unique)
-qed
-
-lemma has_lmeasure_lmeasure: "lmeasurable s \<longleftrightarrow> s has_lmeasure (lmeasure s)" (is "?l = ?r")
-proof assume ?l let ?f = "\<lambda>n. Real (gmeasure (s \<inter> cube n))"
-  have "\<forall>n m. n\<ge>m \<longrightarrow> ?f n \<ge> ?f m" unfolding pinfreal_less_eq apply safe
-    apply(subst if_P) defer apply(rule measure_subset) prefer 3
-    apply(drule cube_subset) using `?l` by auto
-  from lim_pinfreal_increasing[OF this] guess l . note l=this
-  hence "s has_lmeasure l" using `?l` apply-apply(rule has_lmeasureI) by auto
-  thus ?r using lmeasure_unique by auto
-next assume ?r thus ?l unfolding has_lmeasure_def by auto
-qed
-
-lemma lmeasure_subset[dest]: assumes "lmeasurable s" "lmeasurable t" "s \<subseteq> t"
-  shows "lmeasure s \<le> lmeasure t"
-proof(cases "lmeasure t = \<omega>")
-  case False have som:"lmeasure s \<noteq> \<omega>"
-  proof(rule ccontr,unfold not_not) assume as:"lmeasure s = \<omega>"
-    have "t has_lmeasure \<omega>" using assms(2) apply(rule has_lmeasureI)
-      unfolding Lim_omega
-    proof case goal1
-      note assms(1)[unfolded has_lmeasure_lmeasure]
-      note has_lmeasureD(3)[OF this,unfolded as Lim_omega,rule_format,of B]
-      then guess N .. note N = this
-      show ?case apply(rule_tac x=N in exI) apply safe
-        apply(rule order_trans) apply(rule N[rule_format],assumption)
-        unfolding pinfreal_less_eq apply(subst if_P)defer
-        apply(rule measure_subset) using assms by auto
-    qed
-    thus False using lmeasure_unique False by auto
-  qed
-
-  note assms(1)[unfolded has_lmeasure_lmeasure] note has_lmeasureD(3)[OF this]
-  hence "(\<lambda>n. Real (gmeasure (s \<inter> cube n))) ----> Real (real (lmeasure s))"
-    unfolding Real_real'[OF som] .
-  hence l1:"(\<lambda>n. gmeasure (s \<inter> cube n)) ----> real (lmeasure s)"
-    apply-apply(subst(asm) lim_Real) by auto
-
-  note assms(2)[unfolded has_lmeasure_lmeasure] note has_lmeasureD(3)[OF this]
-  hence "(\<lambda>n. Real (gmeasure (t \<inter> cube n))) ----> Real (real (lmeasure t))"
-    unfolding Real_real'[OF False] .
-  hence l2:"(\<lambda>n. gmeasure (t \<inter> cube n)) ----> real (lmeasure t)"
-    apply-apply(subst(asm) lim_Real) by auto
-
-  have "real (lmeasure s) \<le> real (lmeasure t)" apply(rule LIMSEQ_le[OF l1 l2])
-    apply(rule_tac x=0 in exI,safe) apply(rule measure_subset) using assms by auto
-  hence "Real (real (lmeasure s)) \<le> Real (real (lmeasure t))"
-    unfolding pinfreal_less_eq by auto
-  thus ?thesis unfolding Real_real'[OF som] Real_real'[OF False] .
-qed auto
-
-lemma has_lmeasure_negligible_unions_image:
-  assumes "finite s" "\<And>x. x \<in> s ==> lmeasurable(f x)"
-  "\<And>x y. x \<in> s \<Longrightarrow> y \<in> s \<Longrightarrow> x \<noteq> y \<Longrightarrow> negligible((f x) \<inter> (f y))"
-  shows "(\<Union> (f ` s)) has_lmeasure (setsum (\<lambda>x. lmeasure(f x)) s)"
-  unfolding has_lmeasure_def
-proof show lmeaf:"lmeasurable (\<Union>f ` s)" apply(rule lmeasurable_finite_unions)
-    using assms(1-2) by auto
-  show "(\<lambda>n. Real (gmeasure (\<Union>f ` s \<inter> cube n))) ----> (\<Sum>x\<in>s. lmeasure (f x))" (is ?l)
-  proof(cases "\<exists>x\<in>s. lmeasure (f x) = \<omega>")
-    case False hence *:"(\<Sum>x\<in>s. lmeasure (f x)) \<noteq> \<omega>" apply-
-      apply(rule setsum_neq_omega) using assms(1) by auto
-    have gmea:"\<And>x. x\<in>s \<Longrightarrow> gmeasurable (f x)" apply(rule glmeasurable_finite) using False assms(2) by auto
-    have "(\<Sum>x\<in>s. lmeasure (f x)) = (\<Sum>x\<in>s. Real (gmeasure (f x)))" apply(rule setsum_cong2)
-      apply(rule gmeasure_lmeasure) using False assms(2) gmea by auto
-    also have "... = Real (\<Sum>x\<in>s. (gmeasure (f x)))" apply(rule setsum_Real) by auto
-    finally have sum:"(\<Sum>x\<in>s. lmeasure (f x)) = Real (\<Sum>x\<in>s. gmeasure (f x))" .
-    have sum_0:"(\<Sum>x\<in>s. gmeasure (f x)) \<ge> 0" apply(rule setsum_nonneg) by auto
-    have int_un:"\<Union>f ` s has_gmeasure (\<Sum>x\<in>s. gmeasure (f x))"
-      apply(rule has_gmeasure_negligible_unions_image) using assms gmea by auto
-
-    have unun:"\<Union>{\<Union>f ` s \<inter> cube n |n. n \<in> UNIV} = \<Union>f ` s" unfolding simple_image 
-    proof safe fix x y assume as:"x \<in> f y" "y \<in> s"
-      from mem_big_cube[of x] guess n . note n=this
-      thus "x \<in> \<Union>range (\<lambda>n. \<Union>f ` s \<inter> cube n)" unfolding Union_iff
-        apply-apply(rule_tac x="\<Union>f ` s \<inter> cube n" in bexI) using as by auto
-    qed
-    show ?l apply(subst Real_real'[OF *,THEN sym])apply(subst lim_Real) 
-      apply rule apply rule unfolding sum real_Real if_P[OF sum_0] apply(rule sum_0)
-      unfolding measure_unique[OF int_un,THEN sym] apply(subst(2) unun[THEN sym])
-      apply(rule has_gmeasure_nested_unions[THEN conjunct2])
-    proof- fix n::nat
-      show *:"gmeasurable (\<Union>f ` s \<inter> cube n)" using lmeaf unfolding lmeasurable_def by auto
-      thus "gmeasure (\<Union>f ` s \<inter> cube n) \<le> gmeasure (\<Union>f ` s)"
-        apply(rule measure_subset) using int_un by auto
-      show "\<Union>f ` s \<inter> cube n \<subseteq> \<Union>f ` s \<inter> cube (Suc n)"
-        using cube_subset[of n "Suc n"] by auto
-    qed
-
-  next case True then guess X .. note X=this
-    hence sum:"(\<Sum>x\<in>s. lmeasure (f x)) = \<omega>" using setsum_\<omega>[THEN iffD2, of s] assms by fastsimp
-    show ?l unfolding sum Lim_omega
-    proof fix B::real
-      have Xm:"(f X) has_lmeasure \<omega>" using X by (metis assms(2) has_lmeasure_lmeasure)
-      note this[unfolded has_lmeasure_def,THEN conjunct2, unfolded Lim_omega]
-      from this[rule_format,of B] guess N .. note N=this[rule_format]
-      show "\<exists>N. \<forall>n\<ge>N. Real B \<le> Real (gmeasure (\<Union>f ` s \<inter> cube n))"
-        apply(rule_tac x=N in exI)
-      proof safe case goal1 show ?case apply(rule order_trans[OF N[OF goal1]])
-          unfolding pinfreal_less_eq apply(subst if_P) defer
-          apply(rule measure_subset) using has_lmeasureD(2)[OF Xm]
-          using lmeaf unfolding lmeasurable_def using X(1) by auto
-      qed qed qed qed
-
-lemma has_lmeasure_negligible_unions:
-  assumes"finite f" "\<And>s. s \<in> f ==> s has_lmeasure (m s)"
-  "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> negligible (s\<inter>t)"
-  shows "(\<Union> f) has_lmeasure (setsum m f)"
-proof- have *:"setsum m f = setsum lmeasure f" apply(rule setsum_cong2)
-    apply(subst lmeasure_unique[OF assms(2)]) by auto
-  show ?thesis unfolding *
-    apply(rule has_lmeasure_negligible_unions_image[where s=f and f=id,unfolded image_id id_apply])
-    using assms by auto
-qed
-
-lemma has_lmeasure_disjoint_unions:
-  assumes"finite f" "\<And>s. s \<in> f ==> s has_lmeasure (m s)"
-  "\<And>s t. s \<in> f \<Longrightarrow> t \<in> f \<Longrightarrow> s \<noteq> t ==> s \<inter> t = {}"
-  shows "(\<Union> f) has_lmeasure (setsum m f)"
-proof- have *:"setsum m f = setsum lmeasure f" apply(rule setsum_cong2)
-    apply(subst lmeasure_unique[OF assms(2)]) by auto
-  show ?thesis unfolding *
-    apply(rule has_lmeasure_negligible_unions_image[where s=f and f=id,unfolded image_id id_apply])
-    using assms by auto
 qed
 
-lemma has_lmeasure_nested_unions:
-  assumes "\<And>n. lmeasurable(s n)" "\<And>n. s(n) \<subseteq> s(Suc n)"
-  shows "lmeasurable(\<Union> { s n | n. n \<in> UNIV }) \<and>
-  (\<lambda>n. lmeasure(s n)) ----> lmeasure(\<Union> { s(n) | n. n \<in> UNIV })" (is "?mea \<and> ?lim")
-proof- have cube:"\<And>m. \<Union> { s(n) | n. n \<in> UNIV } \<inter> cube m = \<Union> { s(n) \<inter> cube m | n. n \<in> UNIV }" by blast
-  have 3:"\<And>n. \<forall>m\<ge>n. s n \<subseteq> s m" apply(rule transitive_stepwise_le) using assms(2) by auto
-  have mea:"?mea" unfolding lmeasurable_def cube apply rule 
-    apply(rule_tac B1="gmeasure (cube n)" in has_gmeasure_nested_unions[THEN conjunct1])
-    prefer 3 apply rule using assms(1) unfolding lmeasurable_def
-    by(auto intro!:assms(2)[unfolded subset_eq,rule_format])
-  show ?thesis apply(rule,rule mea)
-  proof(cases "lmeasure(\<Union> { s(n) | n. n \<in> UNIV }) = \<omega>")
-    case True show ?lim unfolding True Lim_omega
-    proof(rule ccontr) case goal1 note this[unfolded not_all not_ex]
-      hence "\<exists>B. \<forall>n. \<exists>m\<ge>n. Real B > lmeasure (s m)" by(auto simp add:not_le)
-      from this guess B .. note B=this[rule_format]
-
-      have "\<forall>n. gmeasurable (s n) \<and> gmeasure (s n) \<le> max B 0" 
-      proof safe fix n::nat from B[of n] guess m .. note m=this
-        hence *:"lmeasure (s n) < Real B" apply-apply(rule le_less_trans)
-          apply(rule lmeasure_subset[OF assms(1,1)]) apply(rule 3[rule_format]) by auto
-        thus **:"gmeasurable (s n)" apply-apply(rule glmeasurable_finite[OF assms(1)]) by auto
-        thus "gmeasure (s n) \<le> max B 0"  using * unfolding gmeasure_lmeasure[OF **] Real_max'[of B]
-          unfolding pinfreal_less apply- apply(subst(asm) if_P) by auto
-      qed
-      hence "\<And>n. gmeasurable (s n)" "\<And>n. gmeasure (s n) \<le> max B 0" by auto
-      note g = conjunctD2[OF has_gmeasure_nested_unions[of s, OF this assms(2)]]
-      show False using True unfolding gmeasure_lmeasure[OF g(1)] by auto
-    qed
-  next let ?B = "lmeasure (\<Union>{s n |n. n \<in> UNIV})"
-    case False note gmea_lim = glmeasurable_finite[OF mea this]
-    have ls:"\<And>n. lmeasure (s n) \<le> lmeasure (\<Union>{s n |n. n \<in> UNIV})"
-      apply(rule lmeasure_subset) using assms(1) mea by auto
-    have "\<And>n. lmeasure (s n) \<noteq> \<omega>"
-    proof(rule ccontr,safe) case goal1
-      show False using False ls[of n] unfolding goal1 by auto
-    qed
-    note gmea = glmeasurable_finite[OF assms(1) this]
-
-    have "\<And>n. gmeasure (s n) \<le> real ?B"  unfolding gmeasure_lmeasure[OF gmea_lim]
-      unfolding real_Real apply(subst if_P,rule) apply(rule measure_subset)
-      using gmea gmea_lim by auto
-    note has_gmeasure_nested_unions[of s, OF gmea this assms(2)]
-    thus ?lim unfolding gmeasure_lmeasure[OF gmea] gmeasure_lmeasure[OF gmea_lim]
-      apply-apply(subst lim_Real) by auto
-  qed
-qed
-
-lemma has_lmeasure_countable_negligible_unions: 
-  assumes "\<And>n. lmeasurable(s n)" "\<And>m n. m \<noteq> n \<Longrightarrow> negligible(s m \<inter> s n)"
-  shows "(\<lambda>m. setsum (\<lambda>n. lmeasure(s n)) {..m}) ----> (lmeasure(\<Union> { s(n) |n. n \<in> UNIV }))"
-proof- have *:"\<And>n. (\<Union> (s ` {0..n})) has_lmeasure (setsum (\<lambda>k. lmeasure(s k)) {0..n})"
-    apply(rule has_lmeasure_negligible_unions_image) using assms by auto
-  have **:"(\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV}) = (\<Union>{s n |n. n \<in> UNIV})" unfolding simple_image by fastsimp
-  have "lmeasurable (\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV}) \<and>
-    (\<lambda>n. lmeasure (\<Union>(s ` {0..n}))) ----> lmeasure (\<Union>{\<Union>s ` {0..n} |n. n \<in> UNIV})"
-    apply(rule has_lmeasure_nested_unions) apply(rule has_lmeasureD(1)[OF *])
-    apply(rule Union_mono,rule image_mono) by auto
-  note lem = conjunctD2[OF this,unfolded **] 
-  show ?thesis using lem(2) unfolding lmeasure_unique[OF *] unfolding atLeast0AtMost .
-qed
-
-lemma lmeasure_eq_0: assumes "negligible s" shows "lmeasure s = 0"
-proof- note mea=negligible_imp_lmeasurable[OF assms]
-  have *:"\<And>n. (gmeasure (s \<inter> cube n)) = 0" 
-    unfolding gmeasurable_measure_eq_0[OF mea[unfolded lmeasurable_def,rule_format]]
-    using assms by auto
-  show ?thesis
-    apply(rule lmeasure_unique) unfolding has_lmeasure_def
-    apply(rule,rule mea) unfolding * by auto
+lemma has_gmeasure_iff_lmeasure:
+  "A has_gmeasure m \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m)"
+proof
+  assume "A has_gmeasure m"
+  with has_gmeasure_lmeasure[OF this]
+  have "gmeasurable A" "0 \<le> m" "lmeasure A = Real m" by auto
+  then show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m" by auto
+next
+  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lmeasure A = Real m"
+  then show "A has_gmeasure m" by (intro lmeasure_finite_has_gmeasure) auto
 qed
 
-lemma negligible_img_gmeasurable: fixes s::"'a::ordered_euclidean_space set"
-  assumes "negligible s" shows "gmeasurable s"
-  apply(rule glmeasurable_finite)
-  using lmeasure_eq_0[OF assms] negligible_imp_lmeasurable[OF assms] by auto
-
-
-
-
-section {* Instantiation of _::euclidean_space as measure_space *}
-
-definition lebesgue_space :: "'a::ordered_euclidean_space algebra" where
-  "lebesgue_space = \<lparr> space = UNIV, sets = lmeasurable \<rparr>"
-
-lemma lebesgue_measurable[simp]:"A \<in> sets lebesgue_space \<longleftrightarrow> lmeasurable A"
-  unfolding lebesgue_space_def by(auto simp: mem_def)
-
-lemma mem_gmeasurable[simp]: "A \<in> gmeasurable \<longleftrightarrow> gmeasurable A"
-  unfolding mem_def ..
-
-interpretation lebesgue: measure_space lebesgue_space lmeasure
-  apply(intro_locales) unfolding measure_space_axioms_def countably_additive_def
-  unfolding sigma_algebra_axioms_def algebra_def
-  unfolding lebesgue_measurable
-proof safe
-  fix A::"nat => _" assume as:"range A \<subseteq> sets lebesgue_space" "disjoint_family A"
-    "lmeasurable (UNION UNIV A)"
-  have *:"UNION UNIV A = \<Union>range A" by auto
-  show "(\<Sum>\<^isub>\<infinity>n. lmeasure (A n)) = lmeasure (UNION UNIV A)" 
-    unfolding psuminf_def apply(rule SUP_Lim_pinfreal)
-  proof- fix n m::nat assume mn:"m\<le>n"
-    have *:"\<And>m. (\<Sum>n<m. lmeasure (A n)) = lmeasure (\<Union>A ` {..<m})"
-      apply(subst lmeasure_unique[OF has_lmeasure_negligible_unions[where m=lmeasure]])
-      apply(rule finite_imageI) apply rule apply(subst has_lmeasure_lmeasure[THEN sym])
-    proof- fix m::nat
-      show "(\<Sum>n<m. lmeasure (A n)) = setsum lmeasure (A ` {..<m})"
-        apply(subst setsum_reindex_nonzero) unfolding o_def apply rule
-        apply(rule lmeasure_eq_0) using as(2) unfolding disjoint_family_on_def
-        apply(erule_tac x=x in ballE,safe,erule_tac x=y in ballE) by auto
-    next fix m s assume "s \<in> A ` {..<m}"
-      hence "s \<in> range A" by auto thus "lmeasurable s" using as(1) by fastsimp
-    next fix m s t assume st:"s  \<in> A ` {..<m}" "t \<in> A ` {..<m}" "s \<noteq> t"
-      from st(1-2) guess sa ta unfolding image_iff apply-by(erule bexE)+ note a=this
-      from st(3) have "sa \<noteq> ta" unfolding a by auto
-      thus "negligible (s \<inter> t)" 
-        using as(2) unfolding disjoint_family_on_def a
-        apply(erule_tac x=sa in ballE,erule_tac x=ta in ballE) by auto
-    qed
-
-    have "\<And>m. lmeasurable (\<Union>A ` {..<m})"  apply(rule lmeasurable_finite_unions)
-      apply(rule finite_imageI,rule) using as(1) by fastsimp
-    from this this show "(\<Sum>n<m. lmeasure (A n)) \<le> (\<Sum>n<n. lmeasure (A n))" unfolding *
-      apply(rule lmeasure_subset) apply(rule Union_mono) apply(rule image_mono) using mn by auto
-    
-  next have *:"UNION UNIV A = \<Union>{A n |n. n \<in> UNIV}" by auto
-    show "(\<lambda>n. \<Sum>n<n. lmeasure (A n)) ----> lmeasure (UNION UNIV A)"
-      apply(rule LIMSEQ_imp_Suc) unfolding lessThan_Suc_atMost *
-      apply(rule has_lmeasure_countable_negligible_unions)
-      using as unfolding disjoint_family_on_def subset_eq by auto
-  qed
-
-next show "lmeasure {} = 0" by auto
-next fix A::"nat => _" assume as:"range A \<subseteq> sets lebesgue_space"
-  have *:"UNION UNIV A = (\<Union>{A n |n. n \<in> UNIV})" unfolding simple_image by auto
-  show "lmeasurable (UNION UNIV A)" unfolding * using as unfolding subset_eq
-    using lmeasurable_countable_unions_strong[of A] by auto
-qed(auto simp: lebesgue_space_def mem_def)
-
-
-
-lemma lmeasurbale_closed_interval[intro]:
-  "lmeasurable {a..b::'a::ordered_euclidean_space}"
-  unfolding lmeasurable_def cube_def inter_interval by auto
-
-lemma space_lebesgue_space[simp]:"space lebesgue_space = UNIV"
-  unfolding lebesgue_space_def by auto
-
-abbreviation "gintegral \<equiv> Integration.integral"
+lemma gmeasure_lmeasure: assumes "gmeasurable s" shows "lmeasure s = Real (gmeasure s)"
+proof -
+  note has_gmeasure_measureI[OF assms]
+  note has_gmeasure_lmeasure[OF this]
+  thus ?thesis .
+qed
 
 lemma lebesgue_simple_function_indicator:
   fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
@@ -556,21 +447,614 @@
 
 lemma lmeasure_gmeasure:
   "gmeasurable s \<Longrightarrow> gmeasure s = real (lmeasure s)"
-  apply(subst gmeasure_lmeasure) by auto
+  by (subst gmeasure_lmeasure) auto
 
 lemma lmeasure_finite: assumes "gmeasurable s" shows "lmeasure s \<noteq> \<omega>"
   using gmeasure_lmeasure[OF assms] by auto
 
-lemma negligible_lmeasure: assumes "lmeasurable s"
-  shows "lmeasure s = 0 \<longleftrightarrow> negligible s" (is "?l = ?r")
-proof assume ?l
-  hence *:"gmeasurable s" using glmeasurable_finite[of s] assms by auto
-  show ?r unfolding gmeasurable_measure_eq_0[THEN sym,OF *]
-    unfolding lmeasure_gmeasure[OF *] using `?l` by auto
-next assume ?r
-  note g=negligible_img_gmeasurable[OF this] and measure_eq_0[OF this]
-  hence "real (lmeasure s) = 0" using lmeasure_gmeasure[of s] by auto
-  thus ?l using lmeasure_finite[OF g] apply- apply(rule real_0_imp_eq_0) by auto
+lemma negligible_iff_lebesgue_null_sets:
+  "negligible A \<longleftrightarrow> A \<in> lebesgue.null_sets"
+proof
+  assume "negligible A"
+  from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0]
+  show "A \<in> lebesgue.null_sets" by auto
+next
+  assume A: "A \<in> lebesgue.null_sets"
+  then have *:"gmeasurable A" using lmeasure_finite_gmeasurable[of A] by auto
+  show "negligible A"
+    unfolding gmeasurable_measure_eq_0[OF *, symmetric]
+    unfolding lmeasure_gmeasure[OF *] using A by auto
+qed
+
+lemma
+  fixes a b ::"'a::ordered_euclidean_space"
+  shows lmeasure_atLeastAtMost[simp]: "lmeasure {a..b} = Real (content {a..b})"
+    and lmeasure_greaterThanLessThan[simp]: "lmeasure {a <..< b} = Real (content {a..b})"
+  using has_gmeasure_interval[of a b] by (auto intro!: has_gmeasure_lmeasure)
+
+lemma lmeasure_cube:
+  "lmeasure (cube n::('a::ordered_euclidean_space) set) = (Real ((2 * real n) ^ (DIM('a))))"
+  by (intro has_gmeasure_lmeasure) auto
+
+lemma lmeasure_UNIV[intro]: "lmeasure UNIV = \<omega>"
+  unfolding lmeasure_def SUP_\<omega>
+proof (intro allI impI)
+  fix x assume "x < \<omega>"
+  then obtain r where r: "x = Real r" "0 \<le> r" by (cases x) auto
+  then obtain n where n: "r < of_nat n" using ex_less_of_nat by auto
+  show "\<exists>i\<in>UNIV. x < Real (gmeasure (UNIV \<inter> cube i))"
+  proof (intro bexI[of _ n])
+    have "x < Real (of_nat n)" using n r by auto
+    also have "Real (of_nat n) \<le> Real (gmeasure (UNIV \<inter> cube n))"
+      using gmeasure_cube_ge_n[of n] by (auto simp: real_eq_of_nat[symmetric])
+    finally show "x < Real (gmeasure (UNIV \<inter> cube n))" .
+  qed auto
+qed
+
+lemma atLeastAtMost_singleton_euclidean[simp]:
+  fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}"
+  by (force simp: eucl_le[where 'a='a] euclidean_eq[where 'a='a])
+
+lemma content_singleton[simp]: "content {a} = 0"
+proof -
+  have "content {a .. a} = 0"
+    by (subst content_closed_interval) auto
+  then show ?thesis by simp
+qed
+
+lemma lmeasure_singleton[simp]:
+  fixes a :: "'a::ordered_euclidean_space" shows "lmeasure {a} = 0"
+  using has_gmeasure_interval[of a a] unfolding zero_pinfreal_def
+  by (intro has_gmeasure_lmeasure)
+     (simp add: content_closed_interval DIM_positive)
+
+declare content_real[simp]
+
+lemma
+  fixes a b :: real
+  shows lmeasure_real_greaterThanAtMost[simp]:
+    "lmeasure {a <.. b} = Real (if a \<le> b then b - a else 0)"
+proof cases
+  assume "a < b"
+  then have "lmeasure {a <.. b} = lmeasure {a <..< b} + lmeasure {b}"
+    by (subst lebesgue.measure_additive)
+       (auto intro!: lebesgueI_borel arg_cong[where f=lmeasure])
+  then show ?thesis by auto
+qed auto
+
+lemma
+  fixes a b :: real
+  shows lmeasure_real_atLeastLessThan[simp]:
+    "lmeasure {a ..< b} = Real (if a \<le> b then b - a else 0)" (is ?eqlt)
+proof cases
+  assume "a < b"
+  then have "lmeasure {a ..< b} = lmeasure {a} + lmeasure {a <..< b}"
+    by (subst lebesgue.measure_additive)
+       (auto intro!: lebesgueI_borel arg_cong[where f=lmeasure])
+  then show ?thesis by auto
+qed auto
+
+interpretation borel: measure_space borel lmeasure
+proof
+  show "countably_additive borel lmeasure"
+    using lebesgue.ca unfolding countably_additive_def
+    apply safe apply (erule_tac x=A in allE) by auto
+qed auto
+
+interpretation borel: sigma_finite_measure borel lmeasure
+proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
+  show "range cube \<subseteq> sets borel" by (auto intro: borel_closed)
+  { fix x have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
+  thus "(\<Union>i. cube i) = space borel" by auto
+  show "\<forall>i. lmeasure (cube i) \<noteq> \<omega>" unfolding lmeasure_cube by auto
+qed
+
+interpretation lebesgue: sigma_finite_measure lebesgue lmeasure
+proof
+  from borel.sigma_finite guess A ..
+  moreover then have "range A \<subseteq> sets lebesgue" using lebesgueI_borel by blast
+  ultimately show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. lmeasure (A i) \<noteq> \<omega>)"
+    by auto
+qed
+
+lemma simple_function_has_integral:
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  assumes f:"lebesgue.simple_function f"
+  and f':"\<forall>x. f x \<noteq> \<omega>"
+  and om:"\<forall>x\<in>range f. lmeasure (f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
+  unfolding lebesgue.simple_integral_def
+  apply(subst lebesgue_simple_function_indicator[OF f])
+proof- case goal1
+  have *:"\<And>x. \<forall>y\<in>range f. y * indicator (f -` {y}) x \<noteq> \<omega>"
+    "\<forall>x\<in>range f. x * lmeasure (f -` {x} \<inter> UNIV) \<noteq> \<omega>"
+    using f' om unfolding indicator_def by auto
+  show ?case unfolding space_lebesgue real_of_pinfreal_setsum'[OF *(1),THEN sym]
+    unfolding real_of_pinfreal_setsum'[OF *(2),THEN sym]
+    unfolding real_of_pinfreal_setsum space_lebesgue
+    apply(rule has_integral_setsum)
+  proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD)
+    fix y::'a show "((\<lambda>x. real (f y * indicator (f -` {f y}) x)) has_integral
+      real (f y * lmeasure (f -` {f y} \<inter> UNIV))) UNIV"
+    proof(cases "f y = 0") case False
+      have mea:"gmeasurable (f -` {f y})" apply(rule lmeasure_finite_gmeasurable)
+        using assms unfolding lebesgue.simple_function_def using False by auto
+      have *:"\<And>x. real (indicator (f -` {f y}) x::pinfreal) = (if x \<in> f -` {f y} then 1 else 0)" by auto
+      show ?thesis unfolding real_of_pinfreal_mult[THEN sym]
+        apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def])
+        unfolding Int_UNIV_right lmeasure_gmeasure[OF mea,THEN sym]
+        unfolding measure_integral_univ[OF mea] * apply(rule integrable_integral)
+        unfolding gmeasurable_integrable[THEN sym] using mea .
+    qed auto
+  qed qed
+
+lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s"
+  unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI)
+  using assms by auto
+
+lemma simple_function_has_integral':
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  assumes f:"lebesgue.simple_function f"
+  and i: "lebesgue.simple_integral f \<noteq> \<omega>"
+  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.simple_integral f))) UNIV"
+proof- let ?f = "\<lambda>x. if f x = \<omega> then 0 else f x"
+  { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this
+  have **:"{x. f x \<noteq> ?f x} = f -` {\<omega>}" by auto
+  have **:"lmeasure {x\<in>space lebesgue. f x \<noteq> ?f x} = 0"
+    using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**)
+  show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **])
+    apply(rule lebesgue.simple_function_compose1[OF f])
+    unfolding * defer apply(rule simple_function_has_integral)
+  proof-
+    show "lebesgue.simple_function ?f"
+      using lebesgue.simple_function_compose1[OF f] .
+    show "\<forall>x. ?f x \<noteq> \<omega>" by auto
+    show "\<forall>x\<in>range ?f. lmeasure (?f -` {x} \<inter> UNIV) = \<omega> \<longrightarrow> x = 0"
+    proof (safe, simp, safe, rule ccontr)
+      fix y assume "f y \<noteq> \<omega>" "f y \<noteq> 0"
+      hence "(\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y} = f -` {f y}"
+        by (auto split: split_if_asm)
+      moreover assume "lmeasure ((\<lambda>x. if f x = \<omega> then 0 else f x) -` {if f y = \<omega> then 0 else f y}) = \<omega>"
+      ultimately have "lmeasure (f -` {f y}) = \<omega>" by simp
+      moreover
+      have "f y * lmeasure (f -` {f y}) \<noteq> \<omega>" using i f
+        unfolding lebesgue.simple_integral_def setsum_\<omega> lebesgue.simple_function_def
+        by auto
+      ultimately have "f y = 0" by (auto split: split_if_asm)
+      then show False using `f y \<noteq> 0` by simp
+    qed
+  qed
+qed
+
+lemma (in measure_space) positive_integral_monotone_convergence:
+  fixes f::"nat \<Rightarrow> 'a \<Rightarrow> pinfreal"
+  assumes i: "\<And>i. f i \<in> borel_measurable M" and mono: "\<And>x. mono (\<lambda>n. f n x)"
+  and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
+  shows "u \<in> borel_measurable M"
+  and "(\<lambda>i. positive_integral (f i)) ----> positive_integral u" (is ?ilim)
+proof -
+  from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u]
+  show ?ilim using mono lim i by auto
+  have "(SUP i. f i) = u" using mono lim SUP_Lim_pinfreal
+    unfolding fun_eq_iff SUPR_fun_expand mono_def by auto
+  moreover have "(SUP i. f i) \<in> borel_measurable M"
+    using i by (rule borel_measurable_SUP)
+  ultimately show "u \<in> borel_measurable M" by simp
+qed
+
+lemma positive_integral_has_integral:
+  fixes f::"'a::ordered_euclidean_space => pinfreal"
+  assumes f:"f \<in> borel_measurable lebesgue"
+  and int_om:"lebesgue.positive_integral f \<noteq> \<omega>"
+  and f_om:"\<forall>x. f x \<noteq> \<omega>" (* TODO: remove this *)
+  shows "((\<lambda>x. real (f x)) has_integral (real (lebesgue.positive_integral f))) UNIV"
+proof- let ?i = "lebesgue.positive_integral f"
+  from lebesgue.borel_measurable_implies_simple_function_sequence[OF f]
+  guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2)
+  let ?u = "\<lambda>i x. real (u i x)" and ?f = "\<lambda>x. real (f x)"
+  have u_simple:"\<And>k. lebesgue.simple_integral (u k) = lebesgue.positive_integral (u k)"
+    apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) ..
+  have int_u_le:"\<And>k. lebesgue.simple_integral (u k) \<le> lebesgue.positive_integral f"
+    unfolding u_simple apply(rule lebesgue.positive_integral_mono)
+    using isoton_Sup[OF u(3)] unfolding le_fun_def by auto
+  have u_int_om:"\<And>i. lebesgue.simple_integral (u i) \<noteq> \<omega>"
+  proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed
+
+  note u_int = simple_function_has_integral'[OF u(1) this]
+  have "(\<lambda>x. real (f x)) integrable_on UNIV \<and>
+    (\<lambda>k. Integration.integral UNIV (\<lambda>x. real (u k x))) ----> Integration.integral UNIV (\<lambda>x. real (f x))"
+    apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int)
+  proof safe case goal1 show ?case apply(rule real_of_pinfreal_mono) using u(2,3) by auto
+  next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym])
+      prefer 3 apply(subst Real_real') defer apply(subst Real_real')
+      using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto
+  next case goal3
+    show ?case apply(rule bounded_realI[where B="real (lebesgue.positive_integral f)"])
+      apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int)
+      unfolding integral_unique[OF u_int] defer apply(rule real_of_pinfreal_mono[OF _ int_u_le])
+      using u int_om by auto
+  qed note int = conjunctD2[OF this]
+
+  have "(\<lambda>i. lebesgue.simple_integral (u i)) ----> ?i" unfolding u_simple
+    apply(rule lebesgue.positive_integral_monotone_convergence(2))
+    apply(rule lebesgue.borel_measurable_simple_function[OF u(1)])
+    using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto
+  hence "(\<lambda>i. real (lebesgue.simple_integral (u i))) ----> real ?i" apply-
+    apply(subst lim_Real[THEN sym]) prefer 3
+    apply(subst Real_real') defer apply(subst Real_real')
+    using u f_om int_om u_int_om by auto
+  note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]]
+  show ?thesis unfolding * by(rule integrable_integral[OF int(1)])
+qed
+
+lemma lebesgue_integral_has_integral:
+  fixes f::"'a::ordered_euclidean_space => real"
+  assumes f:"lebesgue.integrable f"
+  shows "(f has_integral (lebesgue.integral f)) UNIV"
+proof- let ?n = "\<lambda>x. - min (f x) 0" and ?p = "\<lambda>x. max (f x) 0"
+  have *:"f = (\<lambda>x. ?p x - ?n x)" apply rule by auto
+  note f = lebesgue.integrableD[OF f]
+  show ?thesis unfolding lebesgue.integral_def apply(subst *)
+  proof(rule has_integral_sub) case goal1
+    have *:"\<forall>x. Real (f x) \<noteq> \<omega>" by auto
+    note lebesgue.borel_measurable_Real[OF f(1)]
+    from positive_integral_has_integral[OF this f(2) *]
+    show ?case unfolding real_Real_max .
+  next case goal2
+    have *:"\<forall>x. Real (- f x) \<noteq> \<omega>" by auto
+    note lebesgue.borel_measurable_uminus[OF f(1)]
+    note lebesgue.borel_measurable_Real[OF this]
+    from positive_integral_has_integral[OF this f(3) *]
+    show ?case unfolding real_Real_max minus_min_eq_max by auto
+  qed
+qed
+
+lemma continuous_on_imp_borel_measurable:
+  fixes f::"'a::ordered_euclidean_space \<Rightarrow> 'b::ordered_euclidean_space"
+  assumes "continuous_on UNIV f"
+  shows "f \<in> borel_measurable lebesgue"
+  apply(rule lebesgue.borel_measurableI)
+  using continuous_open_preimage[OF assms] unfolding vimage_def by auto
+
+lemma (in measure_space) integral_monotone_convergence_pos':
+  assumes i: "\<And>i. integrable (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 (f i)) ----> x"
+  shows "integrable u \<and> integral u = x"
+  using integral_monotone_convergence_pos[OF assms] by auto
+
+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 bij_euclidean_component:
+  "bij_betw (e2p::'a::ordered_euclidean_space \<Rightarrow> _) (UNIV :: 'a set)
+  ({..<DIM('a)} \<rightarrow>\<^isub>E (UNIV :: real set))"
+  unfolding bij_betw_def e2p_def_raw
+proof let ?e = "\<lambda>x.\<lambda>i\<in>{..<DIM('a::ordered_euclidean_space)}. (x::'a)$$i"
+  show "inj ?e" unfolding inj_on_def restrict_def apply(subst euclidean_eq) apply safe
+    apply(drule_tac x=i in fun_cong) by auto
+  { fix x::"nat \<Rightarrow> real" assume x:"\<forall>i. i \<notin> {..<DIM('a)} \<longrightarrow> x i = undefined"
+    hence "x = ?e (\<chi>\<chi> i. x i)" apply-apply(rule,case_tac "xa<DIM('a)") by auto
+    hence "x \<in> range ?e" by fastsimp
+  } thus "range ?e = ({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)}"
+    unfolding extensional_def using DIM_positive by auto
+qed
+
+lemma bij_p2e:
+  "bij_betw (p2e::_ \<Rightarrow> 'a::ordered_euclidean_space) ({..<DIM('a)} \<rightarrow>\<^isub>E (UNIV :: real set))
+  (UNIV :: 'a set)" (is "bij_betw ?p ?U _")
+  unfolding bij_betw_def
+proof show "inj_on ?p ?U" unfolding inj_on_def p2e_def
+    apply(subst euclidean_eq) apply(safe,rule) unfolding extensional_def
+    apply(case_tac "xa<DIM('a)") by auto
+  { fix x::'a have "x \<in> ?p ` extensional {..<DIM('a)}"
+      unfolding image_iff apply(rule_tac x="\<lambda>i. if i<DIM('a) then x$$i else undefined" in bexI)
+      apply(subst euclidean_eq,safe) unfolding p2e_def extensional_def by auto
+  } thus "?p ` ?U = UNIV" by auto
+qed
+
+lemma e2p_p2e[simp]: fixes z::"'a::ordered_euclidean_space"
+  assumes "x \<in> extensional {..<DIM('a)}"
+  shows "e2p (p2e x::'a) = x"
+proof fix i::nat
+  show "e2p (p2e x::'a) i = x i" unfolding e2p_def p2e_def restrict_def
+    using assms unfolding extensional_def by auto
+qed
+
+lemma p2e_e2p[simp]: fixes x::"'a::ordered_euclidean_space"
+  shows "p2e (e2p x) = x"
+  apply(subst euclidean_eq) unfolding e2p_def p2e_def restrict_def by auto
+
+interpretation borel_product: product_sigma_finite "\<lambda>x. borel::real algebra" "\<lambda>x. lmeasure"
+  by default
+
+lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)"
+  unfolding cube_def_raw subset_eq apply safe unfolding mem_interval by auto
+
+lemma borel_vimage_algebra_eq:
+  "sigma_algebra.vimage_algebra
+         (borel :: ('a::ordered_euclidean_space) algebra) ({..<DIM('a)} \<rightarrow>\<^isub>E UNIV) p2e =
+  sigma (product_algebra (\<lambda>x. \<lparr> space = UNIV::real set, sets = range (\<lambda>a. {..<a}) \<rparr>) {..<DIM('a)} )"
+proof- note bor = borel_eq_lessThan
+  def F \<equiv> "product_algebra (\<lambda>x. \<lparr> space = UNIV::real set, sets = range (\<lambda>a. {..<a}) \<rparr>) {..<DIM('a)}"
+  def E \<equiv> "\<lparr>space = (UNIV::'a set), sets = range lessThan\<rparr>"
+  have *:"(({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)}) = space F" unfolding F_def by auto
+  show ?thesis unfolding F_def[symmetric] * bor
+  proof(rule vimage_algebra_sigma,unfold E_def[symmetric])
+    show "sets E \<subseteq> Pow (space E)" "p2e \<in> space F \<rightarrow> space E" unfolding E_def by auto
+  next fix A assume "A \<in> sets F"
+    hence A:"A \<in> (Pi\<^isub>E {..<DIM('a)}) ` ({..<DIM('a)} \<rightarrow> range lessThan)"
+      unfolding F_def product_algebra_def algebra.simps .
+    then guess B unfolding image_iff .. note B=this
+    hence "\<forall>x<DIM('a). B x \<in> range lessThan" by auto
+    hence "\<forall>x. \<exists>xa. x<DIM('a) \<longrightarrow> B x = {..<xa}" unfolding image_iff by auto
+    from choice[OF this] guess b .. note b=this
+    hence b':"\<forall>i<DIM('a). Sup (B i) = b i" using Sup_lessThan by auto
+
+    show "A \<in> (\<lambda>X. p2e -` X \<inter> space F) ` sets E" unfolding image_iff B
+    proof(rule_tac x="{..< \<chi>\<chi> i. Sup (B i)}" in bexI)
+      show "Pi\<^isub>E {..<DIM('a)} B = p2e -` {..<(\<chi>\<chi> i. Sup (B i))::'a} \<inter> space F"
+        unfolding F_def E_def product_algebra_def algebra.simps
+      proof(rule,unfold subset_eq,rule_tac[!] ballI)
+        fix x assume "x \<in> Pi\<^isub>E {..<DIM('a)} B"
+        hence *:"\<forall>i<DIM('a). x i < b i" "\<forall>i\<ge>DIM('a). x i = undefined"
+          unfolding Pi_def extensional_def using b by auto
+        have "(p2e x::'a) < (\<chi>\<chi> i. Sup (B i))" unfolding less_prod_def eucl_less[of "p2e x"]
+          apply safe unfolding euclidean_lambda_beta b'[rule_format] p2e_def using * by auto
+        moreover have "x \<in> extensional {..<DIM('a)}"
+          using *(2) unfolding extensional_def by auto
+        ultimately show "x \<in> p2e -` {..<(\<chi>\<chi> i. Sup (B i)) ::'a} \<inter>
+          (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})" by auto
+      next fix x assume as:"x \<in> p2e -` {..<(\<chi>\<chi> i. Sup (B i))::'a} \<inter>
+          (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})"
+        hence "p2e x < ((\<chi>\<chi> i. Sup (B i))::'a)" by auto
+        hence "\<forall>i<DIM('a). x i \<in> B i" apply-apply(subst(asm) eucl_less)
+          unfolding p2e_def using b b' by auto
+        thus "x \<in> Pi\<^isub>E {..<DIM('a)} B" using as unfolding Pi_def extensional_def by auto
+      qed
+      show "{..<(\<chi>\<chi> i. Sup (B i))::'a} \<in> sets E" unfolding E_def algebra.simps by auto
+    qed
+  next fix A assume "A \<in> sets E"
+    then guess a unfolding E_def algebra.simps image_iff .. note a = this(2)
+    def B \<equiv> "\<lambda>i. {..<a $$ i}"
+    show "p2e -` A \<inter> space F \<in> sets F" unfolding F_def
+      unfolding product_algebra_def algebra.simps image_iff
+      apply(rule_tac x=B in bexI) apply rule unfolding subset_eq apply(rule_tac[1-2] ballI)
+    proof- show "B \<in> {..<DIM('a)} \<rightarrow> range lessThan" unfolding B_def by auto
+      fix x assume as:"x \<in> p2e -` A \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})"
+      hence "p2e x \<in> A" by auto
+      hence "\<forall>i<DIM('a). x i \<in> B i" unfolding B_def a lessThan_iff
+        apply-apply(subst (asm) eucl_less) unfolding p2e_def by auto
+      thus "x \<in> Pi\<^isub>E {..<DIM('a)} B" using as unfolding Pi_def extensional_def by auto
+    next fix x assume x:"x \<in> Pi\<^isub>E {..<DIM('a)} B"
+      moreover have "p2e x \<in> A" unfolding a lessThan_iff p2e_def apply(subst eucl_less)
+        using x unfolding Pi_def extensional_def B_def by auto
+      ultimately show "x \<in> p2e -` A \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})" by auto
+    qed
+  qed
+qed
+
+lemma Real_mult_nonneg: assumes "x \<ge> 0" "y \<ge> 0"
+  shows "Real (x * y) = Real x * Real y" using assms by auto
+
+lemma Real_setprod: assumes "\<forall>x\<in>A. f x \<ge> 0" shows "Real (setprod f A) = setprod (\<lambda>x. Real (f x)) A"
+proof(cases "finite A")
+  case True thus ?thesis using assms
+  proof(induct A) case (insert x A)
+    have "0 \<le> setprod f A" apply(rule setprod_nonneg) using insert by auto
+    thus ?case unfolding setprod_insert[OF insert(1-2)] apply-
+      apply(subst Real_mult_nonneg) prefer 3 apply(subst insert(3)[THEN sym])
+      using insert by auto
+  qed auto
+qed auto
+
+lemma e2p_Int:"e2p ` A \<inter> e2p ` B = e2p ` (A \<inter> B)" (is "?L = ?R")
+  apply(rule image_Int[THEN sym]) using bij_euclidean_component
+  unfolding bij_betw_def by auto
+
+lemma Int_stable_cuboids: fixes x::"'a::ordered_euclidean_space"
+  shows "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a, b::'a). e2p ` {a..b})\<rparr>"
+  unfolding Int_stable_def algebra.select_convs
+proof safe fix a b x y::'a
+  have *:"e2p ` {a..b} \<inter> e2p ` {x..y} =
+    (\<lambda>(a, b). e2p ` {a..b}) (\<chi>\<chi> i. max (a $$ i) (x $$ i), \<chi>\<chi> i. min (b $$ i) (y $$ i)::'a)"
+    unfolding e2p_Int inter_interval by auto
+  show "e2p ` {a..b} \<inter> e2p ` {x..y} \<in> range (\<lambda>(a, b). e2p ` {a..b::'a})" unfolding *
+    apply(rule range_eqI) ..
+qed
+
+lemma Int_stable_cuboids': fixes x::"'a::ordered_euclidean_space"
+  shows "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a, b::'a). {a..b})\<rparr>"
+  unfolding Int_stable_def algebra.select_convs
+  apply safe unfolding inter_interval by auto
+
+lemma product_borel_eq_vimage:
+  "sigma (product_algebra (\<lambda>x. borel) {..<DIM('a::ordered_euclidean_space)}) =
+  sigma_algebra.vimage_algebra borel (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})
+  (p2e:: _ \<Rightarrow> 'a::ordered_euclidean_space)"
+  unfolding borel_vimage_algebra_eq unfolding borel_eq_lessThan
+  apply(subst sigma_product_algebra_sigma_eq[where S="\<lambda>i. \<lambda>n. lessThan (real n)"])
+  unfolding lessThan_iff
+proof- fix i assume i:"i<DIM('a)"
+  show "(\<lambda>n. {..<real n}) \<up> space \<lparr>space = UNIV, sets = range lessThan\<rparr>"
+    by(auto intro!:real_arch_lt isotoneI)
+qed auto
+
+lemma inj_on_disjoint_family_on: assumes "disjoint_family_on A S" "inj f"
+  shows "disjoint_family_on (\<lambda>x. f ` A x) S"
+  unfolding disjoint_family_on_def
+proof(rule,rule,rule)
+  fix x1 x2 assume x:"x1 \<in> S" "x2 \<in> S" "x1 \<noteq> x2"
+  show "f ` A x1 \<inter> f ` A x2 = {}"
+  proof(rule ccontr) case goal1
+    then obtain z where z:"z \<in> f ` A x1 \<inter> f ` A x2" by auto
+    then obtain z1 z2 where z12:"z1 \<in> A x1" "z2 \<in> A x2" "f z1 = z" "f z2 = z" by auto
+    hence "z1 = z2" using assms(2) unfolding inj_on_def by blast
+    hence "x1 = x2" using z12(1-2) using assms[unfolded disjoint_family_on_def] using x by auto
+    thus False using x(3) by auto
+  qed
+qed
+
+declare restrict_extensional[intro]
+
+lemma e2p_extensional[intro]:"e2p (y::'a::ordered_euclidean_space) \<in> extensional {..<DIM('a)}"
+  unfolding e2p_def by auto
+
+lemma e2p_image_vimage: fixes A::"'a::ordered_euclidean_space set"
+  shows "e2p ` A = p2e -` A \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})"
+proof(rule set_eqI,rule)
+  fix x assume "x \<in> e2p ` A" then guess y unfolding image_iff .. note y=this
+  show "x \<in> p2e -` A \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})"
+    apply safe apply(rule vimageI[OF _ y(1)]) unfolding y p2e_e2p by auto
+next fix x assume "x \<in> p2e -` A \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)})"
+  thus "x \<in> e2p ` A" unfolding image_iff apply(rule_tac x="p2e x" in bexI) apply(subst e2p_p2e) by auto
+qed
+
+lemma lmeasure_measure_eq_borel_prod:
+  fixes A :: "('a::ordered_euclidean_space) set"
+  assumes "A \<in> sets borel"
+  shows "lmeasure A = borel_product.product_measure {..<DIM('a)} (e2p ` A :: (nat \<Rightarrow> real) set)"
+proof (rule measure_unique_Int_stable[where X=A and A=cube])
+  interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
+  show "Int_stable \<lparr> space = UNIV :: 'a set, sets = range (\<lambda>(a,b). {a..b}) \<rparr>"
+    (is "Int_stable ?E" ) using Int_stable_cuboids' .
+  show "borel = sigma ?E" using borel_eq_atLeastAtMost .
+  show "\<And>i. lmeasure (cube i) \<noteq> \<omega>" unfolding lmeasure_cube by auto
+  show "\<And>X. X \<in> sets ?E \<Longrightarrow>
+    lmeasure X = borel_product.product_measure {..<DIM('a)} (e2p ` X :: (nat \<Rightarrow> real) set)"
+  proof- case goal1 then obtain a b where X:"X = {a..b}" by auto
+    { presume *:"X \<noteq> {} \<Longrightarrow> ?case"
+      show ?case apply(cases,rule *,assumption) by auto }
+    def XX \<equiv> "\<lambda>i. {a $$ i .. b $$ i}" assume  "X \<noteq> {}"  note X' = this[unfolded X interval_ne_empty]
+    have *:"Pi\<^isub>E {..<DIM('a)} XX = e2p ` X" apply(rule set_eqI)
+    proof fix x assume "x \<in> Pi\<^isub>E {..<DIM('a)} XX"
+      thus "x \<in> e2p ` X" unfolding image_iff apply(rule_tac x="\<chi>\<chi> i. x i" in bexI)
+        unfolding Pi_def extensional_def e2p_def restrict_def X mem_interval XX_def by rule auto
+    next fix x assume "x \<in> e2p ` X" then guess y unfolding image_iff .. note y = this
+      show "x \<in> Pi\<^isub>E {..<DIM('a)} XX" unfolding y using y(1)
+        unfolding Pi_def extensional_def e2p_def restrict_def X mem_interval XX_def by auto
+    qed
+    have "lmeasure X = (\<Prod>x<DIM('a). Real (b $$ x - a $$ x))"  using X' apply- unfolding X
+      unfolding lmeasure_atLeastAtMost content_closed_interval apply(subst Real_setprod) by auto
+    also have "... = (\<Prod>i<DIM('a). lmeasure (XX i))" apply(rule setprod_cong2)
+      unfolding XX_def lmeasure_atLeastAtMost apply(subst content_real) using X' by auto
+    also have "... = borel_product.product_measure {..<DIM('a)} (e2p ` X)" unfolding *[THEN sym]
+      apply(rule fprod.measure_times[THEN sym]) unfolding XX_def by auto
+    finally show ?case .
+  qed
+
+  show "range cube \<subseteq> sets \<lparr>space = UNIV, sets = range (\<lambda>(a, b). {a..b})\<rparr>"
+    unfolding cube_def_raw by auto
+  have "\<And>x. \<exists>xa. x \<in> cube xa" apply(rule_tac x=x in mem_big_cube) by fastsimp
+  thus "cube \<up> space \<lparr>space = UNIV, sets = range (\<lambda>(a, b). {a..b})\<rparr>"
+    apply-apply(rule isotoneI) apply(rule cube_subset_Suc) by auto
+  show "A \<in> sets borel " by fact
+  show "measure_space borel lmeasure" by default
+  show "measure_space borel
+     (\<lambda>a::'a set. finite_product_sigma_finite.measure (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` a))"
+    apply default unfolding countably_additive_def
+  proof safe fix A::"nat \<Rightarrow> 'a set" assume A:"range A \<subseteq> sets borel" "disjoint_family A"
+      "(\<Union>i. A i) \<in> sets borel"
+    note fprod.ca[unfolded countably_additive_def,rule_format]
+    note ca = this[of "\<lambda> n. e2p ` (A n)"]
+    show "(\<Sum>\<^isub>\<infinity>n. finite_product_sigma_finite.measure
+        (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` A n)) =
+           finite_product_sigma_finite.measure (\<lambda>x. borel)
+            (\<lambda>x. lmeasure) {..<DIM('a)} (e2p ` (\<Union>i. A i))" unfolding image_UN
+    proof(rule ca) show "range (\<lambda>n. e2p ` A n) \<subseteq> sets
+       (sigma (product_algebra (\<lambda>x. borel) {..<DIM('a)}))"
+        unfolding product_borel_eq_vimage
+      proof case goal1
+        then guess y unfolding image_iff .. note y=this(2)
+        show ?case unfolding borel.in_vimage_algebra y apply-
+          apply(rule_tac x="A y" in bexI,rule e2p_image_vimage)
+          using A(1) by auto
+      qed
+
+      show "disjoint_family (\<lambda>n. e2p ` A n)" apply(rule inj_on_disjoint_family_on)
+        using bij_euclidean_component using A(2) unfolding bij_betw_def by auto
+      show "(\<Union>n. e2p ` A n) \<in> sets (sigma (product_algebra (\<lambda>x. borel) {..<DIM('a)}))"
+        unfolding product_borel_eq_vimage borel.in_vimage_algebra
+      proof(rule bexI[OF _ A(3)],rule set_eqI,rule)
+        fix x assume x:"x \<in> (\<Union>n. e2p ` A n)" hence "p2e x \<in> (\<Union>i. A i)" by auto
+        moreover have "x \<in> extensional {..<DIM('a)}"
+          using x unfolding extensional_def e2p_def_raw by auto
+        ultimately show "x \<in> p2e -` (\<Union>i. A i) \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter>
+          extensional {..<DIM('a)})" by auto
+      next fix x assume x:"x \<in> p2e -` (\<Union>i. A i) \<inter> (({..<DIM('a)} \<rightarrow> UNIV) \<inter>
+          extensional {..<DIM('a)})"
+        hence "p2e x \<in> (\<Union>i. A i)" by auto
+        hence "\<exists>n. x \<in> e2p ` A n" apply safe apply(rule_tac x=i in exI)
+          unfolding image_iff apply(rule_tac x="p2e x" in bexI)
+          apply(subst e2p_p2e) using x by auto
+        thus "x \<in> (\<Union>n. e2p ` A n)" by auto
+      qed
+    qed
+  qed auto
+qed
+
+lemma e2p_p2e'[simp]: fixes x::"'a::ordered_euclidean_space"
+  assumes "A \<subseteq> extensional {..<DIM('a)}"
+  shows "e2p ` (p2e ` A ::'a set) = A"
+  apply(rule set_eqI) unfolding image_iff Bex_def apply safe defer
+  apply(rule_tac x="p2e x" in exI,safe) using assms by auto
+
+lemma range_p2e:"range (p2e::_\<Rightarrow>'a::ordered_euclidean_space) = UNIV"
+  apply safe defer unfolding image_iff apply(rule_tac x="\<lambda>i. x $$ i" in bexI)
+  unfolding p2e_def by auto
+
+lemma p2e_inv_extensional:"(A::'a::ordered_euclidean_space set)
+  = p2e ` (p2e -` A \<inter> extensional {..<DIM('a)})"
+  unfolding p2e_def_raw apply safe unfolding image_iff
+proof- fix x assume "x\<in>A"
+  let ?y = "\<lambda>i. if i<DIM('a) then x$$i else undefined"
+  have *:"Chi ?y = x" apply(subst euclidean_eq) by auto
+  show "\<exists>xa\<in>Chi -` A \<inter> extensional {..<DIM('a)}. x = Chi xa" apply(rule_tac x="?y" in bexI)
+    apply(subst euclidean_eq) unfolding extensional_def using `x\<in>A` by(auto simp: *)
+qed
+
+lemma borel_fubini_positiv_integral:
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> pinfreal"
+  assumes f: "f \<in> borel_measurable borel"
+  shows "borel.positive_integral f =
+          borel_product.product_positive_integral {..<DIM('a)} (f \<circ> p2e)"
+proof- def U \<equiv> "(({..<DIM('a)} \<rightarrow> UNIV) \<inter> extensional {..<DIM('a)}):: (nat \<Rightarrow> real) set"
+  interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
+  have "\<And>x. \<exists>i::nat. x < real i" by (metis real_arch_lt)
+  hence "(\<lambda>n::nat. {..<real n}) \<up> UNIV" apply-apply(rule isotoneI) by auto
+  hence *:"sigma_algebra.vimage_algebra borel U (p2e:: _ \<Rightarrow> 'a)
+    = sigma (product_algebra (\<lambda>x. borel) {..<DIM('a)})"
+    unfolding U_def apply-apply(subst borel_vimage_algebra_eq)
+    apply(subst sigma_product_algebra_sigma_eq[where S="\<lambda>x. \<lambda>n. {..<(\<chi>\<chi> i. real n)}", THEN sym])
+    unfolding borel_eq_lessThan[THEN sym] by auto
+  show ?thesis unfolding borel.positive_integral_vimage[unfolded space_borel,OF bij_p2e]
+    apply(subst fprod.positive_integral_cong_measure[THEN sym, of "\<lambda>A. lmeasure (p2e ` A)"])
+    unfolding U_def[symmetric] *[THEN sym] o_def
+  proof- fix A assume A:"A \<in> sets (sigma_algebra.vimage_algebra borel U (p2e ::_ \<Rightarrow> 'a))"
+    hence *:"A \<subseteq> extensional {..<DIM('a)}" unfolding U_def by auto
+    from A guess B unfolding borel.in_vimage_algebra U_def .. note B=this
+    have "(p2e ` A::'a set) \<in> sets borel" unfolding B apply(subst Int_left_commute)
+      apply(subst Int_absorb1) unfolding p2e_inv_extensional[of B,THEN sym] using B(1) by auto
+    from lmeasure_measure_eq_borel_prod[OF this] show "lmeasure (p2e ` A::'a set) =
+      finite_product_sigma_finite.measure (\<lambda>x. borel) (\<lambda>x. lmeasure) {..<DIM('a)} A"
+      unfolding e2p_p2e'[OF *] .
+  qed auto
+qed
+
+lemma borel_fubini:
+  fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable borel"
+  shows "borel.integral f = borel_product.product_integral {..<DIM('a)} (f \<circ> p2e)"
+proof- interpret fprod: finite_product_sigma_finite "\<lambda>x. borel" "\<lambda>x. lmeasure" "{..<DIM('a)}" by default auto
+  have 1:"(\<lambda>x. Real (f x)) \<in> borel_measurable borel" using f by auto
+  have 2:"(\<lambda>x. Real (- f x)) \<in> borel_measurable borel" using f by auto
+  show ?thesis unfolding fprod.integral_def borel.integral_def
+    unfolding borel_fubini_positiv_integral[OF 1] borel_fubini_positiv_integral[OF 2]
+    unfolding o_def ..
 qed
 
 end
--- a/src/HOL/Probability/Measure.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Measure.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -1,10 +1,47 @@
-header {*Measures*}
+(* Author: Lawrence C Paulson; Armin Heller, Johannes Hoelzl, TU Muenchen *)
 
 theory Measure
   imports Caratheodory
 begin
 
-text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*}
+lemma inj_on_image_eq_iff:
+  assumes "inj_on f S"
+  assumes "A \<subseteq> S" "B \<subseteq> S"
+  shows "(f ` A = f ` B) \<longleftrightarrow> (A = B)"
+proof -
+  have "inj_on f (A \<union> B)"
+    using assms by (auto intro: subset_inj_on)
+  from inj_on_Un_image_eq_iff[OF this]
+  show ?thesis .
+qed
+
+lemma image_vimage_inter_eq:
+  assumes "f ` S = T" "X \<subseteq> T"
+  shows "f ` (f -` X \<inter> S) = X"
+proof (intro antisym)
+  have "f ` (f -` X \<inter> S) \<subseteq> f ` (f -` X)" by auto
+  also have "\<dots> = X \<inter> range f" by simp
+  also have "\<dots> = X" using assms by auto
+  finally show "f ` (f -` X \<inter> S) \<subseteq> X" by auto
+next
+  show "X \<subseteq> f ` (f -` X \<inter> S)"
+  proof
+    fix x assume "x \<in> X"
+    then have "x \<in> T" using `X \<subseteq> T` by auto
+    then obtain y where "x = f y" "y \<in> S"
+      using assms by auto
+    then have "{y} \<subseteq> f -` X \<inter> S" using `x \<in> X` by auto
+    moreover have "x \<in> f ` {y}" using `x = f y` by auto
+    ultimately show "x \<in> f ` (f -` X \<inter> S)" by auto
+  qed
+qed
+
+text {*
+  This formalisation of measure theory is based on the work of Hurd/Coble wand
+  was later translated by Lawrence Paulson to Isabelle/HOL. Later it was
+  modified to use the positive infinite reals and to prove the uniqueness of
+  cut stable measures.
+*}
 
 section {* Equations for the measure function @{text \<mu>} *}
 
@@ -157,11 +194,19 @@
 qed
 
 lemma (in measure_space) measure_up:
-  assumes "P \<in> sets M" "\<And>i. B i \<in> sets M" "B \<up> P"
+  assumes "\<And>i. B i \<in> sets M" "B \<up> P"
   shows "(\<lambda>i. \<mu> (B i)) \<up> \<mu> P"
   using assms unfolding isoton_def
   by (auto intro!: measure_mono continuity_from_below)
 
+lemma (in measure_space) continuity_from_below':
+  assumes A: "range A \<subseteq> sets M" "\<And>i. A i \<subseteq> A (Suc i)"
+  shows "(\<lambda>i. (\<mu> (A i))) ----> (\<mu> (\<Union>i. A i))"
+proof- let ?A = "\<Union>i. A i"
+  have " (\<lambda>i. \<mu> (A i)) \<up> \<mu> ?A" apply(rule measure_up)
+    using assms unfolding complete_lattice_class.isoton_def by auto
+  thus ?thesis by(rule isotone_Lim(1))
+qed
 
 lemma (in measure_space) continuity_from_above:
   assumes A: "range A \<subseteq> sets M"
@@ -196,7 +241,7 @@
 qed
 
 lemma (in measure_space) measure_down:
-  assumes "P \<in> sets M" "\<And>i. B i \<in> sets M" "B \<down> P"
+  assumes "\<And>i. B i \<in> sets M" "B \<down> P"
   and finite: "\<And>i. \<mu> (B i) \<noteq> \<omega>"
   shows "(\<lambda>i. \<mu> (B i)) \<down> \<mu> P"
   using assms unfolding antiton_def
@@ -414,6 +459,11 @@
   finally show ?thesis .
 qed
 
+lemma (in measure_space) measure_eq_0:
+  assumes "N \<in> sets M" and "\<mu> N = 0" and "K \<subseteq> N" and "K \<in> sets M"
+  shows "\<mu> K = 0"
+using measure_mono[OF assms(3,4,1)] assms(2) by auto
+
 lemma (in measure_space) measure_finitely_subadditive:
   assumes "finite I" "A ` I \<subseteq> sets M"
   shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
@@ -427,7 +477,7 @@
   finally show ?case .
 qed simp
 
-lemma (in measure_space) measurable_countably_subadditive:
+lemma (in measure_space) measure_countably_subadditive:
   assumes "range f \<subseteq> sets M"
   shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
 proof -
@@ -445,6 +495,11 @@
   finally show ?thesis .
 qed
 
+lemma (in measure_space) measure_UN_eq_0:
+  assumes "\<And> i :: nat. \<mu> (N i) = 0" and "range N \<subseteq> sets M"
+  shows "\<mu> (\<Union> i. N i) = 0"
+using measure_countably_subadditive[OF assms(2)] assms(1) by auto
+
 lemma (in measure_space) measure_inter_full_set:
   assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
   assumes T: "\<mu> T = \<mu> (space M)"
@@ -470,6 +525,340 @@
   qed
 qed
 
+lemma measure_unique_Int_stable:
+  fixes M E :: "'a algebra" and A :: "nat \<Rightarrow> 'a set"
+  assumes "Int_stable E" "M = sigma E"
+  and A: "range  A \<subseteq> sets E" "A \<up> space E"
+  and ms: "measure_space M \<mu>" "measure_space M \<nu>"
+  and eq: "\<And>X. X \<in> sets E \<Longrightarrow> \<mu> X = \<nu> X"
+  and finite: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+  assumes "X \<in> sets M"
+  shows "\<mu> X = \<nu> X"
+proof -
+  let "?D F" = "{D. D \<in> sets M \<and> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)}"
+  interpret M: measure_space M \<mu> by fact
+  interpret M': measure_space M \<nu> by fact
+  have "space E = space M"
+    using `M = sigma E` by simp
+  have sets_E: "sets E \<subseteq> Pow (space E)"
+  proof
+    fix X assume "X \<in> sets E"
+    then have "X \<in> sets M" unfolding `M = sigma E`
+      unfolding sigma_def by (auto intro!: sigma_sets.Basic)
+    with M.sets_into_space show "X \<in> Pow (space E)"
+      unfolding `space E = space M` by auto
+  qed
+  have A': "range A \<subseteq> sets M" using `M = sigma E` A
+    by (auto simp: sets_sigma intro!: sigma_sets.Basic)
+  { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<omega>"
+    then have [intro]: "F \<in> sets M" unfolding `M = sigma E` sets_sigma
+      by (intro sigma_sets.Basic)
+    have "\<nu> F \<noteq> \<omega>" using `\<mu> F \<noteq> \<omega>` `F \<in> sets E` eq by simp
+    interpret D: dynkin_system "\<lparr>space=space E, sets=?D F\<rparr>"
+    proof (rule dynkin_systemI, simp_all)
+      fix A assume "A \<in> sets M \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
+      then show "A \<subseteq> space E"
+        unfolding `space E = space M` using M.sets_into_space by auto
+    next
+      have "F \<inter> space E = F" using `F \<in> sets E` sets_E by auto
+      then show "space E \<in> sets M \<and> \<mu> (F \<inter> space E) = \<nu> (F \<inter> space E)"
+        unfolding `space E = space M` using `F \<in> sets E` eq by auto
+    next
+      fix A assume *: "A \<in> sets M \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
+      then have **: "F \<inter> (space M - A) = F - (F \<inter> A)"
+        and [intro]: "F \<inter> A \<in> sets M"
+        using `F \<in> sets E` sets_E `space E = space M` by auto
+      have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: M'.measure_mono)
+      then have "\<nu> (F \<inter> A) \<noteq> \<omega>" using `\<nu> F \<noteq> \<omega>` by auto
+      have "\<mu> (F \<inter> A) \<le> \<mu> F" by (auto intro!: M.measure_mono)
+      then have "\<mu> (F \<inter> A) \<noteq> \<omega>" using `\<mu> F \<noteq> \<omega>` by auto
+      then have "\<mu> (F \<inter> (space M - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding **
+        using `F \<inter> A \<in> sets M` by (auto intro!: M.measure_Diff)
+      also have "\<dots> = \<nu> F - \<nu> (F \<inter> A)" using eq `F \<in> sets E` * by simp
+      also have "\<dots> = \<nu> (F \<inter> (space M - A))" unfolding **
+        using `F \<inter> A \<in> sets M` `\<nu> (F \<inter> A) \<noteq> \<omega>` by (auto intro!: M'.measure_Diff[symmetric])
+      finally show "space E - A \<in> sets M \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))"
+        using `space E = space M` * by auto
+    next
+      fix A :: "nat \<Rightarrow> 'a set"
+      assume "disjoint_family A" "range A \<subseteq> {X \<in> sets M. \<mu> (F \<inter> X) = \<nu> (F \<inter> X)}"
+      then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sets M" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
+        "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. \<mu> (F \<inter> A i) = \<nu> (F \<inter> A i)" "range A \<subseteq> sets M"
+        by ((fastsimp simp: disjoint_family_on_def)+)
+      then show "(\<Union>x. A x) \<in> sets M \<and> \<mu> (F \<inter> (\<Union>x. A x)) = \<nu> (F \<inter> (\<Union>x. A x))"
+        by (auto simp: M.measure_countably_additive[symmetric]
+                       M'.measure_countably_additive[symmetric]
+            simp del: UN_simps)
+    qed
+    have *: "sigma E = \<lparr>space = space E, sets = ?D F\<rparr>"
+      using `M = sigma E` `F \<in> sets E` `Int_stable E`
+      by (intro D.dynkin_lemma)
+         (auto simp add: sets_sigma Int_stable_def eq intro: sigma_sets.Basic)
+    have "\<And>D. D \<in> sets M \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)"
+      unfolding `M = sigma E` by (auto simp: *) }
+  note * = this
+  { fix i have "\<mu> (A i \<inter> X) = \<nu> (A i \<inter> X)"
+      using *[of "A i" X] `X \<in> sets M` A finite by auto }
+  moreover
+  have "(\<lambda>i. A i \<inter> X) \<up> X"
+    using `X \<in> sets M` M.sets_into_space A `space E = space M`
+    by (auto simp: isoton_def)
+  then have "(\<lambda>i. \<mu> (A i \<inter> X)) \<up> \<mu> X" "(\<lambda>i. \<nu> (A i \<inter> X)) \<up> \<nu> X"
+    using `X \<in> sets M` A' by (auto intro!: M.measure_up M'.measure_up M.Int)
+  ultimately show ?thesis by (simp add: isoton_def)
+qed
+
+section "Isomorphisms between measure spaces"
+
+lemma (in measure_space) measure_space_isomorphic:
+  fixes f :: "'c \<Rightarrow> 'a"
+  assumes "bij_betw f S (space M)"
+  shows "measure_space (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A))"
+    (is "measure_space ?T ?\<mu>")
+proof -
+  have "f \<in> S \<rightarrow> space M" using assms unfolding bij_betw_def by auto
+  then interpret T: sigma_algebra ?T by (rule sigma_algebra_vimage)
+  show ?thesis
+  proof
+    show "\<mu> (f`{}) = 0" by simp
+    show "countably_additive ?T (\<lambda>A. \<mu> (f ` A))"
+    proof (unfold countably_additive_def, intro allI impI)
+      fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets ?T" "disjoint_family A"
+      then have "\<forall>i. \<exists>F'. F' \<in> sets M \<and> A i = f -` F' \<inter> S"
+        by (auto simp: image_iff image_subset_iff Bex_def vimage_algebra_def)
+      from choice[OF this] obtain F where F: "\<And>i. F i \<in> sets M" and A: "\<And>i. A i = f -` F i \<inter> S" by auto
+      then have [simp]: "\<And>i. f ` (A i) = F i"
+        using sets_into_space assms
+        by (force intro!: image_vimage_inter_eq[where T="space M"] simp: bij_betw_def)
+      have "disjoint_family F"
+      proof (intro disjoint_family_on_bisimulation[OF `disjoint_family A`])
+        fix n m assume "A n \<inter> A m = {}"
+        then have "f -` (F n \<inter> F m) \<inter> S = {}" unfolding A by auto
+        moreover
+        have "F n \<in> sets M" "F m \<in> sets M" using F by auto
+        then have "f`S = space M" "F n \<inter> F m \<subseteq> space M"
+          using sets_into_space assms by (auto simp: bij_betw_def)
+        note image_vimage_inter_eq[OF this, symmetric]
+        ultimately show "F n \<inter> F m = {}" by simp
+      qed
+      with F show "(\<Sum>\<^isub>\<infinity> i. \<mu> (f ` A i)) = \<mu> (f ` (\<Union>i. A i))"
+        by (auto simp add: image_UN intro!: measure_countably_additive)
+    qed
+  qed
+qed
+
+section "@{text \<mu>}-null sets"
+
+abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}"
+
+definition (in measure_space)
+  almost_everywhere :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "AE " 10) where
+  "almost_everywhere P \<longleftrightarrow> (\<exists>N\<in>null_sets. { x \<in> space M. \<not> P x } \<subseteq> N)"
+
+lemma (in measure_space) AE_I':
+  "N \<in> null_sets \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x. P x)"
+  unfolding almost_everywhere_def by auto
+
+lemma (in measure_space) AE_iff_null_set:
+  assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
+  shows "(AE x. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets"
+proof
+  assume "AE x. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "\<mu> N = 0"
+    unfolding almost_everywhere_def by auto
+  moreover have "\<mu> ?P \<le> \<mu> N"
+    using assms N(1,2) by (auto intro: measure_mono)
+  ultimately show "?P \<in> null_sets" using assms by auto
+next
+  assume "?P \<in> null_sets" with assms show "AE x. P x" by (auto intro: AE_I')
+qed
+
+lemma (in measure_space) null_sets_Un[intro]:
+  assumes "N \<in> null_sets" "N' \<in> null_sets"
+  shows "N \<union> N' \<in> null_sets"
+proof (intro conjI CollectI)
+  show "N \<union> N' \<in> sets M" using assms by auto
+  have "\<mu> (N \<union> N') \<le> \<mu> N + \<mu> N'"
+    using assms by (intro measure_subadditive) auto
+  then show "\<mu> (N \<union> N') = 0"
+    using assms by auto
+qed
+
+lemma UN_from_nat: "(\<Union>i. N i) = (\<Union>i. N (Countable.from_nat i))"
+proof -
+  have "(\<Union>i. N i) = (\<Union>i. (N \<circ> Countable.from_nat) i)"
+    unfolding SUPR_def image_compose
+    unfolding surj_from_nat ..
+  then show ?thesis by simp
+qed
+
+lemma (in measure_space) null_sets_UN[intro]:
+  assumes "\<And>i::'i::countable. N i \<in> null_sets"
+  shows "(\<Union>i. N i) \<in> null_sets"
+proof (intro conjI CollectI)
+  show "(\<Union>i. N i) \<in> sets M" using assms by auto
+  have "\<mu> (\<Union>i. N i) \<le> (\<Sum>\<^isub>\<infinity> n. \<mu> (N (Countable.from_nat n)))"
+    unfolding UN_from_nat[of N]
+    using assms by (intro measure_countably_subadditive) auto
+  then show "\<mu> (\<Union>i. N i) = 0"
+    using assms by auto
+qed
+
+lemma (in measure_space) null_set_Int1:
+  assumes "B \<in> null_sets" "A \<in> sets M" shows "A \<inter> B \<in> null_sets"
+using assms proof (intro CollectI conjI)
+  show "\<mu> (A \<inter> B) = 0" using assms by (intro measure_eq_0[of B "A \<inter> B"]) auto
+qed auto
+
+lemma (in measure_space) null_set_Int2:
+  assumes "B \<in> null_sets" "A \<in> sets M" shows "B \<inter> A \<in> null_sets"
+  using assms by (subst Int_commute) (rule null_set_Int1)
+
+lemma (in measure_space) measure_Diff_null_set:
+  assumes "B \<in> null_sets" "A \<in> sets M"
+  shows "\<mu> (A - B) = \<mu> A"
+proof -
+  have *: "A - B = (A - (A \<inter> B))" by auto
+  have "A \<inter> B \<in> null_sets" using assms by (rule null_set_Int1)
+  then show ?thesis
+    unfolding * using assms
+    by (subst measure_Diff) auto
+qed
+
+lemma (in measure_space) null_set_Diff:
+  assumes "B \<in> null_sets" "A \<in> sets M" shows "B - A \<in> null_sets"
+using assms proof (intro CollectI conjI)
+  show "\<mu> (B - A) = 0" using assms by (intro measure_eq_0[of B "B - A"]) auto
+qed auto
+
+lemma (in measure_space) measure_Un_null_set:
+  assumes "A \<in> sets M" "B \<in> null_sets"
+  shows "\<mu> (A \<union> B) = \<mu> A"
+proof -
+  have *: "A \<union> B = A \<union> (B - A)" by auto
+  have "B - A \<in> null_sets" using assms(2,1) by (rule null_set_Diff)
+  then show ?thesis
+    unfolding * using assms
+    by (subst measure_additive[symmetric]) auto
+qed
+
+lemma (in measure_space) AE_True[intro, simp]: "AE x. True"
+  unfolding almost_everywhere_def by auto
+
+lemma (in measure_space) AE_E[consumes 1]:
+  assumes "AE x. P x"
+  obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "\<mu> N = 0" "N \<in> sets M"
+  using assms unfolding almost_everywhere_def by auto
+
+lemma (in measure_space) AE_I:
+  assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "\<mu> N = 0" "N \<in> sets M"
+  shows "AE x. P x"
+  using assms unfolding almost_everywhere_def by auto
+
+lemma (in measure_space) AE_mp:
+  assumes AE_P: "AE x. P x" and AE_imp: "AE x. P x \<longrightarrow> Q x"
+  shows "AE x. Q x"
+proof -
+  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
+    and A: "A \<in> sets M" "\<mu> A = 0"
+    by (auto elim!: AE_E)
+
+  from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
+    and B: "B \<in> sets M" "\<mu> B = 0"
+    by (auto elim!: AE_E)
+
+  show ?thesis
+  proof (intro AE_I)
+    show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B
+      using measure_subadditive[of A B] by auto
+    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
+      using P imp by auto
+  qed
+qed
+
+lemma (in measure_space) AE_iffI:
+  assumes P: "AE x. P x" and Q: "AE x. P x \<longleftrightarrow> Q x" shows "AE x. Q x"
+  using AE_mp[OF Q, of "\<lambda>x. P x \<longrightarrow> Q x"] AE_mp[OF P, of Q] by auto
+
+lemma (in measure_space) AE_disjI1:
+  assumes P: "AE x. P x" shows "AE x. P x \<or> Q x"
+  by (rule AE_mp[OF P]) auto
+
+lemma (in measure_space) AE_disjI2:
+  assumes P: "AE x. Q x" shows "AE x. P x \<or> Q x"
+  by (rule AE_mp[OF P]) auto
+
+lemma (in measure_space) AE_conjI:
+  assumes AE_P: "AE x. P x" and AE_Q: "AE x. Q x"
+  shows "AE x. P x \<and> Q x"
+proof -
+  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
+    and A: "A \<in> sets M" "\<mu> A = 0"
+    by (auto elim!: AE_E)
+
+  from AE_Q obtain B where Q: "{x\<in>space M. \<not> Q x} \<subseteq> B"
+    and B: "B \<in> sets M" "\<mu> B = 0"
+    by (auto elim!: AE_E)
+
+  show ?thesis
+  proof (intro AE_I)
+    show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B
+      using measure_subadditive[of A B] by auto
+    show "{x\<in>space M. \<not> (P x \<and> Q x)} \<subseteq> A \<union> B"
+      using P Q by auto
+  qed
+qed
+
+lemma (in measure_space) AE_E2:
+  assumes "AE x. P x" "{x\<in>space M. P x} \<in> sets M"
+  shows "\<mu> {x\<in>space M. \<not> P x} = 0" (is "\<mu> ?P = 0")
+proof -
+  obtain A where A: "?P \<subseteq> A" "A \<in> sets M" "\<mu> A = 0"
+    using assms by (auto elim!: AE_E)
+  have "?P = space M - {x\<in>space M. P x}" by auto
+  then have "?P \<in> sets M" using assms by auto
+  with assms `?P \<subseteq> A` `A \<in> sets M` have "\<mu> ?P \<le> \<mu> A"
+    by (auto intro!: measure_mono)
+  then show "\<mu> ?P = 0" using A by simp
+qed
+
+lemma (in measure_space) AE_space[simp, intro]: "AE x. x \<in> space M"
+  by (rule AE_I[where N="{}"]) auto
+
+lemma (in measure_space) AE_cong:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> P x" shows "AE x. P x"
+proof -
+  have [simp]: "\<And>x. (x \<in> space M \<longrightarrow> P x) \<longleftrightarrow> True" using assms by auto
+  show ?thesis
+    by (rule AE_mp[OF AE_space]) auto
+qed
+
+lemma (in measure_space) AE_conj_iff[simp]:
+  shows "(AE x. P x \<and> Q x) \<longleftrightarrow> (AE x. P x) \<and> (AE x. Q x)"
+proof (intro conjI iffI AE_conjI)
+  assume *: "AE x. P x \<and> Q x"
+  from * show "AE x. P x" by (rule AE_mp) auto
+  from * show "AE x. Q x" by (rule AE_mp) auto
+qed auto
+
+lemma (in measure_space) all_AE_countable:
+  "(\<forall>i::'i::countable. AE x. P i x) \<longleftrightarrow> (AE x. \<forall>i. P i x)"
+proof
+  assume "\<forall>i. AE x. P i x"
+  from this[unfolded almost_everywhere_def Bex_def, THEN choice]
+  obtain N where N: "\<And>i. N i \<in> null_sets" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
+  have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
+  also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
+  finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
+  moreover from N have "(\<Union>i. N i) \<in> null_sets"
+    by (intro null_sets_UN) auto
+  ultimately show "AE x. \<forall>i. P i x"
+    unfolding almost_everywhere_def by auto
+next
+  assume *: "AE x. \<forall>i. P i x"
+  show "\<forall>i. AE x. P i x"
+    by (rule allI, rule AE_mp[OF *]) simp
+qed
+
 lemma (in measure_space) restricted_measure_space:
   assumes "S \<in> sets M"
   shows "measure_space (restricted_space S) \<mu>"
@@ -490,6 +879,7 @@
 qed
 
 lemma (in measure_space) measure_space_vimage:
+  fixes M' :: "'b algebra"
   assumes "f \<in> measurable M M'"
   and "sigma_algebra M'"
   shows "measure_space M' (\<lambda>A. \<mu> (f -` A \<inter> space M))" (is "measure_space M' ?T")
@@ -502,7 +892,7 @@
 
     show "countably_additive M' ?T"
     proof (unfold countably_additive_def, safe)
-      fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A"
+      fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets M'" "disjoint_family A"
       hence *: "\<And>i. f -` (A i) \<inter> space M \<in> sets M"
         using `f \<in> measurable M M'` by (auto simp: measurable_def)
       moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
@@ -564,6 +954,20 @@
   qed
 qed
 
+lemma (in sigma_finite_measure) sigma_finite_measure_cong:
+  assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A"
+  shows "sigma_finite_measure M \<mu>'"
+proof -
+  interpret \<mu>': measure_space M \<mu>'
+    using cong by (rule measure_space_cong)
+  from sigma_finite guess A .. note A = this
+  then have "\<And>i. A i \<in> sets M" by auto
+  with A have fin: "(\<forall>i. \<mu>' (A i) \<noteq> \<omega>)" using cong by auto
+  show ?thesis
+    apply default
+    using A fin by auto
+qed
+
 lemma (in sigma_finite_measure) disjoint_sigma_finite:
   "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
     (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
@@ -583,6 +987,50 @@
   show ?thesis by (auto intro!: exI[of _ "disjointed A"])
 qed
 
+lemma (in sigma_finite_measure) sigma_finite_up:
+  "\<exists>F. range F \<subseteq> sets M \<and> F \<up> space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<omega>)"
+proof -
+  obtain F :: "nat \<Rightarrow> 'a set" where
+    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. \<mu> (F i) \<noteq> \<omega>"
+    using sigma_finite by auto
+  then show ?thesis unfolding isoton_def
+  proof (intro exI[of _ "\<lambda>n. \<Union>i\<le>n. F i"] conjI allI)
+    from F have "\<And>x. x \<in> space M \<Longrightarrow> \<exists>i. x \<in> F i" by auto
+    then show "(\<Union>n. \<Union> i\<le>n. F i) = space M"
+      using F by fastsimp
+  next
+    fix n
+    have "\<mu> (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. \<mu> (F i))" using F
+      by (auto intro!: measure_finitely_subadditive)
+    also have "\<dots> < \<omega>"
+      using F by (auto simp: setsum_\<omega>)
+    finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<omega>" by simp
+  qed force+
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_measure_isomorphic:
+  assumes f: "bij_betw f S (space M)"
+  shows "sigma_finite_measure (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A))"
+proof -
+  interpret M: measure_space "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)"
+    using f by (rule measure_space_isomorphic)
+  show ?thesis
+  proof default
+    from sigma_finite guess A::"nat \<Rightarrow> 'a set" .. note A = this
+    show "\<exists>A::nat\<Rightarrow>'b set. range A \<subseteq> sets (vimage_algebra S f) \<and> (\<Union>i. A i) = space (vimage_algebra S f) \<and> (\<forall>i. \<mu> (f ` A i) \<noteq> \<omega>)"
+    proof (intro exI conjI)
+      show "(\<Union>i::nat. f -` A i \<inter> S) = space (vimage_algebra S f)"
+        using A f[THEN bij_betw_imp_funcset] by (auto simp: vimage_UN[symmetric])
+      show "range (\<lambda>i. f -` A i \<inter> S) \<subseteq> sets (vimage_algebra S f)"
+        using A by (auto simp: vimage_algebra_def)
+      have "\<And>i. f ` (f -` A i \<inter> S) = A i"
+        using f A sets_into_space
+        by (intro image_vimage_inter_eq) (auto simp: bij_betw_def)
+      then show "\<forall>i. \<mu> (f ` (f -` A i \<inter> S)) \<noteq> \<omega>"  using A by simp
+    qed
+  qed
+qed
+
 section "Real measure values"
 
 lemma (in measure_space) real_measure_Union:
@@ -638,12 +1086,12 @@
   finally show ?thesis .
 qed
 
-lemma (in measure_space) real_measurable_countably_subadditive:
+lemma (in measure_space) real_measure_countably_subadditive:
   assumes "range f \<subseteq> sets M" and "(\<Sum>\<^isub>\<infinity> i. \<mu> (f i)) \<noteq> \<omega>"
   shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
 proof -
   have "real (\<mu> (\<Union>i. f i)) \<le> real (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
-    using assms by (auto intro!: real_of_pinfreal_mono measurable_countably_subadditive)
+    using assms by (auto intro!: real_of_pinfreal_mono measure_countably_subadditive)
   also have "\<dots> = (\<Sum> i. real (\<mu> (f i)))"
     using assms by (auto intro!: sums_unique psuminf_imp_suminf)
   finally show ?thesis .
@@ -725,6 +1173,17 @@
   show "\<mu> (space ?r) \<noteq> \<omega>" using finite_measure[OF `S \<in> sets M`] by auto
 qed
 
+lemma (in measure_space) restricted_to_finite_measure:
+  assumes "S \<in> sets M" "\<mu> S \<noteq> \<omega>"
+  shows "finite_measure (restricted_space S) \<mu>"
+proof -
+  have "measure_space (restricted_space S) \<mu>"
+    using `S \<in> sets M` by (rule restricted_measure_space)
+  then show ?thesis
+    unfolding finite_measure_def finite_measure_axioms_def
+    using assms by auto
+qed
+
 lemma (in finite_measure) real_finite_measure_Diff:
   assumes "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
   shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
@@ -761,10 +1220,10 @@
   shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
   using measurable measurable[THEN finite_measure] by (rule real_measure_subadditive)
 
-lemma (in finite_measure) real_finite_measurable_countably_subadditive:
+lemma (in finite_measure) real_finite_measure_countably_subadditive:
   assumes "range f \<subseteq> sets M" and "summable (\<lambda>i. real (\<mu> (f i)))"
   shows "real (\<mu> (\<Union>i. f i)) \<le> (\<Sum> i. real (\<mu> (f i)))"
-proof (rule real_measurable_countably_subadditive[OF assms(1)])
+proof (rule real_measure_countably_subadditive[OF assms(1)])
   have *: "\<And>i. f i \<in> sets M" using assms by auto
   have "(\<lambda>i. real (\<mu> (f i))) sums (\<Sum>i. real (\<mu> (f i)))"
     using assms(2) by (rule summable_sums)
@@ -817,11 +1276,11 @@
     {f \<in> measurable A B. (\<forall>y \<in> sets B. \<mu> (f -` y \<inter> space A) = \<nu> y)}"
 
 lemma (in finite_measure) measure_preserving_lift:
-  fixes f :: "'a \<Rightarrow> 'a2" and A :: "('a2, 'b2) algebra_scheme"
+  fixes f :: "'a \<Rightarrow> 'a2" and A :: "'a2 algebra"
   assumes "algebra A"
-  assumes "finite_measure (sigma (space A) (sets A)) \<nu>" (is "finite_measure ?sA _")
+  assumes "finite_measure (sigma A) \<nu>" (is "finite_measure ?sA _")
   assumes fmp: "f \<in> measure_preserving M \<mu> A \<nu>"
-  shows "f \<in> measure_preserving M \<mu> (sigma (space A) (sets A)) \<nu>"
+  shows "f \<in> measure_preserving M \<mu> (sigma A) \<nu>"
 proof -
   interpret sA: finite_measure ?sA \<nu> by fact
   interpret A: algebra A by fact
@@ -916,24 +1375,8 @@
 
 section "Finite spaces"
 
-locale finite_measure_space = measure_space +
-  assumes finite_space: "finite (space M)"
-  and sets_eq_Pow: "sets M = Pow (space M)"
-  and finite_single_measure: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
-
-lemma (in finite_measure_space) sets_image_space_eq_Pow:
-  "sets (image_space X) = Pow (space (image_space X))"
-proof safe
-  fix x S assume "S \<in> sets (image_space X)" "x \<in> S"
-  then show "x \<in> space (image_space X)"
-    using sets_into_space by (auto intro!: imageI simp: image_space_def)
-next
-  fix S assume "S \<subseteq> space (image_space X)"
-  then obtain S' where "S = X`S'" "S'\<in>sets M"
-    by (auto simp: subset_image_iff sets_eq_Pow image_space_def)
-  then show "S \<in> sets (image_space X)"
-    by (auto simp: image_space_def)
-qed
+locale finite_measure_space = measure_space + finite_sigma_algebra +
+  assumes finite_single_measure[simp]: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
 
 lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
   using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
@@ -945,22 +1388,22 @@
     and "\<mu> {} = 0"
   shows "finite_measure_space M \<mu>"
     unfolding finite_measure_space_def finite_measure_space_axioms_def
-proof (safe del: notI)
+proof (intro allI impI conjI)
   show "measure_space M \<mu>"
   proof (rule sigma_algebra.finite_additivity_sufficient)
+    have *: "\<lparr>space = space M, sets = sets M\<rparr> = M" by auto
     show "sigma_algebra M"
-      apply (rule sigma_algebra_cong)
-      apply (rule sigma_algebra_Pow[of "space M"])
-      using assms by simp_all
+      using sigma_algebra_Pow[of "space M" "more M"] assms(2)[symmetric] by (simp add: *)
     show "finite (space M)" by fact
     show "positive \<mu>" unfolding positive_def by fact
     show "additive M \<mu>" unfolding additive_def using assms by simp
   qed
-  show "finite (space M)" by fact
-  { fix A x assume "A \<in> sets M" "x \<in> A" then show "x \<in> space M"
-      using assms by auto }
-  { fix A assume "A \<subseteq> space M" then show "A \<in> sets M"
-      using assms by auto }
+  then interpret measure_space M \<mu> .
+  show "finite_sigma_algebra M"
+  proof
+    show "finite (space M)" by fact
+    show "sets M = Pow (space M)" using assms by auto
+  qed
   { fix x assume *: "x \<in> space M"
     with add[of "{x}" "space M - {x}"] space
     show "\<mu> {x} \<noteq> \<omega>" by (auto simp: insert_absorb[OF *] Diff_subset) }
--- a/src/HOL/Probability/Positive_Infinite_Real.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Positive_Infinite_Real.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -6,6 +6,14 @@
   imports Complex_Main Nat_Bijection Multivariate_Analysis
 begin
 
+lemma range_const[simp]: "range (\<lambda>x. c) = {c}" by auto
+
+lemma (in complete_lattice) SUPR_const[simp]: "(SUP i. c) = c"
+  unfolding SUPR_def by simp
+
+lemma (in complete_lattice) INFI_const[simp]: "(INF i. c) = c"
+  unfolding INFI_def by simp
+
 lemma (in complete_lattice) Sup_start:
   assumes *: "\<And>x. f x \<le> f 0"
   shows "(SUP n. f n) = f 0"
@@ -415,6 +423,13 @@
   fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
   by (cases x, cases y) (auto simp: zero_less_mult_iff)
 
+lemma pinfreal_ord_one[simp]:
+  "Real p < 1 \<longleftrightarrow> p < 1"
+  "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
+  "1 < Real p \<longleftrightarrow> 1 < p"
+  "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
+  by (simp_all add: one_pinfreal_def del: Real_1)
+
 subsection {* @{text "x - y"} on @{typ pinfreal} *}
 
 instantiation pinfreal :: minus
@@ -540,6 +555,26 @@
   shows "real a \<le> real b"
 using assms by (cases b, cases a) auto
 
+lemma setprod_pinfreal_0:
+  "(\<Prod>i\<in>I. f i) = (0::pinfreal) \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = 0)"
+proof cases
+  assume "finite I" then show ?thesis
+  proof (induct I)
+    case (insert i I)
+    then show ?case by simp
+  qed simp
+qed simp
+
+lemma setprod_\<omega>:
+  "(\<Prod>i\<in>I. f i) = \<omega> \<longleftrightarrow> finite I \<and> (\<exists>i\<in>I. f i = \<omega>) \<and> (\<forall>i\<in>I. f i \<noteq> 0)"
+proof cases
+  assume "finite I" then show ?thesis
+  proof (induct I)
+    case (insert i I) then show ?case
+      by (auto simp: setprod_pinfreal_0)
+  qed simp
+qed simp
+
 instance pinfreal :: "semiring_char_0"
 proof
   fix m n
@@ -1573,7 +1608,25 @@
 definition (in complete_lattice) antiton (infix "\<down>" 50) where
   "A \<down> X \<longleftrightarrow> (\<forall>i. A i \<ge> A (Suc i)) \<and> (INF i. A i) = X"
 
-lemma range_const[simp]: "range (\<lambda>x. c) = {c}" by auto
+lemma isotoneI[intro?]: "\<lbrakk> \<And>i. f i \<le> f (Suc i) ; (SUP i. f i) = F \<rbrakk> \<Longrightarrow> f \<up> F"
+  unfolding isoton_def by auto
+
+lemma (in complete_lattice) isotonD[dest]:
+  assumes "A \<up> X" shows "A i \<le> A (Suc i)" "(SUP i. A i) = X"
+  using assms unfolding isoton_def by auto
+
+lemma isotonD'[dest]:
+  assumes "(A::_=>_) \<up> X" shows "A i x \<le> A (Suc i) x" "(SUP i. A i) = X"
+  using assms unfolding isoton_def le_fun_def by auto
+
+lemma isoton_mono_le:
+  assumes "f \<up> x" "i \<le> j"
+  shows "f i \<le> f j"
+  using `f \<up> x`[THEN isotonD(1)] lift_Suc_mono_le[of f, OF _ `i \<le> j`] by auto
+
+lemma isoton_const:
+  shows "(\<lambda> i. c) \<up> c"
+unfolding isoton_def by auto
 
 lemma isoton_cmult_right:
   assumes "f \<up> (x::pinfreal)"
@@ -1605,12 +1658,23 @@
   shows "(\<lambda>i x. f i x * indicator A x) \<up> (\<lambda>x. g x * indicator A x :: pinfreal)"
   using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left)
 
-lemma pinfreal_ord_one[simp]:
-  "Real p < 1 \<longleftrightarrow> p < 1"
-  "Real p \<le> 1 \<longleftrightarrow> p \<le> 1"
-  "1 < Real p \<longleftrightarrow> 1 < p"
-  "1 \<le> Real p \<longleftrightarrow> 1 \<le> p"
-  by (simp_all add: one_pinfreal_def del: Real_1)
+lemma isoton_setsum:
+  fixes f :: "'a \<Rightarrow> nat \<Rightarrow> pinfreal"
+  assumes "finite A" "A \<noteq> {}"
+  assumes "\<And> x. x \<in> A \<Longrightarrow> f x \<up> y x"
+  shows "(\<lambda> i. (\<Sum> x \<in> A. f x i)) \<up> (\<Sum> x \<in> A. y x)"
+using assms
+proof (induct A rule:finite_ne_induct)
+  case singleton thus ?case by auto
+next
+  case (insert a A) note asms = this
+  hence *: "(\<lambda> i. \<Sum> x \<in> A. f x i) \<up> (\<Sum> x \<in> A. y x)" by auto
+  have **: "(\<lambda> i. f a i) \<up> y a" using asms by simp
+  have "(\<lambda> i. f a i + (\<Sum> x \<in> A. f x i)) \<up> (y a + (\<Sum> x \<in> A. y x))"
+    using * ** isoton_add by auto
+  thus "(\<lambda> i. \<Sum> x \<in> insert a A. f x i) \<up> (\<Sum> x \<in> insert a A. y x)"
+    using asms by fastsimp
+qed
 
 lemma isoton_Sup:
   assumes "f \<up> u"
@@ -1981,14 +2045,6 @@
   apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe)
   apply(rule_tac y="Real (max 1 B)" in order_trans) by auto
 
-lemma (in complete_lattice) isotonD[dest]:
-  assumes "A \<up> X" shows "A i \<le> A (Suc i)" "(SUP i. A i) = X"
-  using assms unfolding isoton_def by auto
-
-lemma isotonD'[dest]:
-  assumes "(A::_=>_) \<up> X" shows "A i x \<le> A (Suc i) x" "(SUP i. A i) = X"
-  using assms unfolding isoton_def le_fun_def by auto
-
 lemma pinfreal_LimI_finite:
   assumes "x \<noteq> \<omega>" "\<And>r. 0 < r \<Longrightarrow> \<exists>N. \<forall>n\<ge>N. u n < x + r \<and> x < u n + r"
   shows "u ----> x"
--- a/src/HOL/Probability/Probability.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Probability.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -3,5 +3,4 @@
   Information
   "ex/Dining_Cryptographers"
 begin
-
 end
--- a/src/HOL/Probability/Probability_Space.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Probability_Space.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -1,34 +1,48 @@
 theory Probability_Space
-imports Lebesgue_Integration Radon_Nikodym
+imports Lebesgue_Integration Radon_Nikodym Product_Measure
 begin
 
+lemma real_of_pinfreal_inverse[simp]:
+  fixes X :: pinfreal
+  shows "real (inverse X) = 1 / real X"
+  by (cases X) (auto simp: inverse_eq_divide)
+
+lemma real_of_pinfreal_le_0[simp]: "real (X :: pinfreal) \<le> 0 \<longleftrightarrow> (X = 0 \<or> X = \<omega>)"
+  by (cases X) auto
+
+lemma real_of_pinfreal_less_0[simp]: "\<not> (real (X :: pinfreal) < 0)"
+  by (cases X) auto
+
 locale prob_space = measure_space +
   assumes measure_space_1: "\<mu> (space M) = 1"
 
+lemma abs_real_of_pinfreal[simp]: "\<bar>real (X :: pinfreal)\<bar> = real X"
+  by simp
+
+lemma zero_less_real_of_pinfreal: "0 < real (X :: pinfreal) \<longleftrightarrow> X \<noteq> 0 \<and> X \<noteq> \<omega>"
+  by (cases X) auto
+
 sublocale prob_space < finite_measure
 proof
   from measure_space_1 show "\<mu> (space M) \<noteq> \<omega>" by simp
 qed
 
-context prob_space
-begin
+abbreviation (in prob_space) "events \<equiv> sets M"
+abbreviation (in prob_space) "prob \<equiv> \<lambda>A. real (\<mu> A)"
+abbreviation (in prob_space) "prob_preserving \<equiv> measure_preserving"
+abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
+abbreviation (in prob_space) "expectation \<equiv> integral"
 
-abbreviation "events \<equiv> sets M"
-abbreviation "prob \<equiv> \<lambda>A. real (\<mu> A)"
-abbreviation "prob_preserving \<equiv> measure_preserving"
-abbreviation "random_variable \<equiv> \<lambda> s X. X \<in> measurable M s"
-abbreviation "expectation \<equiv> integral"
-
-definition
+definition (in prob_space)
   "indep A B \<longleftrightarrow> A \<in> events \<and> B \<in> events \<and> prob (A \<inter> B) = prob A * prob B"
 
-definition
+definition (in prob_space)
   "indep_families F G \<longleftrightarrow> (\<forall> A \<in> F. \<forall> B \<in> G. indep A B)"
 
-definition
+definition (in prob_space)
   "distribution X = (\<lambda>s. \<mu> ((X -` s) \<inter> (space M)))"
 
-abbreviation
+abbreviation (in prob_space)
   "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
 
 lemma (in prob_space) distribution_cong:
@@ -44,10 +58,14 @@
   unfolding distribution_def fun_eq_iff
   using assms by (auto intro!: arg_cong[where f="\<mu>"])
 
-lemma prob_space: "prob (space M) = 1"
+lemma (in prob_space) distribution_id[simp]:
+  assumes "N \<in> sets M" shows "distribution (\<lambda>x. x) N = \<mu> N"
+  using assms by (auto simp: distribution_def intro!: arg_cong[where f=\<mu>])
+
+lemma (in prob_space) prob_space: "prob (space M) = 1"
   unfolding measure_space_1 by simp
 
-lemma measure_le_1[simp, intro]:
+lemma (in prob_space) measure_le_1[simp, intro]:
   assumes "A \<in> events" shows "\<mu> A \<le> 1"
 proof -
   have "\<mu> A \<le> \<mu> (space M)"
@@ -56,21 +74,21 @@
   finally show ?thesis .
 qed
 
-lemma prob_compl:
+lemma (in prob_space) prob_compl:
   assumes "A \<in> events"
   shows "prob (space M - A) = 1 - prob A"
   using `A \<in> events`[THEN sets_into_space] `A \<in> events` measure_space_1
   by (subst real_finite_measure_Diff) auto
 
-lemma indep_space:
+lemma (in prob_space) indep_space:
   assumes "s \<in> events"
   shows "indep (space M) s"
   using assms prob_space by (simp add: indep_def)
 
-lemma prob_space_increasing: "increasing M prob"
+lemma (in prob_space) prob_space_increasing: "increasing M prob"
   by (auto intro!: real_measure_mono simp: increasing_def)
 
-lemma prob_zero_union:
+lemma (in prob_space) prob_zero_union:
   assumes "s \<in> events" "t \<in> events" "prob t = 0"
   shows "prob (s \<union> t) = prob s"
 using assms
@@ -82,13 +100,13 @@
   ultimately show ?thesis by simp
 qed
 
-lemma prob_eq_compl:
+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 prob_one_inter:
+lemma (in prob_space) prob_one_inter:
   assumes events:"s \<in> events" "t \<in> events"
   assumes "prob t = 1"
   shows "prob (s \<inter> t) = prob s"
@@ -101,7 +119,7 @@
     using events by (auto intro!: prob_eq_compl[of "s \<inter> t" s])
 qed
 
-lemma prob_eq_bigunion_image:
+lemma (in prob_space) prob_eq_bigunion_image:
   assumes "range f \<subseteq> events" "range g \<subseteq> events"
   assumes "disjoint_family f" "disjoint_family g"
   assumes "\<And> n :: nat. prob (f n) = prob (g n)"
@@ -115,27 +133,27 @@
   show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp
 qed
 
-lemma prob_countably_zero:
+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"
 proof (rule antisym)
   show "prob (\<Union> i :: nat. c i) \<le> 0"
-    using real_finite_measurable_countably_subadditive[OF assms(1)]
+    using real_finite_measure_countably_subadditive[OF assms(1)]
     by (simp add: assms(2) suminf_zero summable_zero)
   show "0 \<le> prob (\<Union> i :: nat. c i)" by (rule real_pinfreal_nonneg)
 qed
 
-lemma indep_sym:
+lemma (in prob_space) indep_sym:
    "indep a b \<Longrightarrow> indep b a"
 unfolding indep_def using Int_commute[of a b] by auto
 
-lemma indep_refl:
+lemma (in prob_space) indep_refl:
   assumes "a \<in> events"
   shows "indep a a = (prob a = 0) \<or> (prob a = 1)"
 using assms unfolding indep_def by auto
 
-lemma prob_equiprobable_finite_unions:
+lemma (in prob_space) prob_equiprobable_finite_unions:
   assumes "s \<in> events"
   assumes s_finite: "finite s" "\<And>x. x \<in> s \<Longrightarrow> {x} \<in> events"
   assumes "\<And> x y. \<lbrakk>x \<in> s; y \<in> s\<rbrakk> \<Longrightarrow> (prob {x} = prob {y})"
@@ -152,7 +170,7 @@
   finally show ?thesis by simp
 qed simp
 
-lemma prob_real_sum_image_fn:
+lemma (in prob_space) prob_real_sum_image_fn:
   assumes "e \<in> events"
   assumes "\<And> x. x \<in> s \<Longrightarrow> e \<inter> f x \<in> events"
   assumes "finite s"
@@ -173,12 +191,12 @@
   finally show ?thesis .
 qed
 
-lemma distribution_prob_space:
-  assumes S: "sigma_algebra S" "random_variable S X"
+lemma (in prob_space) distribution_prob_space:
+  assumes "random_variable S X"
   shows "prob_space S (distribution X)"
 proof -
   interpret S: measure_space S "distribution X"
-    using measure_space_vimage[OF S(2,1)] unfolding distribution_def .
+    using measure_space_vimage[of X S] assms unfolding distribution_def by simp
   show ?thesis
   proof
     have "X -` space S \<inter> space M = space M"
@@ -188,7 +206,19 @@
   qed
 qed
 
-lemma distribution_lebesgue_thm1:
+lemma (in prob_space) AE_distribution:
+  assumes X: "random_variable MX X" and "measure_space.almost_everywhere MX (distribution X) (\<lambda>x. Q x)"
+  shows "AE x. Q (X x)"
+proof -
+  interpret X: prob_space MX "distribution X" using X by (rule distribution_prob_space)
+  obtain N where N: "N \<in> sets MX" "distribution X N = 0" "{x\<in>space MX. \<not> Q x} \<subseteq> N"
+    using assms unfolding X.almost_everywhere_def by auto
+  show "AE x. Q (X x)"
+    using X[unfolded measurable_def] N unfolding distribution_def
+    by (intro AE_I'[where N="X -` N \<inter> space M"]) auto
+qed
+
+lemma (in prob_space) distribution_lebesgue_thm1:
   assumes "random_variable s X"
   assumes "A \<in> sets s"
   shows "real (distribution X A) = expectation (indicator (X -` A \<inter> space M))"
@@ -196,34 +226,34 @@
 using assms unfolding measurable_def
 using integral_indicator by auto
 
-lemma distribution_lebesgue_thm2:
-  assumes "sigma_algebra S" "random_variable S X" and "A \<in> sets S"
+lemma (in prob_space) distribution_lebesgue_thm2:
+  assumes "random_variable S X" and "A \<in> sets S"
   shows "distribution X A =
     measure_space.positive_integral S (distribution X) (indicator A)"
   (is "_ = measure_space.positive_integral _ ?D _")
 proof -
-  interpret S: prob_space S "distribution X" using assms(1,2) by (rule distribution_prob_space)
+  interpret S: prob_space S "distribution X" using assms(1) by (rule distribution_prob_space)
 
   show ?thesis
     using S.positive_integral_indicator(1)
     using assms unfolding distribution_def by auto
 qed
 
-lemma finite_expectation1:
-  assumes "finite (X`space M)" and rv: "random_variable borel_space X"
+lemma (in prob_space) finite_expectation1:
+  assumes f: "finite (X`space M)" and rv: "random_variable borel X"
   shows "expectation X = (\<Sum> r \<in> X ` space M. r * prob (X -` {r} \<inter> space M))"
-proof (rule integral_on_finite(2)[OF assms(2,1)])
+proof (rule integral_on_finite(2)[OF rv[THEN conjunct2] f])
   fix x have "X -` {x} \<inter> space M \<in> sets M"
     using rv unfolding measurable_def by auto
   thus "\<mu> (X -` {x} \<inter> space M) \<noteq> \<omega>" using finite_measure by simp
 qed
 
-lemma finite_expectation:
-  assumes "finite (space M)" "random_variable borel_space X"
+lemma (in prob_space) finite_expectation:
+  assumes "finite (space M)" "random_variable borel X"
   shows "expectation X = (\<Sum> r \<in> X ` (space M). r * real (distribution X {r}))"
   using assms unfolding distribution_def using finite_expectation1 by auto
 
-lemma prob_x_eq_1_imp_prob_y_eq_0:
+lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
   assumes "{x} \<in> events"
   assumes "prob {x} = 1"
   assumes "{y} \<in> events"
@@ -231,17 +261,17 @@
   shows "prob {y} = 0"
   using prob_one_inter[of "{y}" "{x}"] assms by auto
 
-lemma distribution_empty[simp]: "distribution X {} = 0"
+lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
   unfolding distribution_def by simp
 
-lemma distribution_space[simp]: "distribution X (X ` space M) = 1"
+lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1"
 proof -
   have "X -` X ` space M \<inter> space M = space M" by auto
   thus ?thesis unfolding distribution_def by (simp add: measure_space_1)
 qed
 
-lemma distribution_one:
-  assumes "random_variable M X" and "A \<in> events"
+lemma (in prob_space) distribution_one:
+  assumes "random_variable M' X" and "A \<in> sets M'"
   shows "distribution X A \<le> 1"
 proof -
   have "distribution X A \<le> \<mu> (space M)" unfolding distribution_def
@@ -249,31 +279,27 @@
   thus ?thesis by (simp add: measure_space_1)
 qed
 
-lemma distribution_finite:
-  assumes "random_variable M X" and "A \<in> events"
+lemma (in prob_space) distribution_finite:
+  assumes "random_variable M' X" and "A \<in> sets M'"
   shows "distribution X A \<noteq> \<omega>"
   using distribution_one[OF assms] by auto
 
-lemma distribution_x_eq_1_imp_distribution_y_eq_0:
+lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0:
   assumes X: "random_variable \<lparr>space = X ` (space M), sets = Pow (X ` (space M))\<rparr> X"
     (is "random_variable ?S X")
   assumes "distribution X {x} = 1"
   assumes "y \<noteq> x"
   shows "distribution X {y} = 0"
 proof -
-  have "sigma_algebra ?S" by (rule sigma_algebra_Pow)
-  from distribution_prob_space[OF this X]
+  from distribution_prob_space[OF X]
   interpret S: prob_space ?S "distribution X" by simp
-
   have x: "{x} \<in> sets ?S"
   proof (rule ccontr)
     assume "{x} \<notin> sets ?S"
     hence "X -` {x} \<inter> space M = {}" by auto
     thus "False" using assms unfolding distribution_def by auto
   qed
-
   have [simp]: "{y} \<inter> {x} = {}" "{x} - {y} = {x}" using `y \<noteq> x` by auto
-
   show ?thesis
   proof cases
     assume "{y} \<in> sets ?S"
@@ -287,10 +313,341 @@
   qed
 qed
 
-end
+lemma (in prob_space) joint_distribution_Times_le_fst:
+  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
+    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
+  unfolding distribution_def
+proof (intro measure_mono)
+  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
+  show "X -` A \<inter> space M \<in> events"
+    using X A unfolding measurable_def by simp
+  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
+    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
+  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<in> events"
+    unfolding * apply (rule Int)
+    using assms unfolding measurable_def by auto
+qed
+
+lemma (in prob_space) joint_distribution_commute:
+  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+
+lemma (in prob_space) joint_distribution_Times_le_snd:
+  assumes X: "random_variable MX X" and Y: "random_variable MY Y"
+    and A: "A \<in> sets MX" and B: "B \<in> sets MY"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
+  using assms
+  by (subst joint_distribution_commute)
+     (simp add: swap_product joint_distribution_Times_le_fst)
+
+lemma (in prob_space) random_variable_pairI:
+  assumes "random_variable MX X"
+  assumes "random_variable MY Y"
+  shows "random_variable (sigma (pair_algebra MX MY)) (\<lambda>x. (X x, Y x))"
+proof
+  interpret MX: sigma_algebra MX using assms by simp
+  interpret MY: sigma_algebra MY using assms by simp
+  interpret P: pair_sigma_algebra MX MY by default
+  show "sigma_algebra (sigma (pair_algebra MX MY))" by default
+  have sa: "sigma_algebra M" by default
+  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (sigma (pair_algebra MX MY))"
+    unfolding P.measurable_pair[OF sa] using assms by (simp add: comp_def)
+qed
+
+lemma (in prob_space) distribution_order:
+  assumes "random_variable MX X" "random_variable MY Y"
+  assumes "{x} \<in> sets MX" "{y} \<in> sets MY"
+  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
+    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
+    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
+    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
+    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+  using joint_distribution_Times_le_snd[OF assms]
+  using joint_distribution_Times_le_fst[OF assms]
+  by auto
+
+lemma (in prob_space) joint_distribution_commute_singleton:
+  "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+
+lemma (in prob_space) joint_distribution_assoc_singleton:
+  "joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)} =
+   joint_distribution (\<lambda>x. (X x, Y x)) Z {((x, y), z)}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+
+locale pair_prob_space = M1: prob_space M1 p1 + M2: prob_space M2 p2 for M1 p1 M2 p2
+
+sublocale pair_prob_space \<subseteq> pair_sigma_finite M1 p1 M2 p2 by default
+
+sublocale pair_prob_space \<subseteq> P: prob_space P pair_measure
+proof
+  show "pair_measure (space P) = 1"
+    by (simp add: pair_algebra_def pair_measure_times M1.measure_space_1 M2.measure_space_1)
+qed
+
+lemma countably_additiveI[case_names countably]:
+  assumes "\<And>A. \<lbrakk> range A \<subseteq> sets M ; disjoint_family A ; (\<Union>i. A i) \<in> sets M\<rbrakk> \<Longrightarrow>
+    (\<Sum>\<^isub>\<infinity>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+  shows "countably_additive M \<mu>"
+  using assms unfolding countably_additive_def by auto
+
+lemma (in prob_space) joint_distribution_prob_space:
+  assumes "random_variable MX X" "random_variable MY Y"
+  shows "prob_space (sigma (pair_algebra MX MY)) (joint_distribution X Y)"
+proof -
+  interpret X: prob_space MX "distribution X" by (intro distribution_prob_space assms)
+  interpret Y: prob_space MY "distribution Y" by (intro distribution_prob_space assms)
+  interpret XY: pair_sigma_finite MX "distribution X" MY "distribution Y" by default
+  show ?thesis
+  proof
+    let "?X A" = "(\<lambda>x. (X x, Y x)) -` A \<inter> space M"
+    show "joint_distribution X Y {} = 0" by (simp add: distribution_def)
+    show "countably_additive XY.P (joint_distribution X Y)"
+    proof (rule countably_additiveI)
+      fix A :: "nat \<Rightarrow> ('b \<times> 'c) set"
+      assume A: "range A \<subseteq> sets XY.P" and df: "disjoint_family A"
+      have "(\<Sum>\<^isub>\<infinity>n. \<mu> (?X (A n))) = \<mu> (\<Union>x. ?X (A x))"
+      proof (intro measure_countably_additive)
+        from assms have *: "(\<lambda>x. (X x, Y x)) \<in> measurable M XY.P"
+          by (intro XY.measurable_prod_sigma) (simp_all add: comp_def, default)
+        show "range (\<lambda>n. ?X (A n)) \<subseteq> events"
+          using measurable_sets[OF *] A by auto
+        show "disjoint_family (\<lambda>n. ?X (A n))"
+          by (intro disjoint_family_on_bisimulation[OF df]) auto
+      qed
+      then show "(\<Sum>\<^isub>\<infinity>n. joint_distribution X Y (A n)) = joint_distribution X Y (\<Union>i. A i)"
+        by (simp add: distribution_def vimage_UN)
+    qed
+    have "?X (space MX \<times> space MY) = space M"
+      using assms by (auto simp: measurable_def)
+    then show "joint_distribution X Y (space XY.P) = 1"
+      by (simp add: space_pair_algebra distribution_def measure_space_1)
+  qed
+qed
+
+section "Probability spaces on finite sets"
 
 locale finite_prob_space = prob_space + finite_measure_space
 
+abbreviation (in prob_space) "finite_random_variable M' X \<equiv> finite_sigma_algebra M' \<and> X \<in> measurable M M'"
+
+lemma (in prob_space) finite_random_variableD:
+  assumes "finite_random_variable M' X" shows "random_variable M' X"
+proof -
+  interpret M': finite_sigma_algebra M' using assms by simp
+  then show "random_variable M' X" using assms by simp default
+qed
+
+lemma (in prob_space) distribution_finite_prob_space:
+  assumes "finite_random_variable MX X"
+  shows "finite_prob_space MX (distribution X)"
+proof -
+  interpret X: prob_space MX "distribution X"
+    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
+  interpret MX: finite_sigma_algebra MX
+    using assms by simp
+  show ?thesis
+  proof
+    fix x assume "x \<in> space MX"
+    then have "X -` {x} \<inter> space M \<in> sets M"
+      using assms unfolding measurable_def by simp
+    then show "distribution X {x} \<noteq> \<omega>"
+      unfolding distribution_def by simp
+  qed
+qed
+
+lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
+  assumes "simple_function X"
+  shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
+proof (intro conjI)
+  have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
+  interpret X: sigma_algebra "\<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr>"
+    by (rule sigma_algebra_Pow)
+  show "finite_sigma_algebra \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr>"
+    by default auto
+  show "X \<in> measurable M \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr>"
+  proof (unfold measurable_def, clarsimp)
+    fix A assume A: "A \<subseteq> X`space M"
+    then have "finite A" by (rule finite_subset) simp
+    then have "X -` (\<Union>a\<in>A. {a}) \<inter> space M \<in> events"
+      unfolding vimage_UN UN_extend_simps
+      apply (rule finite_UN)
+      using A assms unfolding simple_function_def by auto
+    then show "X -` A \<inter> space M \<in> events" by simp
+  qed
+qed
+
+lemma (in prob_space) simple_function_imp_random_variable[simp, intro]:
+  assumes "simple_function X"
+  shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
+  using simple_function_imp_finite_random_variable[OF assms]
+  by (auto dest!: finite_random_variableD)
+
+lemma (in prob_space) sum_over_space_real_distribution:
+  "simple_function X \<Longrightarrow> (\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
+  unfolding distribution_def prob_space[symmetric]
+  by (subst real_finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def simple_function_def
+           intro!: arg_cong[where f=prob])
+
+lemma (in prob_space) finite_random_variable_pairI:
+  assumes "finite_random_variable MX X"
+  assumes "finite_random_variable MY Y"
+  shows "finite_random_variable (sigma (pair_algebra MX MY)) (\<lambda>x. (X x, Y x))"
+proof
+  interpret MX: finite_sigma_algebra MX using assms by simp
+  interpret MY: finite_sigma_algebra MY using assms by simp
+  interpret P: pair_finite_sigma_algebra MX MY by default
+  show "finite_sigma_algebra (sigma (pair_algebra MX MY))" by default
+  have sa: "sigma_algebra M" by default
+  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (sigma (pair_algebra MX MY))"
+    unfolding P.measurable_pair[OF sa] using assms by (simp add: comp_def)
+qed
+
+lemma (in prob_space) finite_random_variable_imp_sets:
+  "finite_random_variable MX X \<Longrightarrow> x \<in> space MX \<Longrightarrow> {x} \<in> sets MX"
+  unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
+
+lemma (in prob_space) finite_random_variable_vimage:
+  assumes X: "finite_random_variable MX X" shows "X -` A \<inter> space M \<in> events"
+proof -
+  interpret X: finite_sigma_algebra MX using X by simp
+  from X have vimage: "\<And>A. A \<subseteq> space MX \<Longrightarrow> X -` A \<inter> space M \<in> events" and
+    "X \<in> space M \<rightarrow> space MX"
+    by (auto simp: measurable_def)
+  then have *: "X -` A \<inter> space M = X -` (A \<inter> space MX) \<inter> space M"
+    by auto
+  show "X -` A \<inter> space M \<in> events"
+    unfolding * by (intro vimage) auto
+qed
+
+lemma (in prob_space) joint_distribution_finite_Times_le_fst:
+  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution X A"
+  unfolding distribution_def
+proof (intro measure_mono)
+  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<subseteq> X -` A \<inter> space M" by force
+  show "X -` A \<inter> space M \<in> events"
+    using finite_random_variable_vimage[OF X] .
+  have *: "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M =
+    (X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
+  show "(\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M \<in> events"
+    unfolding * apply (rule Int)
+    using assms[THEN finite_random_variable_vimage] by auto
+qed
+
+lemma (in prob_space) joint_distribution_finite_Times_le_snd:
+  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y"
+  shows "joint_distribution X Y (A \<times> B) \<le> distribution Y B"
+  using assms
+  by (subst joint_distribution_commute)
+     (simp add: swap_product joint_distribution_finite_Times_le_fst)
+
+lemma (in prob_space) finite_distribution_order:
+  assumes "finite_random_variable MX X" "finite_random_variable MY Y"
+  shows "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution X {x}"
+    and "r \<le> joint_distribution X Y {(x, y)} \<Longrightarrow> r \<le> distribution Y {y}"
+    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution X {x}"
+    and "r < joint_distribution X Y {(x, y)} \<Longrightarrow> r < distribution Y {y}"
+    and "distribution X {x} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+    and "distribution Y {y} = 0 \<Longrightarrow> joint_distribution X Y {(x, y)} = 0"
+  using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"]
+  using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"]
+  by auto
+
+lemma (in prob_space) finite_distribution_finite:
+  assumes "finite_random_variable M' X"
+  shows "distribution X {x} \<noteq> \<omega>"
+proof -
+  have "distribution X {x} \<le> \<mu> (space M)"
+    unfolding distribution_def
+    using finite_random_variable_vimage[OF assms]
+    by (intro measure_mono) auto
+  then show ?thesis
+    by auto
+qed
+
+lemma (in prob_space) setsum_joint_distribution:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "random_variable MY Y" "B \<in> sets MY"
+  shows "(\<Sum>a\<in>space MX. joint_distribution X Y ({a} \<times> B)) = distribution Y B"
+  unfolding distribution_def
+proof (subst measure_finitely_additive'')
+  interpret MX: finite_sigma_algebra MX using X by auto
+  show "finite (space MX)" using MX.finite_space .
+  let "?d i" = "(\<lambda>x. (X x, Y x)) -` ({i} \<times> B) \<inter> space M"
+  { fix i assume "i \<in> space MX"
+    moreover have "?d i = (X -` {i} \<inter> space M) \<inter> (Y -` B \<inter> space M)" by auto
+    ultimately show "?d i \<in> events"
+      using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y
+      using MX.sets_eq_Pow by auto }
+  show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def)
+  show "\<mu> (\<Union>i\<in>space MX. ?d i) = \<mu> (Y -` B \<inter> space M)"
+    using X[unfolded measurable_def]
+    by (auto intro!: arg_cong[where f=\<mu>])
+qed
+
+lemma (in prob_space) setsum_joint_distribution_singleton:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
+  shows "(\<Sum>a\<in>space MX. joint_distribution X Y {(a, b)}) = distribution Y {b}"
+  using setsum_joint_distribution[OF X
+    finite_random_variableD[OF Y(1)]
+    finite_random_variable_imp_sets[OF Y]] by simp
+
+lemma (in prob_space) setsum_real_joint_distribution:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "random_variable MY Y" "B \<in> sets MY"
+  shows "(\<Sum>a\<in>space MX. real (joint_distribution X Y ({a} \<times> B))) = real (distribution Y B)"
+proof -
+  interpret MX: finite_sigma_algebra MX using X by auto
+  show ?thesis
+    unfolding setsum_joint_distribution[OF assms, symmetric]
+    using distribution_finite[OF random_variable_pairI[OF finite_random_variableD[OF X] Y(1)]] Y(2)
+    by (simp add: space_pair_algebra in_sigma pair_algebraI MX.sets_eq_Pow real_of_pinfreal_setsum)
+qed
+
+lemma (in prob_space) setsum_real_joint_distribution_singleton:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "finite_random_variable MY Y" "b \<in> space MY"
+  shows "(\<Sum>a\<in>space MX. real (joint_distribution X Y {(a,b)})) = real (distribution Y {b})"
+  using setsum_real_joint_distribution[OF X
+    finite_random_variableD[OF Y(1)]
+    finite_random_variable_imp_sets[OF Y]] by simp
+
+locale pair_finite_prob_space = M1: finite_prob_space M1 p1 + M2: finite_prob_space M2 p2 for M1 p1 M2 p2
+
+sublocale pair_finite_prob_space \<subseteq> pair_prob_space M1 p1 M2 p2 by default
+sublocale pair_finite_prob_space \<subseteq> pair_finite_space M1 p1 M2 p2  by default
+sublocale pair_finite_prob_space \<subseteq> finite_prob_space P pair_measure by default
+
+lemma (in prob_space) joint_distribution_finite_prob_space:
+  assumes X: "finite_random_variable MX X"
+  assumes Y: "finite_random_variable MY Y"
+  shows "finite_prob_space (sigma (pair_algebra MX MY)) (joint_distribution X Y)"
+proof -
+  interpret X: finite_prob_space MX "distribution X"
+    using X by (rule distribution_finite_prob_space)
+  interpret Y: finite_prob_space MY "distribution Y"
+    using Y by (rule distribution_finite_prob_space)
+  interpret P: prob_space "sigma (pair_algebra MX MY)" "joint_distribution X Y"
+    using assms[THEN finite_random_variableD] by (rule joint_distribution_prob_space)
+  interpret XY: pair_finite_prob_space MX "distribution X" MY "distribution Y"
+    by default
+  show ?thesis
+  proof
+    fix x assume "x \<in> space XY.P"
+    moreover have "(\<lambda>x. (X x, Y x)) \<in> measurable M XY.P"
+      using X Y by (subst XY.measurable_pair) (simp_all add: o_def, default)
+    ultimately have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M \<in> sets M"
+      unfolding measurable_def by simp
+    then show "joint_distribution X Y {x} \<noteq> \<omega>"
+      unfolding distribution_def by simp
+  qed
+qed
+
 lemma finite_prob_space_eq:
   "finite_prob_space M \<mu> \<longleftrightarrow> finite_measure_space M \<mu> \<and> \<mu> (space M) = 1"
   unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
@@ -550,10 +907,6 @@
   "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
   by (simp add: finite_prob_space_eq finite_measure_space)
 
-lemma (in prob_space) joint_distribution_commute:
-  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
-
 lemma (in finite_prob_space) real_distribution_order':
   shows "real (distribution X {x}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
   and "real (distribution Y {y}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
@@ -586,6 +939,16 @@
                               (joint_distribution X Y)"
   using finite_space by (auto intro!: finite_product_measure_space)
 
+lemma (in finite_prob_space) finite_product_prob_space_of_images:
+  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr>
+                     (joint_distribution X Y)"
+  (is "finite_prob_space ?S _")
+proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
+  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
+  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
+    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
+qed
+
 section "Conditional Expectation and Probability"
 
 lemma (in prob_space) conditional_expectation_exists:
--- a/src/HOL/Probability/Product_Measure.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Product_Measure.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -2,367 +2,184 @@
 imports Lebesgue_Integration
 begin
 
-definition "dynkin M \<longleftrightarrow>
-  space M \<in> sets M \<and>
-  (\<forall> A \<in> sets M. A \<subseteq> space M) \<and>
-  (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M) \<and>
-  (\<forall>A. disjoint_family A \<and> range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
-
-lemma dynkinI:
-  assumes "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
-  assumes "space M \<in> sets M" and "\<forall> b \<in> sets M. \<forall> a \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M"
-  assumes "\<And> a. (\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {})
-          \<Longrightarrow> (\<And> i :: nat. a i \<in> sets M) \<Longrightarrow> UNION UNIV a \<in> sets M"
-  shows "dynkin M"
-using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff)
-
-lemma dynkin_subset:
-  assumes "dynkin M"
-  shows "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
-using assms unfolding dynkin_def by auto
-
-lemma dynkin_space:
-  assumes "dynkin M"
-  shows "space M \<in> sets M"
-using assms unfolding dynkin_def by auto
-
-lemma dynkin_diff:
-  assumes "dynkin M"
-  shows "\<And> a b. \<lbrakk> a \<in> sets M ; b \<in> sets M ; a \<subseteq> b \<rbrakk> \<Longrightarrow> b - a \<in> sets M"
-using assms unfolding dynkin_def by auto
-
-lemma dynkin_empty:
-  assumes "dynkin M"
-  shows "{} \<in> sets M"
-using dynkin_diff[OF assms dynkin_space[OF assms] dynkin_space[OF assms]] by auto
-
-lemma dynkin_UN:
-  assumes "dynkin M"
-  assumes "\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}"
-  assumes "\<And> i :: nat. a i \<in> sets M"
-  shows "UNION UNIV a \<in> sets M"
-using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff)
-
-definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)"
+lemma in_sigma[intro, simp]: "A \<in> sets M \<Longrightarrow> A \<in> sets (sigma M)"
+  unfolding sigma_def by (auto intro!: sigma_sets.Basic)
 
-lemma dynkin_trivial:
-  shows "dynkin \<lparr> space = A, sets = Pow A \<rparr>"
-by (rule dynkinI) auto
-
-lemma dynkin_lemma:
-  fixes D :: "'a algebra"
-  assumes stab: "Int_stable E"
-  and spac: "space E = space D"
-  and subsED: "sets E \<subseteq> sets D"
-  and subsDE: "sets D \<subseteq> sigma_sets (space E) (sets E)"
-  and dyn: "dynkin D"
-  shows "sigma (space E) (sets E) = D"
-proof -
-  def sets_\<delta>E == "\<Inter> {sets d | d :: 'a algebra. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
-  def \<delta>E == "\<lparr> space = space E, sets = sets_\<delta>E \<rparr>"
-  have "\<lparr> space = space E, sets = Pow (space E) \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
-    using dynkin_trivial spac subsED dynkin_subset[OF dyn] by fastsimp
-  hence not_empty: "{sets (d :: 'a algebra) | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d} \<noteq> {}"
-    using exI[of "\<lambda> x. space x = space E \<and> dynkin x \<and> sets E \<subseteq> sets x" "\<lparr> space = space E, sets = Pow (space E) \<rparr>", simplified]
-    by auto
-  have \<delta>E_D: "sets_\<delta>E \<subseteq> sets D"
-    unfolding sets_\<delta>E_def using assms by auto
-  have \<delta>ynkin: "dynkin \<delta>E"
-  proof (rule dynkinI, safe)
-    fix A x assume asm: "A \<in> sets \<delta>E" "x \<in> A"
-    { fix d :: "('a, 'b) algebra_scheme" assume "A \<in> sets d" "dynkin d \<and> space d = space E"
-      hence "A \<subseteq> space d" using dynkin_subset by auto }
-    show "x \<in> space \<delta>E" using asm unfolding \<delta>E_def sets_\<delta>E_def using not_empty
-      by simp (metis dynkin_subset in_mono mem_def)
-  next
-    show "space \<delta>E \<in> sets \<delta>E"
-      unfolding \<delta>E_def sets_\<delta>E_def
-      using dynkin_space by fastsimp
-  next
-    fix a b assume "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" "a \<subseteq> b"
-    thus "b - a \<in> sets \<delta>E"
-      unfolding \<delta>E_def sets_\<delta>E_def by (auto intro:dynkin_diff)
-  next
-    fix a assume asm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> sets \<delta>E"
-    thus "UNION UNIV a \<in> sets \<delta>E"
-      unfolding \<delta>E_def sets_\<delta>E_def apply (auto intro!:dynkin_UN[OF _ asm(1)])
-      by blast
-  qed
+lemma (in sigma_algebra) sigma_eq[simp]: "sigma M = M"
+  unfolding sigma_def sigma_sets_eq by simp
 
-  def Dy == "\<lambda> d. {A | A. A \<in> sets_\<delta>E \<and> A \<inter> d \<in> sets_\<delta>E}"
-  { fix d assume dasm: "d \<in> sets_\<delta>E"
-    have "dynkin \<lparr> space = space E, sets = Dy d \<rparr>"
-    proof (rule dynkinI, safe, simp_all)
-      fix A x assume "A \<in> Dy d" "x \<in> A"
-      thus "x \<in> space E" unfolding Dy_def sets_\<delta>E_def using not_empty
-        by simp (metis dynkin_subset in_mono mem_def)
-    next
-      show "space E \<in> Dy d"
-        unfolding Dy_def \<delta>E_def sets_\<delta>E_def
-      proof auto
-        fix d assume asm: "dynkin d" "space d = space E" "sets E \<subseteq> sets d"
-        hence "space d \<in> sets d" using dynkin_space[OF asm(1)] by auto
-        thus "space E \<in> sets d" using asm by auto
-      next
-        fix da :: "'a algebra" assume asm: "dynkin da" "space da = space E" "sets E \<subseteq> sets da"
-        have d: "d = space E \<inter> d"
-          using dasm dynkin_subset[OF asm(1)] asm(2) dynkin_subset[OF \<delta>ynkin]
-          unfolding \<delta>E_def by auto
-        hence "space E \<inter> d \<in> sets \<delta>E" unfolding \<delta>E_def
-          using dasm by auto
-        have "sets \<delta>E \<subseteq> sets da" unfolding \<delta>E_def sets_\<delta>E_def using asm
-          by auto
-        thus "space E \<inter> d \<in> sets da" using dasm asm d dynkin_subset[OF \<delta>ynkin]
-          unfolding \<delta>E_def by auto
-      qed
-    next
-      fix a b assume absm: "a \<in> Dy d" "b \<in> Dy d" "a \<subseteq> b"
-      hence "a \<in> sets \<delta>E" "b \<in> sets \<delta>E"
-        unfolding Dy_def \<delta>E_def by auto
-      hence *: "b - a \<in> sets \<delta>E"
-        using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto
-      have "a \<inter> d \<in> sets \<delta>E" "b \<inter> d \<in> sets \<delta>E"
-        using absm unfolding Dy_def \<delta>E_def by auto
-      hence "(b \<inter> d) - (a \<inter> d) \<in> sets \<delta>E"
-        using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto
-      hence **: "(b - a) \<inter> d \<in> sets \<delta>E" by (auto simp add:Diff_Int_distrib2)
-      thus "b - a \<in> Dy d"
-        using * ** unfolding Dy_def \<delta>E_def by auto
-    next
-      fix a assume aasm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> Dy d"
-      hence "\<And> i. a i \<in> sets \<delta>E"
-        unfolding Dy_def \<delta>E_def by auto
-      from dynkin_UN[OF \<delta>ynkin aasm(1) this]
-      have *: "UNION UNIV a \<in> sets \<delta>E" by auto
-      from aasm
-      have aE: "\<forall> i. a i \<inter> d \<in> sets \<delta>E"
-        unfolding Dy_def \<delta>E_def by auto
-      from aasm
-      have "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> (a i \<inter> d) \<inter> (a j \<inter> d) = {}" by auto
-      from dynkin_UN[OF \<delta>ynkin this]
-      have "UNION UNIV (\<lambda> i. a i \<inter> d) \<in> sets \<delta>E"
-        using aE by auto
-      hence **: "UNION UNIV a \<inter> d \<in> sets \<delta>E" by auto
-      from * ** show "UNION UNIV a \<in> Dy d" unfolding Dy_def \<delta>E_def by auto
-    qed } note Dy_nkin = this
-  have E_\<delta>E: "sets E \<subseteq> sets \<delta>E"
-    unfolding \<delta>E_def sets_\<delta>E_def by auto
-  { fix d assume dasm: "d \<in> sets \<delta>E"
-    { fix e assume easm: "e \<in> sets E"
-      hence deasm: "e \<in> sets \<delta>E"
-        unfolding \<delta>E_def sets_\<delta>E_def by auto
-      have subset: "Dy e \<subseteq> sets \<delta>E"
-        unfolding Dy_def \<delta>E_def by auto
-      { fix e' assume e'asm: "e' \<in> sets E"
-        have "e' \<inter> e \<in> sets E"
-          using easm e'asm stab unfolding Int_stable_def by auto
-        hence "e' \<inter> e \<in> sets \<delta>E"
-          unfolding \<delta>E_def sets_\<delta>E_def by auto
-        hence "e' \<in> Dy e" using e'asm unfolding Dy_def \<delta>E_def sets_\<delta>E_def by auto }
-      hence E_Dy: "sets E \<subseteq> Dy e" by auto
-      have "\<lparr> space = space E, sets = Dy e \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
-        using Dy_nkin[OF deasm[unfolded \<delta>E_def, simplified]] E_\<delta>E E_Dy by auto
-      hence "sets_\<delta>E \<subseteq> Dy e"
-        unfolding sets_\<delta>E_def by auto (metis E_Dy simps(1) simps(2) spac)
-      hence "sets \<delta>E = Dy e" using subset unfolding \<delta>E_def by auto
-      hence "d \<inter> e \<in> sets \<delta>E"
-        using dasm easm deasm unfolding Dy_def \<delta>E_def by auto
-      hence "e \<in> Dy d" using deasm
-        unfolding Dy_def \<delta>E_def
-        by (auto simp add:Int_commute) }
-    hence "sets E \<subseteq> Dy d" by auto
-    hence "sets \<delta>E \<subseteq> Dy d" using Dy_nkin[OF dasm[unfolded \<delta>E_def, simplified]]
-      unfolding \<delta>E_def sets_\<delta>E_def
-      by auto (metis `sets E <= Dy d` simps(1) simps(2) spac)
-    hence *: "sets \<delta>E = Dy d"
-      unfolding Dy_def \<delta>E_def by auto
-    fix a assume aasm: "a \<in> sets \<delta>E"
-    hence "a \<inter> d \<in> sets \<delta>E"
-      using * dasm unfolding Dy_def \<delta>E_def by auto } note \<delta>E_stab = this
-  { fix A :: "nat \<Rightarrow> 'a set" assume Asm: "range A \<subseteq> sets \<delta>E" "\<And>A. A \<in> sets \<delta>E \<Longrightarrow> A \<subseteq> space \<delta>E"
-      "\<And>a. a \<in> sets \<delta>E \<Longrightarrow> space \<delta>E - a \<in> sets \<delta>E"
-    "{} \<in> sets \<delta>E" "space \<delta>E \<in> sets \<delta>E"
-    let "?A i" = "A i \<inter> (\<Inter> j \<in> {..< i}. space \<delta>E - A j)"
-    { fix i :: nat
-      have *: "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<inter> space \<delta>E \<in> sets \<delta>E"
-        apply (induct i)
-        using lessThan_Suc Asm \<delta>E_stab apply fastsimp
-        apply (subst lessThan_Suc)
-        apply (subst INT_insert)
-        apply (subst Int_assoc)
-        apply (subst \<delta>E_stab)
-        using lessThan_Suc Asm \<delta>E_stab Asm
-        apply (fastsimp simp add:Int_assoc dynkin_diff[OF \<delta>ynkin])
-        prefer 2 apply simp
-        apply (rule dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]])
-        using Asm by auto
-      have **: "\<And> i. A i \<subseteq> space \<delta>E" using Asm by auto
-      have "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<subseteq> space \<delta>E \<or> (\<Inter> j \<in> {..< i}. A j) = UNIV \<and> i = 0"
-        apply (cases i)
-        using Asm ** dynkin_subset[OF \<delta>ynkin, of "A (i - 1)"]
-        by auto
-      hence Aisets: "?A i \<in> sets \<delta>E"
-        apply (cases i)
-        using Asm * apply fastsimp
-        apply (rule \<delta>E_stab)
-        using Asm * **
-        by (auto simp add:Int_absorb2)
-      have "?A i = disjointed A i" unfolding disjointed_def
-      atLeast0LessThan using Asm by auto
-      hence "?A i = disjointed A i" "?A i \<in> sets \<delta>E"
-        using Aisets by auto
-    } note Ai = this
-    from dynkin_UN[OF \<delta>ynkin _ this(2)] this disjoint_family_disjointed[of A]
-    have "(\<Union> i. ?A i) \<in> sets \<delta>E"
-      by (auto simp add:disjoint_family_on_def disjointed_def)
-    hence "(\<Union> i. A i) \<in> sets \<delta>E"
-      using Ai(1) UN_disjointed_eq[of A] by auto } note \<delta>E_UN = this
-  { fix a b assume asm: "a \<in> sets \<delta>E" "b \<in> sets \<delta>E"
-    let ?ab = "\<lambda> i. if (i::nat) = 0 then a else if i = 1 then b else {}"
-    have *: "(\<Union> i. ?ab i) \<in> sets \<delta>E"
-      apply (rule \<delta>E_UN)
-      using asm \<delta>E_UN dynkin_empty[OF \<delta>ynkin] 
-      dynkin_subset[OF \<delta>ynkin] 
-      dynkin_space[OF \<delta>ynkin]
-      dynkin_diff[OF \<delta>ynkin] by auto
-    have "(\<Union> i. ?ab i) = a \<union> b" apply auto
-      apply (case_tac "i = 0")
-      apply auto
-      apply (case_tac "i = 1")
-      by auto
-    hence "a \<union> b \<in> sets \<delta>E" using * by auto} note \<delta>E_Un = this
-  have "sigma_algebra \<delta>E"
-    apply unfold_locales
-    using dynkin_subset[OF \<delta>ynkin]
-    using dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]]
-    using dynkin_diff[OF \<delta>ynkin, of "space \<delta>E" "space \<delta>E", OF dynkin_space[OF \<delta>ynkin] dynkin_space[OF \<delta>ynkin]]
-    using dynkin_space[OF \<delta>ynkin]
-    using \<delta>E_UN \<delta>E_Un
-    by auto
-  from sigma_algebra.sigma_subset[OF this E_\<delta>E] \<delta>E_D subsDE spac
-  show ?thesis by (auto simp add:\<delta>E_def sigma_def)
+lemma vimage_algebra_sigma:
+  assumes E: "sets E \<subseteq> Pow (space E)"
+    and f: "f \<in> space F \<rightarrow> space E"
+    and "\<And>A. A \<in> sets F \<Longrightarrow> A \<in> (\<lambda>X. f -` X \<inter> space F) ` sets E"
+    and "\<And>A. A \<in> sets E \<Longrightarrow> f -` A \<inter> space F \<in> sets F"
+  shows "sigma_algebra.vimage_algebra (sigma E) (space F) f = sigma F"
+proof -
+  interpret sigma_algebra "sigma E"
+    using assms by (intro sigma_algebra_sigma) auto
+  have eq: "sets F = (\<lambda>X. f -` X \<inter> space F) ` sets E"
+    using assms by auto
+  show "vimage_algebra (space F) f = sigma F"
+    unfolding vimage_algebra_def using assms
+    by (simp add: sigma_def eq sigma_sets_vimage)
 qed
 
-lemma measure_eq:
-  assumes fin: "\<mu> (space M) = \<nu> (space M)" "\<nu> (space M) < \<omega>"
-  assumes E: "M = sigma (space E) (sets E)" "Int_stable E"
-  assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e"
-  assumes ms: "measure_space M \<mu>" "measure_space M \<nu>"
-  assumes A: "A \<in> sets M"
-  shows "\<mu> A = \<nu> A"
-proof -
-  interpret M: measure_space M \<mu>
-    using ms by simp
-  interpret M': measure_space M \<nu>
-    using ms by simp
+lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)"
+  by auto
+
+lemma Pair_vimage_times[simp]: "\<And>A B x. Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
+  by auto
+
+lemma rev_Pair_vimage_times[simp]: "\<And>A B y. (\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
+  by auto
+
+lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))"
+  by (cases x) simp
+
+abbreviation
+  "Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A"
 
-  let ?D_sets = "{A. A \<in> sets M \<and> \<mu> A = \<nu> A}"
-  have \<delta>: "dynkin \<lparr> space = space M , sets = ?D_sets \<rparr>"
-  proof (rule dynkinI, safe, simp_all)
-    fix A x assume "A \<in> sets M \<and> \<mu> A = \<nu> A" "x \<in> A"
-    thus "x \<in> space M" using assms M.sets_into_space by auto
-  next
-    show "\<mu> (space M) = \<nu> (space M)"
-      using fin by auto
-  next
-    fix a b
-    assume asm: "a \<in> sets M \<and> \<mu> a = \<nu> a"
-      "b \<in> sets M \<and> \<mu> b = \<nu> b" "a \<subseteq> b"
-    hence "a \<subseteq> space M"
-      using M.sets_into_space by auto
-    from M.measure_mono[OF this]
-    have "\<mu> a \<le> \<mu> (space M)"
-      using asm by auto
-    hence afin: "\<mu> a < \<omega>"
-      using fin by auto
-    have *: "b = b - a \<union> a" using asm by auto
-    have **: "(b - a) \<inter> a = {}" using asm by auto
-    have iv: "\<mu> (b - a) + \<mu> a = \<mu> b"
-      using M.measure_additive[of "b - a" a]
-        conjunct1[OF asm(1)] conjunct1[OF asm(2)] * **
-      by auto
-    have v: "\<nu> (b - a) + \<nu> a = \<nu> b"
-      using M'.measure_additive[of "b - a" a]
-        conjunct1[OF asm(1)] conjunct1[OF asm(2)] * **
-      by auto
-    from iv v have "\<mu> (b - a) = \<nu> (b - a)" using asm afin
-      pinfreal_add_cancel_right[of "\<mu> (b - a)" "\<nu> a" "\<nu> (b - a)"]
-      by auto
-    thus "b - a \<in> sets M \<and> \<mu> (b - a) = \<nu> (b - a)"
-      using asm by auto
-  next
-    fix a assume "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}"
-      "\<And>i. a i \<in> sets M \<and> \<mu> (a i) = \<nu> (a i)"
-    thus "(\<Union>x. a x) \<in> sets M \<and> \<mu> (\<Union>x. a x) = \<nu> (\<Union>x. a x)"
-      using M.measure_countably_additive
-        M'.measure_countably_additive
-        M.countable_UN
-      apply (auto simp add:disjoint_family_on_def image_def)
-      apply (subst M.measure_countably_additive[symmetric])
-      apply (auto simp add:disjoint_family_on_def)
-      apply (subst M'.measure_countably_additive[symmetric])
-      by (auto simp add:disjoint_family_on_def)
-  qed
-  have *: "sets E \<subseteq> ?D_sets"
-    using eq E sigma_sets.Basic[of _ "sets E"]
-    by (auto simp add:sigma_def)
-  have **: "?D_sets \<subseteq> sets M" by auto
-  have "M = \<lparr> space = space M , sets = ?D_sets \<rparr>"
-    unfolding E(1)
-    apply (rule dynkin_lemma[OF E(2)])
-    using eq E space_sigma \<delta> sigma_sets.Basic
-    by (auto simp add:sigma_def)
-  from subst[OF this, of "\<lambda> M. A \<in> sets M", OF A]
-  show ?thesis by auto
-qed
-(*
-lemma
-  assumes sfin: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And> i :: nat. \<nu> (A i) < \<omega>"
-  assumes A: "\<And>  i. \<mu> (A i) = \<nu> (A i)" "\<And> i. A i \<subseteq> A (Suc i)"
-  assumes E: "M = sigma (space E) (sets E)" "Int_stable E"
-  assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e"
-  assumes ms: "measure_space (M :: 'a algebra) \<mu>" "measure_space M \<nu>"
-  assumes B: "B \<in> sets M"
-  shows "\<mu> B = \<nu> B"
-proof -
-  interpret M: measure_space M \<mu> by (rule ms)
-  interpret M': measure_space M \<nu> by (rule ms)
-  have *: "M = \<lparr> space = space M, sets = sets M \<rparr>" by auto
-  { fix i :: nat
-    have **: "M\<lparr> space := A i, sets := op \<inter> (A i) ` sets M \<rparr> =
-      \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>"
-      by auto
-    have mu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<mu>"
-      using M.restricted_measure_space[of "A i", simplified **]
-        sfin by auto
-    have nu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<nu>"
-      using M'.restricted_measure_space[of "A i", simplified **]
-        sfin by auto
-    let ?M = "\<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>"
-    have "\<mu> (A i \<inter> B) = \<nu> (A i \<inter> B)"
-      apply (rule measure_eq[of \<mu> ?M \<nu> "\<lparr> space = space E \<inter> A i, sets = op \<inter> (A i) ` sets E\<rparr>" "A i \<inter> B", simplified])
-      using assms nu_i mu_i
-      apply (auto simp add:image_def) (* TODO *) sorry
-    show ?thesis sorry
-qed
-*)
-definition prod_sets where
-  "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}"
+abbreviation
+  funcset_extensional :: "['a set, 'b set] => ('a => 'b) set"
+    (infixr "->\<^isub>E" 60) where
+  "A ->\<^isub>E B \<equiv> Pi\<^isub>E A (%_. B)"
+
+notation (xsymbols)
+  funcset_extensional  (infixr "\<rightarrow>\<^isub>E" 60)
+
+lemma extensional_empty[simp]: "extensional {} = {\<lambda>x. undefined}"
+  by safe (auto simp add: extensional_def fun_eq_iff)
+
+lemma extensional_insert[intro, simp]:
+  assumes "a \<in> extensional (insert i I)"
+  shows "a(i := b) \<in> extensional (insert i I)"
+  using assms unfolding extensional_def by auto
+
+lemma extensional_Int[simp]:
+  "extensional I \<inter> extensional I' = extensional (I \<inter> I')"
+  unfolding extensional_def by auto
 
 definition
-  "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))"
+  "merge I x J y = (\<lambda>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"
+  unfolding merge_def by auto
+
+lemma merge_commute:
+  "I \<inter> J = {} \<Longrightarrow> merge I x J y = merge J y I 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"
+  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"
+  by (auto simp: Pi_def)
+
+lemma extensional_merge[simp]: "merge I x J 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"
+  by (auto simp: restrict_def intro!: ext)
+
+lemma extensional_insert_undefined[intro, simp]:
+  assumes "a \<in> extensional (insert i I)"
+  shows "a(i := undefined) \<in> extensional I"
+  using assms unfolding extensional_def by auto
+
+lemma extensional_insert_cancel[intro, simp]:
+  assumes "a \<in> extensional I"
+  shows "a \<in> extensional (insert i I)"
+  using assms unfolding extensional_def by auto
+
+lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)"
+  by auto
+
+lemma Pi_cancel_fupd_range[simp]: "i \<notin> I \<Longrightarrow> x \<in> Pi I (B(i := b)) \<longleftrightarrow> x \<in> Pi I B"
+  by (auto simp: Pi_def)
+
+lemma Pi_split_insert_domain[simp]: "x \<in> Pi (insert i I) X \<longleftrightarrow> x \<in> Pi I X \<and> x i \<in> X i"
+  by (auto simp: Pi_def)
 
-lemma
-  fixes M1 :: "'a algebra" and M2 :: "'b algebra"
-  assumes "algebra M1" "algebra M2"
-  shows measureable_fst[intro!, simp]:
-    "fst \<in> measurable (prod_measure_space M1 M2) M1" (is ?fst)
-  and measureable_snd[intro!, simp]:
-    "snd \<in> measurable (prod_measure_space M1 M2) M2" (is ?snd)
+lemma Pi_split_domain[simp]: "x \<in> Pi (I \<union> J) X \<longleftrightarrow> x \<in> Pi I X \<and> x \<in> Pi J X"
+  by (auto simp: Pi_def)
+
+lemma Pi_cancel_fupd[simp]: "i \<notin> I \<Longrightarrow> x(i := a) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B"
+  by (auto simp: Pi_def)
+
+section "Binary products"
+
+definition
+  "pair_algebra A B = \<lparr> space = space A \<times> space B,
+                           sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<rparr>"
+
+locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2
+  for M1 M2
+
+abbreviation (in pair_sigma_algebra)
+  "E \<equiv> pair_algebra M1 M2"
+
+abbreviation (in pair_sigma_algebra)
+  "P \<equiv> sigma E"
+
+sublocale pair_sigma_algebra \<subseteq> sigma_algebra P
+  using M1.sets_into_space M2.sets_into_space
+  by (force simp: pair_algebra_def intro!: sigma_algebra_sigma)
+
+lemma pair_algebraI[intro, simp]:
+  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_algebra A B)"
+  by (auto simp add: pair_algebra_def)
+
+lemma space_pair_algebra:
+  "space (pair_algebra A B) = space A \<times> space B"
+  by (simp add: pair_algebra_def)
+
+lemma pair_algebra_Int_snd:
+  assumes "sets S1 \<subseteq> Pow (space S1)"
+  shows "pair_algebra S1 (algebra.restricted_space S2 A) =
+         algebra.restricted_space (pair_algebra S1 S2) (space S1 \<times> A)"
+  (is "?L = ?R")
+proof (intro algebra.equality set_eqI iffI)
+  fix X assume "X \<in> sets ?L"
+  then obtain A1 A2 where X: "X = A1 \<times> (A \<inter> A2)" and "A1 \<in> sets S1" "A2 \<in> sets S2"
+    by (auto simp: pair_algebra_def)
+  then show "X \<in> sets ?R" unfolding pair_algebra_def
+    using assms apply simp by (intro image_eqI[of _ _ "A1 \<times> A2"]) auto
+next
+  fix X assume "X \<in> sets ?R"
+  then obtain A1 A2 where "X = space S1 \<times> A \<inter> A1 \<times> A2" "A1 \<in> sets S1" "A2 \<in> sets S2"
+    by (auto simp: pair_algebra_def)
+  moreover then have "X = A1 \<times> (A \<inter> A2)"
+    using assms by auto
+  ultimately show "X \<in> sets ?L"
+    unfolding pair_algebra_def by auto
+qed (auto simp add: pair_algebra_def)
+
+lemma (in pair_sigma_algebra)
+  shows measurable_fst[intro!, simp]:
+    "fst \<in> measurable P M1" (is ?fst)
+  and measurable_snd[intro!, simp]:
+    "snd \<in> measurable P M2" (is ?snd)
 proof -
-  interpret M1: algebra M1 by fact
-  interpret M2: algebra M2 by fact
-
   { fix X assume "X \<in> sets M1"
     then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
       apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"])
@@ -372,186 +189,1449 @@
     then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
       apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X])
       using M2.sets_into_space by force+ }
-  ultimately show ?fst ?snd
-    by (force intro!: sigma_sets.Basic
-              simp: measurable_def prod_measure_space_def prod_sets_def sets_sigma)+
+  ultimately have "?fst \<and> ?snd"
+    by (fastsimp simp: measurable_def sets_sigma space_pair_algebra
+                 intro!: sigma_sets.Basic)
+  then show ?fst ?snd by auto
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair:
+  assumes "sigma_algebra M"
+  shows "f \<in> measurable M P \<longleftrightarrow>
+    (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
+proof -
+  interpret M: sigma_algebra M by fact
+  from assms show ?thesis
+  proof (safe intro!: measurable_comp[where b=P])
+    assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
+    show "f \<in> measurable M P"
+    proof (rule M.measurable_sigma)
+      show "sets (pair_algebra M1 M2) \<subseteq> Pow (space E)"
+        unfolding pair_algebra_def using M1.sets_into_space M2.sets_into_space by auto
+      show "f \<in> space M \<rightarrow> space E"
+        using f s by (auto simp: mem_Times_iff measurable_def comp_def space_sigma space_pair_algebra)
+      fix A assume "A \<in> sets E"
+      then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
+        unfolding pair_algebra_def by auto
+      moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
+        using f `B \<in> sets M1` unfolding measurable_def by auto
+      moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
+        using s `C \<in> sets M2` unfolding measurable_def by auto
+      moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)"
+        unfolding `A = B \<times> C` by (auto simp: vimage_Times)
+      ultimately show "f -` A \<inter> space M \<in> sets M" by auto
+    qed
+  qed
+qed
+
+lemma (in pair_sigma_algebra) measurable_prod_sigma:
+  assumes "sigma_algebra M"
+  assumes 1: "(fst \<circ> f) \<in> measurable M M1" and 2: "(snd \<circ> f) \<in> measurable M M2"
+  shows "f \<in> measurable M P"
+proof -
+  interpret M: sigma_algebra M by fact
+  from 1 have fn1: "fst \<circ> f \<in> space M \<rightarrow> space M1"
+     and q1: "\<forall>y\<in>sets M1. (fst \<circ> f) -` y \<inter> space M \<in> sets M"
+    by (auto simp add: measurable_def)
+  from 2 have fn2: "snd \<circ> f \<in> space M \<rightarrow> space M2"
+     and q2: "\<forall>y\<in>sets M2. (snd \<circ> f) -` y \<inter> space M \<in> sets M"
+    by (auto simp add: measurable_def)
+  show ?thesis
+  proof (rule M.measurable_sigma)
+    show "sets E \<subseteq> Pow (space E)"
+      using M1.space_closed M2.space_closed
+      by (auto simp add: sigma_algebra_iff pair_algebra_def)
+  next
+    show "f \<in> space M \<rightarrow> space E"
+      by (simp add: space_pair_algebra) (rule prod_final [OF fn1 fn2])
+  next
+    fix z
+    assume z: "z \<in> sets E"
+    thus "f -` z \<inter> space M \<in> sets M"
+    proof (auto simp add: pair_algebra_def vimage_Times)
+      fix x y
+      assume x: "x \<in> sets M1" and y: "y \<in> sets M2"
+      have "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M =
+            ((fst \<circ> f) -` x \<inter> space M) \<inter> ((snd \<circ> f) -` y \<inter> space M)"
+        by blast
+      also have "...  \<in> sets M" using x y q1 q2
+        by blast
+      finally show "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M \<in> sets M" .
+    qed
+  qed
+qed
+
+lemma pair_algebraE:
+  assumes "X \<in> sets (pair_algebra M1 M2)"
+  obtains A B where "X = A \<times> B" "A \<in> sets M1" "B \<in> sets M2"
+  using assms unfolding pair_algebra_def by auto
+
+lemma (in pair_sigma_algebra) pair_algebra_swap:
+  "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_algebra M2 M1)"
+proof (safe elim!: pair_algebraE)
+  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
+  moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 = B \<times> A"
+    using M1.sets_into_space M2.sets_into_space by auto
+  ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space M2 \<times> space M1 \<in> sets (pair_algebra M2 M1)"
+    by (auto intro: pair_algebraI)
+next
+  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
+  then show "B \<times> A \<in> (\<lambda>X. (\<lambda>(x, y). (y, x)) -` X \<inter> space M2 \<times> space M1) ` sets E"
+    using M1.sets_into_space M2.sets_into_space
+    by (auto intro!: image_eqI[where x="A \<times> B"] pair_algebraI)
+qed
+
+lemma (in pair_sigma_algebra) sets_pair_sigma_algebra_swap:
+  assumes Q: "Q \<in> sets P"
+  shows "(\<lambda>(x,y). (y, x)) ` Q \<in> sets (sigma (pair_algebra M2 M1))" (is "_ \<in> sets ?Q")
+proof -
+  have *: "(\<lambda>(x,y). (y, x)) \<in> space M2 \<times> space M1 \<rightarrow> (space M1 \<times> space M2)"
+       "sets (pair_algebra M1 M2) \<subseteq> Pow (space M1 \<times> space M2)"
+    using M1.sets_into_space M2.sets_into_space by (auto elim!: pair_algebraE)
+  from Q sets_into_space show ?thesis
+    by (auto intro!: image_eqI[where x=Q]
+             simp: pair_algebra_swap[symmetric] sets_sigma
+                   sigma_sets_vimage[OF *] space_pair_algebra)
+qed
+
+lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable:
+  shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (sigma (pair_algebra M2 M1))"
+    (is "?f \<in> measurable ?P ?Q")
+  unfolding measurable_def
+proof (intro CollectI conjI Pi_I ballI)
+  fix x assume "x \<in> space ?P" then show "(case x of (x, y) \<Rightarrow> (y, x)) \<in> space ?Q"
+    unfolding pair_algebra_def by auto
+next
+  fix A assume "A \<in> sets ?Q"
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  have "?f -` A \<inter> space ?P = (\<lambda>(x,y). (y, x)) ` A"
+    using Q.sets_into_space `A \<in> sets ?Q` by (auto simp: pair_algebra_def)
+  with Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets ?Q`]
+  show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
+qed
+
+lemma (in pair_sigma_algebra) measurable_cut_fst:
+  assumes "Q \<in> sets P" shows "Pair x -` Q \<in> sets M2"
+proof -
+  let ?Q' = "{Q. Q \<subseteq> space P \<and> Pair x -` Q \<in> sets M2}"
+  let ?Q = "\<lparr> space = space P, sets = ?Q' \<rparr>"
+  interpret Q: sigma_algebra ?Q
+    proof qed (auto simp: vimage_UN vimage_Diff space_pair_algebra)
+  have "sets E \<subseteq> sets ?Q"
+    using M1.sets_into_space M2.sets_into_space
+    by (auto simp: pair_algebra_def space_pair_algebra)
+  then have "sets P \<subseteq> sets ?Q"
+    by (subst pair_algebra_def, intro Q.sets_sigma_subset)
+       (simp_all add: pair_algebra_def)
+  with assms show ?thesis by auto
+qed
+
+lemma (in pair_sigma_algebra) measurable_cut_snd:
+  assumes Q: "Q \<in> sets P" shows "(\<lambda>x. (x, y)) -` Q \<in> sets M1" (is "?cut Q \<in> sets M1")
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  have "Pair y -` (\<lambda>(x, y). (y, x)) ` Q = (\<lambda>x. (x, y)) -` Q" by auto
+  with Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y]
+  show ?thesis by simp
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair_image_snd:
+  assumes m: "f \<in> measurable P M" and "x \<in> space M1"
+  shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
+  unfolding measurable_def
+proof (intro CollectI conjI Pi_I ballI)
+  fix y assume "y \<in> space M2" with `f \<in> measurable P M` `x \<in> space M1`
+  show "f (x, y) \<in> space M" unfolding measurable_def pair_algebra_def by auto
+next
+  fix A assume "A \<in> sets M"
+  then have "Pair x -` (f -` A \<inter> space P) \<in> sets M2" (is "?C \<in> _")
+    using `f \<in> measurable P M`
+    by (intro measurable_cut_fst) (auto simp: measurable_def)
+  also have "?C = (\<lambda>y. f (x, y)) -` A \<inter> space M2"
+    using `x \<in> space M1` by (auto simp: pair_algebra_def)
+  finally show "(\<lambda>y. f (x, y)) -` A \<inter> space M2 \<in> sets M2" .
+qed
+
+lemma (in pair_sigma_algebra) measurable_pair_image_fst:
+  assumes m: "f \<in> measurable P M" and "y \<in> space M2"
+  shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  from Q.measurable_pair_image_snd[OF measurable_comp `y \<in> space M2`,
+                                      OF Q.pair_sigma_algebra_swap_measurable m]
+  show ?thesis by simp
+qed
+
+lemma (in pair_sigma_algebra) Int_stable_pair_algebra: "Int_stable E"
+  unfolding Int_stable_def
+proof (intro ballI)
+  fix A B assume "A \<in> sets E" "B \<in> sets E"
+  then obtain A1 A2 B1 B2 where "A = A1 \<times> A2" "B = B1 \<times> B2"
+    "A1 \<in> sets M1" "A2 \<in> sets M2" "B1 \<in> sets M1" "B2 \<in> sets M2"
+    unfolding pair_algebra_def by auto
+  then show "A \<inter> B \<in> sets E"
+    by (auto simp add: times_Int_times pair_algebra_def)
+qed
+
+lemma finite_measure_cut_measurable:
+  fixes M1 :: "'a algebra" and M2 :: "'b algebra"
+  assumes "sigma_finite_measure M1 \<mu>1" "finite_measure M2 \<mu>2"
+  assumes "Q \<in> sets (sigma (pair_algebra M1 M2))"
+  shows "(\<lambda>x. \<mu>2 (Pair x -` Q)) \<in> borel_measurable M1"
+    (is "?s Q \<in> _")
+proof -
+  interpret M1: sigma_finite_measure M1 \<mu>1 by fact
+  interpret M2: finite_measure M2 \<mu>2 by fact
+  interpret pair_sigma_algebra M1 M2 by default
+  have [intro]: "sigma_algebra M1" by fact
+  have [intro]: "sigma_algebra M2" by fact
+  let ?D = "\<lparr> space = space P, sets = {A\<in>sets P. ?s A \<in> borel_measurable M1}  \<rparr>"
+  note space_pair_algebra[simp]
+  interpret dynkin_system ?D
+  proof (intro dynkin_systemI)
+    fix A assume "A \<in> sets ?D" then show "A \<subseteq> space ?D"
+      using sets_into_space by simp
+  next
+    from top show "space ?D \<in> sets ?D"
+      by (auto simp add: if_distrib intro!: M1.measurable_If)
+  next
+    fix A assume "A \<in> sets ?D"
+    with sets_into_space have "\<And>x. \<mu>2 (Pair x -` (space M1 \<times> space M2 - A)) =
+        (if x \<in> space M1 then \<mu>2 (space M2) - ?s A x else 0)"
+      by (auto intro!: M2.finite_measure_compl measurable_cut_fst
+               simp: vimage_Diff)
+    with `A \<in> sets ?D` top show "space ?D - A \<in> sets ?D"
+      by (auto intro!: Diff M1.measurable_If M1.borel_measurable_pinfreal_diff)
+  next
+    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
+    moreover then have "\<And>x. \<mu>2 (\<Union>i. Pair x -` F i) = (\<Sum>\<^isub>\<infinity> i. ?s (F i) x)"
+      by (intro M2.measure_countably_additive[symmetric])
+         (auto intro!: measurable_cut_fst simp: disjoint_family_on_def)
+    ultimately show "(\<Union>i. F i) \<in> sets ?D"
+      by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
+  qed
+  have "P = ?D"
+  proof (intro dynkin_lemma)
+    show "Int_stable E" by (rule Int_stable_pair_algebra)
+    from M1.sets_into_space have "\<And>A. A \<in> sets M1 \<Longrightarrow> {x \<in> space M1. x \<in> A} = A"
+      by auto
+    then show "sets E \<subseteq> sets ?D"
+      by (auto simp: pair_algebra_def sets_sigma if_distrib
+               intro: sigma_sets.Basic intro!: M1.measurable_If)
+  qed auto
+  with `Q \<in> sets P` have "Q \<in> sets ?D" by simp
+  then show "?s Q \<in> borel_measurable M1" by simp
+qed
+
+subsection {* Binary products of $\sigma$-finite measure spaces *}
+
+locale pair_sigma_finite = M1: sigma_finite_measure M1 \<mu>1 + M2: sigma_finite_measure M2 \<mu>2
+  for M1 \<mu>1 M2 \<mu>2
+
+sublocale pair_sigma_finite \<subseteq> pair_sigma_algebra M1 M2
+  by default
+
+lemma (in pair_sigma_finite) measure_cut_measurable_fst:
+  assumes "Q \<in> sets P" shows "(\<lambda>x. \<mu>2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
+proof -
+  have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
+  have M1: "sigma_finite_measure M1 \<mu>1" by default
+
+  from M2.disjoint_sigma_finite guess F .. note F = this
+  let "?C x i" = "F i \<inter> Pair x -` Q"
+  { fix i
+    let ?R = "M2.restricted_space (F i)"
+    have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i"
+      using F M2.sets_into_space by auto
+    have "(\<lambda>x. \<mu>2 (Pair x -` (space M1 \<times> F i \<inter> Q))) \<in> borel_measurable M1"
+    proof (intro finite_measure_cut_measurable[OF M1])
+      show "finite_measure (M2.restricted_space (F i)) \<mu>2"
+        using F by (intro M2.restricted_to_finite_measure) auto
+      have "space M1 \<times> F i \<in> sets P"
+        using M1.top F by blast
+      from sigma_sets_Int[symmetric,
+        OF this[unfolded pair_sigma_algebra_def sets_sigma]]
+      show "(space M1 \<times> F i) \<inter> Q \<in> sets (sigma (pair_algebra M1 ?R))"
+        using `Q \<in> sets P`
+        using pair_algebra_Int_snd[OF M1.space_closed, of "F i" M2]
+        by (auto simp: pair_algebra_def sets_sigma)
+    qed
+    moreover have "\<And>x. Pair x -` (space M1 \<times> F i \<inter> Q) = ?C x i"
+      using `Q \<in> sets P` sets_into_space by (auto simp: space_pair_algebra)
+    ultimately have "(\<lambda>x. \<mu>2 (?C x i)) \<in> borel_measurable M1"
+      by simp }
+  moreover
+  { fix x
+    have "(\<Sum>\<^isub>\<infinity>i. \<mu>2 (?C x i)) = \<mu>2 (\<Union>i. ?C x i)"
+    proof (intro M2.measure_countably_additive)
+      show "range (?C x) \<subseteq> sets M2"
+        using F `Q \<in> sets P` by (auto intro!: M2.Int measurable_cut_fst)
+      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 `Q \<in> sets P`
+      by (auto simp: space_pair_algebra)
+    finally have "\<mu>2 (Pair x -` Q) = (\<Sum>\<^isub>\<infinity>i. \<mu>2 (?C x i))"
+      by simp }
+  ultimately show ?thesis
+    by (auto intro!: M1.borel_measurable_psuminf)
+qed
+
+lemma (in pair_sigma_finite) measure_cut_measurable_snd:
+  assumes "Q \<in> sets P" shows "(\<lambda>y. \<mu>1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
+proof -
+  interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
+  have [simp]: "\<And>y. (Pair y -` (\<lambda>(x, y). (y, x)) ` Q) = (\<lambda>x. (x, y)) -` Q"
+    by auto
+  note sets_pair_sigma_algebra_swap[OF assms]
+  from Q.measure_cut_measurable_fst[OF this]
+  show ?thesis by simp
+qed
+
+lemma (in pair_sigma_algebra) pair_sigma_algebra_measurable:
+  assumes "f \<in> measurable P M"
+  shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (sigma (pair_algebra M2 M1)) M"
+proof -
+  interpret Q: pair_sigma_algebra M2 M1 by default
+  have *: "(\<lambda>(x,y). f (y, x)) = f \<circ> (\<lambda>(x,y). (y, x))" by (simp add: fun_eq_iff)
+  show ?thesis
+    using Q.pair_sigma_algebra_swap_measurable assms
+    unfolding * by (rule measurable_comp)
 qed
 
-lemma (in sigma_algebra) measureable_prod:
-  fixes M1 :: "'a algebra" and M2 :: "'b algebra"
-  assumes "algebra M1" "algebra M2"
-  shows "f \<in> measurable M (prod_measure_space M1 M2) \<longleftrightarrow>
-    (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
-using assms proof (safe intro!: measurable_comp[where b="prod_measure_space M1 M2"])
-  interpret M1: algebra M1 by fact
-  interpret M2: algebra M2 by fact
-  assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
+lemma (in pair_sigma_algebra) pair_sigma_algebra_swap:
+  "sigma (pair_algebra M2 M1) = (vimage_algebra (space M2 \<times> space M1) (\<lambda>(x, y). (y, x)))"
+  unfolding vimage_algebra_def
+  apply (simp add: sets_sigma)
+  apply (subst sigma_sets_vimage[symmetric])
+  apply (fastsimp simp: pair_algebra_def)
+  using M1.sets_into_space M2.sets_into_space apply (fastsimp simp: pair_algebra_def)
+proof -
+  have "(\<lambda>X. (\<lambda>(x, y). (y, x)) -` X \<inter> space M2 \<times> space M1) ` sets E
+    = sets (pair_algebra M2 M1)" (is "?S = _")
+    by (rule pair_algebra_swap)
+  then show "sigma (pair_algebra M2 M1) = \<lparr>space = space M2 \<times> space M1,
+       sets = sigma_sets (space M2 \<times> space M1) ?S\<rparr>"
+    by (simp add: pair_algebra_def sigma_def)
+qed
+
+definition (in pair_sigma_finite)
+  "pair_measure A = M1.positive_integral (\<lambda>x.
+    M2.positive_integral (\<lambda>y. indicator A (x, y)))"
+
+lemma (in pair_sigma_finite) pair_measure_alt:
+  assumes "A \<in> sets P"
+  shows "pair_measure A = M1.positive_integral (\<lambda>x. \<mu>2 (Pair x -` A))"
+  unfolding pair_measure_def
+proof (rule M1.positive_integral_cong)
+  fix x assume "x \<in> space M1"
+  have *: "\<And>y. indicator A (x, y) = (indicator (Pair x -` A) y :: pinfreal)"
+    unfolding indicator_def by auto
+  show "M2.positive_integral (\<lambda>y. indicator A (x, y)) = \<mu>2 (Pair x -` A)"
+    unfolding *
+    apply (subst M2.positive_integral_indicator)
+    apply (rule measurable_cut_fst[OF assms])
+    by simp
+qed
+
+lemma (in pair_sigma_finite) pair_measure_times:
+  assumes A: "A \<in> sets M1" and "B \<in> sets M2"
+  shows "pair_measure (A \<times> B) = \<mu>1 A * \<mu>2 B"
+proof -
+  from assms have "pair_measure (A \<times> B) =
+      M1.positive_integral (\<lambda>x. \<mu>2 B * indicator A x)"
+    by (auto intro!: M1.positive_integral_cong simp: pair_measure_alt)
+  with assms show ?thesis
+    by (simp add: M1.positive_integral_cmult_indicator ac_simps)
+qed
+
+lemma (in pair_sigma_finite) sigma_finite_up_in_pair_algebra:
+  "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> F \<up> space E \<and>
+    (\<forall>i. pair_measure (F i) \<noteq> \<omega>)"
+proof -
+  obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
+    F1: "range F1 \<subseteq> sets M1" "F1 \<up> space M1" "\<And>i. \<mu>1 (F1 i) \<noteq> \<omega>" and
+    F2: "range F2 \<subseteq> sets M2" "F2 \<up> space M2" "\<And>i. \<mu>2 (F2 i) \<noteq> \<omega>"
+    using M1.sigma_finite_up M2.sigma_finite_up by auto
+  then have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)"
+    unfolding isoton_def by auto
+  let ?F = "\<lambda>i. F1 i \<times> F2 i"
+  show ?thesis unfolding isoton_def space_pair_algebra
+  proof (intro exI[of _ ?F] conjI allI)
+    show "range ?F \<subseteq> sets E" using F1 F2
+      by (fastsimp intro!: pair_algebraI)
+  next
+    have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
+    proof (intro subsetI)
+      fix x assume "x \<in> space M1 \<times> space M2"
+      then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"
+        by (auto simp: space)
+      then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"
+        using `F1 \<up> space M1` `F2 \<up> space M2`
+        by (auto simp: max_def dest: isoton_mono_le)
+      then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"
+        by (intro SigmaI) (auto simp add: min_max.sup_commute)
+      then show "x \<in> (\<Union>i. ?F i)" by auto
+    qed
+    then show "(\<Union>i. ?F i) = space M1 \<times> space M2"
+      using space by (auto simp: space)
+  next
+    fix i show "F1 i \<times> F2 i \<subseteq> F1 (Suc i) \<times> F2 (Suc i)"
+      using `F1 \<up> space M1` `F2 \<up> space M2` unfolding isoton_def
+      by auto
+  next
+    fix i
+    from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
+    with F1 F2 show "pair_measure (F1 i \<times> F2 i) \<noteq> \<omega>"
+      by (simp add: pair_measure_times)
+  qed
+qed
+
+sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P pair_measure
+proof
+  show "pair_measure {} = 0"
+    unfolding pair_measure_def by auto
+
+  show "countably_additive P pair_measure"
+    unfolding countably_additive_def
+  proof (intro allI impI)
+    fix F :: "nat \<Rightarrow> ('a \<times> 'b) set"
+    assume F: "range F \<subseteq> sets P" "disjoint_family F"
+    from F have *: "\<And>i. F i \<in> sets P" "(\<Union>i. F i) \<in> sets P" by auto
+    moreover from F have "\<And>i. (\<lambda>x. \<mu>2 (Pair x -` F i)) \<in> borel_measurable M1"
+      by (intro measure_cut_measurable_fst) 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. x \<in> space M1 \<Longrightarrow> range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
+      using F by (auto intro!: measurable_cut_fst)
+    ultimately show "(\<Sum>\<^isub>\<infinity>n. pair_measure (F n)) = pair_measure (\<Union>i. F i)"
+      by (simp add: pair_measure_alt vimage_UN M1.positive_integral_psuminf[symmetric]
+                    M2.measure_countably_additive
+               cong: M1.positive_integral_cong)
+  qed
+
+  from sigma_finite_up_in_pair_algebra guess F :: "nat \<Rightarrow> ('a \<times> 'c) set" .. note F = this
+  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. pair_measure (F i) \<noteq> \<omega>)"
+  proof (rule exI[of _ F], intro conjI)
+    show "range F \<subseteq> sets P" using F by auto
+    show "(\<Union>i. F i) = space P"
+      using F by (auto simp: space_pair_algebra isoton_def)
+    show "\<forall>i. pair_measure (F i) \<noteq> \<omega>" using F by auto
+  qed
+qed
 
-  show "f \<in> measurable M (prod_measure_space M1 M2)" unfolding prod_measure_space_def
-  proof (rule measurable_sigma)
-    show "prod_sets (sets M1) (sets M2) \<subseteq> Pow (space M1 \<times> space M2)"
-      unfolding prod_sets_def using M1.sets_into_space M2.sets_into_space by auto
-    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 assume "A \<in> prod_sets (sets M1) (sets M2)"
-    then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
-      unfolding prod_sets_def by auto
-    moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
-      using f `B \<in> sets M1` unfolding measurable_def by auto
-    moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
-      using s `C \<in> sets M2` unfolding measurable_def by auto
-    moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)"
-      unfolding `A = B \<times> C` by (auto simp: vimage_Times)
-    ultimately show "f -` A \<inter> space M \<in> sets M" by auto
+lemma (in pair_sigma_finite) pair_measure_alt2:
+  assumes "A \<in> sets P"
+  shows "pair_measure A = M2.positive_integral (\<lambda>y. \<mu>1 ((\<lambda>x. (x, y)) -` A))"
+    (is "_ = ?\<nu> A")
+proof -
+  from sigma_finite_up_in_pair_algebra guess F :: "nat \<Rightarrow> ('a \<times> 'c) set" .. note F = this
+  show ?thesis
+  proof (rule measure_unique_Int_stable[where \<nu>="?\<nu>", OF Int_stable_pair_algebra],
+         simp_all add: pair_sigma_algebra_def[symmetric])
+    show "range F \<subseteq> sets E" "F \<up> space E" "\<And>i. pair_measure (F i) \<noteq> \<omega>"
+      using F by auto
+    show "measure_space P pair_measure" by default
+  next
+    show "measure_space P ?\<nu>"
+    proof
+      show "?\<nu> {} = 0" by auto
+      show "countably_additive P ?\<nu>"
+        unfolding countably_additive_def
+      proof (intro allI impI)
+        fix F :: "nat \<Rightarrow> ('a \<times> 'b) set"
+        assume F: "range F \<subseteq> sets P" "disjoint_family F"
+        from F have *: "\<And>i. F i \<in> sets P" "(\<Union>i. F i) \<in> sets P" by auto
+        moreover from F have "\<And>i. (\<lambda>y. \<mu>1 ((\<lambda>x. (x, y)) -` F i)) \<in> borel_measurable M2"
+          by (intro measure_cut_measurable_snd) auto
+        moreover have "\<And>y. disjoint_family (\<lambda>i. (\<lambda>x. (x, y)) -` F i)"
+          by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
+        moreover have "\<And>y. y \<in> space M2 \<Longrightarrow> range (\<lambda>i. (\<lambda>x. (x, y)) -` F i) \<subseteq> sets M1"
+          using F by (auto intro!: measurable_cut_snd)
+        ultimately show "(\<Sum>\<^isub>\<infinity>n. ?\<nu> (F n)) = ?\<nu> (\<Union>i. F i)"
+          by (simp add: vimage_UN M2.positive_integral_psuminf[symmetric]
+                        M1.measure_countably_additive
+                   cong: M2.positive_integral_cong)
+      qed
+    qed
+  next
+    fix X assume "X \<in> sets E"
+    then obtain A B where X: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
+      unfolding pair_algebra_def by auto
+    show "pair_measure X = ?\<nu> X"
+    proof -
+      from AB have "?\<nu> (A \<times> B) =
+          M2.positive_integral (\<lambda>y. \<mu>1 A * indicator B y)"
+        by (auto intro!: M2.positive_integral_cong)
+      with AB show ?thesis
+        unfolding pair_measure_times[OF AB] X
+        by (simp add: M2.positive_integral_cmult_indicator ac_simps)
+    qed
+  qed fact
+qed
+
+section "Fubinis theorem"
+
+lemma (in pair_sigma_finite) simple_function_cut:
+  assumes f: "simple_function f"
+  shows "(\<lambda>x. M2.positive_integral (\<lambda> y. f (x, y))) \<in> borel_measurable M1"
+    and "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))
+      = positive_integral f"
+proof -
+  have f_borel: "f \<in> borel_measurable P"
+    using f by (rule borel_measurable_simple_function)
+  let "?F z" = "f -` {z} \<inter> space P"
+  let "?F' 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 P" using `x \<in> space M1`
+      by (simp add: space_pair_algebra)
+    moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
+      by (intro borel_measurable_vimage measurable_cut_fst)
+    ultimately have "M2.simple_function (\<lambda> y. f (x, y))"
+      apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
+      apply (rule simple_function_indicator_representation[OF f])
+      using `x \<in> space M1` by (auto simp del: space_sigma) }
+  note M2_sf = this
+  { fix x assume x: "x \<in> space M1"
+    then have "M2.positive_integral (\<lambda> y. f (x, y)) =
+        (\<Sum>z\<in>f ` space P. z * \<mu>2 (?F' x z))"
+      unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x]]
+      unfolding M2.simple_integral_def
+    proof (safe intro!: setsum_mono_zero_cong_left)
+      from f show "finite (f ` space P)" by (rule simple_functionD)
+    next
+      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
+        using `x \<in> space M1` by (auto simp: space_pair_algebra)
+    next
+      fix x' y assume "(x', y) \<in> space P"
+        "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
+      then have *: "?F' x (f (x', y)) = {}"
+        by (force simp: space_pair_algebra)
+      show  "f (x', y) * \<mu>2 (?F' x (f (x', y))) = 0"
+        unfolding * by simp
+    qed (simp add: vimage_compose[symmetric] comp_def
+                   space_pair_algebra) }
+  note eq = this
+  moreover have "\<And>z. ?F z \<in> sets P"
+    by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
+  moreover then have "\<And>z. (\<lambda>x. \<mu>2 (?F' x z)) \<in> borel_measurable M1"
+    by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
+  ultimately
+  show "(\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) \<in> borel_measurable M1"
+    and "M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))
+    = positive_integral f"
+    by (auto simp del: vimage_Int cong: measurable_cong
+             intro!: M1.borel_measurable_pinfreal_setsum
+             simp add: M1.positive_integral_setsum simple_integral_def
+                       M1.positive_integral_cmult
+                       M1.positive_integral_cong[OF eq]
+                       positive_integral_eq_simple_integral[OF f]
+                       pair_measure_alt[symmetric])
+qed
+
+lemma (in pair_sigma_finite) positive_integral_fst_measurable:
+  assumes f: "f \<in> borel_measurable P"
+  shows "(\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) \<in> borel_measurable M1"
+      (is "?C f \<in> borel_measurable M1")
+    and "M1.positive_integral (\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) =
+      positive_integral f"
+proof -
+  from borel_measurable_implies_simple_function_sequence[OF f]
+  obtain F where F: "\<And>i. simple_function (F i)" "F \<up> f" by auto
+  then have F_borel: "\<And>i. F i \<in> borel_measurable P"
+    and F_mono: "\<And>i x. F i x \<le> F (Suc i) x"
+    and F_SUPR: "\<And>x. (SUP i. F i x) = f x"
+    unfolding isoton_def le_fun_def SUPR_fun_expand
+    by (auto intro: borel_measurable_simple_function)
+  note sf = simple_function_cut[OF F(1)]
+  then have "(SUP i. ?C (F i)) \<in> borel_measurable M1"
+    using F(1) by (auto intro!: M1.borel_measurable_SUP)
+  moreover
+  { fix x assume "x \<in> space M1"
+    have isotone: "(\<lambda> i y. F i (x, y)) \<up> (\<lambda>y. f (x, y))"
+      using `F \<up> f` unfolding isoton_fun_expand
+      by (auto simp: isoton_def)
+    note measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
+    from M2.positive_integral_isoton[OF isotone this]
+    have "(SUP i. ?C (F i) x) = ?C f x"
+      by (simp add: isoton_def) }
+  note SUPR_C = this
+  ultimately show "?C f \<in> borel_measurable M1"
+    unfolding SUPR_fun_expand by (simp cong: measurable_cong)
+  have "positive_integral (\<lambda>x. SUP i. F i x) = (SUP i. positive_integral (F i))"
+    using F_borel F_mono
+    by (auto intro!: positive_integral_monotone_convergence_SUP[symmetric])
+  also have "(SUP i. positive_integral (F i)) =
+    (SUP i. M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. F i (x, y))))"
+    unfolding sf(2) by simp
+  also have "\<dots> = M1.positive_integral (\<lambda>x. SUP i. M2.positive_integral (\<lambda>y. F i (x, y)))"
+    by (auto intro!: M1.positive_integral_monotone_convergence_SUP[OF _ sf(1)]
+                     M2.positive_integral_mono F_mono)
+  also have "\<dots> = M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. SUP i. F i (x, y)))"
+    using F_borel F_mono
+    by (auto intro!: M2.positive_integral_monotone_convergence_SUP
+                     M1.positive_integral_cong measurable_pair_image_snd)
+  finally show "M1.positive_integral (\<lambda> x. M2.positive_integral (\<lambda> y. f (x, y))) =
+      positive_integral f"
+    unfolding F_SUPR by simp
+qed
+
+lemma (in pair_sigma_finite) positive_integral_snd_measurable:
+  assumes f: "f \<in> borel_measurable P"
+  shows "M2.positive_integral (\<lambda>y. M1.positive_integral (\<lambda>x. f (x, y))) =
+      positive_integral f"
+proof -
+  interpret Q: pair_sigma_finite M2 \<mu>2 M1 \<mu>1 by default
+
+  have s: "\<And>x y. (case (x, y) of (x, y) \<Rightarrow> f (y, x)) = f (y, x)" by simp
+  have t: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> (y, x))) = (\<lambda>(x, y). f (y, x))" by (auto simp: fun_eq_iff)
+
+  have bij: "bij_betw (\<lambda>(x, y). (y, x)) (space M2 \<times> space M1) (space P)"
+    by (auto intro!: inj_onI simp: space_pair_algebra bij_betw_def)
+
+  note pair_sigma_algebra_measurable[OF f]
+  from Q.positive_integral_fst_measurable[OF this]
+  have "M2.positive_integral (\<lambda>y. M1.positive_integral (\<lambda>x. f (x, y))) =
+    Q.positive_integral (\<lambda>(x, y). f (y, x))"
+    by simp
+  also have "\<dots> = positive_integral f"
+    unfolding positive_integral_vimage[OF bij, of f] t
+    unfolding pair_sigma_algebra_swap[symmetric]
+  proof (rule Q.positive_integral_cong_measure[symmetric])
+    fix A assume "A \<in> sets Q.P"
+    from this Q.sets_pair_sigma_algebra_swap[OF this]
+    show "pair_measure ((\<lambda>(x, y). (y, x)) ` A) = Q.pair_measure A"
+      by (auto intro!: M1.positive_integral_cong arg_cong[where f=\<mu>2]
+               simp: pair_measure_alt Q.pair_measure_alt2)
+  qed
+  finally show ?thesis .
+qed
+
+lemma (in pair_sigma_finite) Fubini:
+  assumes f: "f \<in> borel_measurable P"
+  shows "M2.positive_integral (\<lambda>y. M1.positive_integral (\<lambda>x. f (x, y))) =
+      M1.positive_integral (\<lambda>x. M2.positive_integral (\<lambda>y. f (x, y)))"
+  unfolding positive_integral_snd_measurable[OF assms]
+  unfolding positive_integral_fst_measurable[OF assms] ..
+
+lemma (in pair_sigma_finite) AE_pair:
+  assumes "almost_everywhere (\<lambda>x. Q x)"
+  shows "M1.almost_everywhere (\<lambda>x. M2.almost_everywhere (\<lambda>y. Q (x, y)))"
+proof -
+  obtain N where N: "N \<in> sets P" "pair_measure N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
+    using assms unfolding almost_everywhere_def by auto
+  show ?thesis
+  proof (rule M1.AE_I)
+    from N measure_cut_measurable_fst[OF `N \<in> sets P`]
+    show "\<mu>1 {x\<in>space M1. \<mu>2 (Pair x -` N) \<noteq> 0} = 0"
+      by (simp add: M1.positive_integral_0_iff pair_measure_alt vimage_def)
+    show "{x \<in> space M1. \<mu>2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
+      by (intro M1.borel_measurable_pinfreal_neq_const measure_cut_measurable_fst N)
+    { fix x assume "x \<in> space M1" "\<mu>2 (Pair x -` N) = 0"
+      have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
+      proof (rule M2.AE_I)
+        show "\<mu>2 (Pair x -` N) = 0" by fact
+        show "Pair x -` N \<in> sets M2" by (intro measurable_cut_fst N)
+        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
+          using N `x \<in> space M1` unfolding space_sigma space_pair_algebra by auto
+      qed }
+    then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. \<mu>2 (Pair x -` N) \<noteq> 0}"
+      by auto
   qed
 qed
 
+section "Finite product spaces"
+
+section "Products"
+
+locale product_sigma_algebra =
+  fixes M :: "'i \<Rightarrow> 'a algebra"
+  assumes sigma_algebras: "\<And>i. sigma_algebra (M i)"
+
+locale finite_product_sigma_algebra = product_sigma_algebra M for M :: "'i \<Rightarrow> 'a algebra" +
+  fixes I :: "'i set"
+  assumes finite_index: "finite I"
+
+syntax
+  "_PiE"  :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3PIE _:_./ _)" 10)
+
+syntax (xsymbols)
+  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
+
+syntax (HTML output)
+  "_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set"  ("(3\<Pi>\<^isub>E _\<in>_./ _)"   10)
+
+translations
+  "PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)"
+
 definition
-  "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))"
+  "product_algebra M I = \<lparr> space = (\<Pi>\<^isub>E i\<in>I. space (M i)), sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)) \<rparr>"
+
+abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra M I"
+abbreviation (in finite_product_sigma_algebra) "P \<equiv> sigma G"
+
+sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras)
+
+lemma (in finite_product_sigma_algebra) product_algebra_into_space:
+  "sets G \<subseteq> Pow (space G)"
+  using M.sets_into_space unfolding product_algebra_def
+  by auto blast
+
+sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P
+  using product_algebra_into_space by (rule sigma_algebra_sigma)
+
+lemma space_product_algebra[simp]:
+  "space (product_algebra M I) = Pi\<^isub>E I (\<lambda>i. space (M i))"
+  unfolding product_algebra_def by simp
+
+lemma (in finite_product_sigma_algebra) P_empty:
+  "I = {} \<Longrightarrow> P = \<lparr> space = {\<lambda>k. undefined}, sets = { {}, {\<lambda>k. undefined} }\<rparr>"
+  unfolding product_algebra_def by (simp add: sigma_def)
+
+lemma (in finite_product_sigma_algebra) in_P[simp, intro]:
+  "\<lbrakk> \<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i) \<rbrakk> \<Longrightarrow> Pi\<^isub>E I A \<in> sets P"
+  by (auto simp: product_algebra_def sets_sigma intro!: sigma_sets.Basic)
+
+lemma bij_betw_prod_fold:
+  assumes "i \<notin> I"
+  shows "bij_betw (\<lambda>x. (x(i:=undefined), x i)) (\<Pi>\<^isub>E j\<in>insert i I. space (M j)) ((\<Pi>\<^isub>E j\<in>I. space (M j)) \<times> space (M i))"
+    (is "bij_betw ?f ?P ?F")
+    using `i \<notin> I`
+proof (unfold bij_betw_def, intro conjI set_eqI iffI)
+  show "inj_on ?f ?P"
+  proof (safe intro!: inj_onI ext)
+    fix a b x assume "a(i:=undefined) = b(i:=undefined)" "a i = b i"
+    then show "a x = b x"
+      by (cases "x = i") (auto simp: fun_eq_iff split: split_if_asm)
+  qed
+next
+  fix X assume *: "X \<in> ?F" show "X \<in> ?f ` ?P"
+  proof (cases X)
+    case (Pair a b) with * `i \<notin> I` show ?thesis
+      by (auto intro!: image_eqI[where x="a (i := b)"] ext simp: extensional_def)
+  qed
+qed auto
+
+section "Generating set generates also product algebra"
+
+lemma pair_sigma_algebra_sigma:
+  assumes 1: "S1 \<up> (space E1)" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
+  assumes 2: "S2 \<up> (space E2)" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
+  shows "sigma (pair_algebra (sigma E1) (sigma E2)) = sigma (pair_algebra E1 E2)"
+    (is "?S = ?E")
+proof -
+  interpret M1: sigma_algebra "sigma E1" using E1 by (rule sigma_algebra_sigma)
+  interpret M2: sigma_algebra "sigma E2" using E2 by (rule sigma_algebra_sigma)
+  have P: "sets (pair_algebra E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
+    using E1 E2 by (auto simp add: pair_algebra_def)
+  interpret E: sigma_algebra ?E unfolding pair_algebra_def
+    using E1 E2 by (intro sigma_algebra_sigma) auto
+
+  { fix A assume "A \<in> sets E1"
+    then have "fst -` A \<inter> space ?E = A \<times> (\<Union>i. S2 i)"
+      using E1 2 unfolding isoton_def pair_algebra_def by auto
+    also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
+    also have "\<dots> \<in> sets ?E" unfolding pair_algebra_def sets_sigma
+      using 2 `A \<in> sets E1`
+      by (intro sigma_sets.Union)
+         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+    finally have "fst -` A \<inter> space ?E \<in> sets ?E" . }
+  moreover
+  { fix B assume "B \<in> sets E2"
+    then have "snd -` B \<inter> space ?E = (\<Union>i. S1 i) \<times> B"
+      using E2 1 unfolding isoton_def pair_algebra_def by auto
+    also have "\<dots> = (\<Union>i. S1 i \<times> B)" by auto
+    also have "\<dots> \<in> sets ?E"
+      using 1 `B \<in> sets E2` unfolding pair_algebra_def sets_sigma
+      by (intro sigma_sets.Union)
+         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+    finally have "snd -` B \<inter> space ?E \<in> sets ?E" . }
+  ultimately have proj:
+    "fst \<in> measurable ?E (sigma E1) \<and> snd \<in> measurable ?E (sigma E2)"
+    using E1 E2 by (subst (1 2) E.measurable_iff_sigma)
+                   (auto simp: pair_algebra_def sets_sigma)
+
+  { fix A B assume A: "A \<in> sets (sigma E1)" and B: "B \<in> sets (sigma E2)"
+    with proj have "fst -` A \<inter> space ?E \<in> sets ?E" "snd -` B \<inter> space ?E \<in> sets ?E"
+      unfolding measurable_def by simp_all
+    moreover have "A \<times> B = (fst -` A \<inter> space ?E) \<inter> (snd -` B \<inter> space ?E)"
+      using A B M1.sets_into_space M2.sets_into_space
+      by (auto simp: pair_algebra_def)
+    ultimately have "A \<times> B \<in> sets ?E" by auto }
+  then have "sigma_sets (space ?E) (sets (pair_algebra (sigma E1) (sigma E2))) \<subseteq> sets ?E"
+    by (intro E.sigma_sets_subset) (auto simp add: pair_algebra_def sets_sigma)
+  then have subset: "sets ?S \<subseteq> sets ?E"
+    by (simp add: sets_sigma pair_algebra_def)
+
+  have "sets ?S = sets ?E"
+  proof (intro set_eqI iffI)
+    fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
+      unfolding sets_sigma
+    proof induct
+      case (Basic A) then show ?case
+        by (auto simp: pair_algebra_def sets_sigma intro: sigma_sets.Basic)
+    qed (auto intro: sigma_sets.intros simp: pair_algebra_def)
+  next
+    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
+  qed
+  then show ?thesis
+    by (simp add: pair_algebra_def sigma_def)
+qed
+
+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"
+  apply auto
+  apply (drule_tac x=x in Pi_mem)
+  apply (simp_all split: split_if_asm)
+  apply (drule_tac x=i in Pi_mem)
+  apply (auto dest!: Pi_mem)
+  done
+
+lemma Pi_UN:
+  fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set"
+  assumes "finite I" and mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i"
+  shows "(\<Union>n. Pi I (A n)) = (\<Pi> i\<in>I. \<Union>n. A n i)"
+proof (intro set_eqI iffI)
+  fix f assume "f \<in> (\<Pi> i\<in>I. \<Union>n. A n i)"
+  then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" by auto
+  from bchoice[OF this] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto
+  obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k"
+    using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto
+  have "f \<in> Pi I (A k)"
+  proof (intro Pi_I)
+    fix i assume "i \<in> I"
+    from mono[OF this, of "n i" k] k[OF this] n[OF this]
+    show "f i \<in> A k i" by auto
+  qed
+  then show "f \<in> (\<Union>n. Pi I (A n))" by auto
+qed auto
+
+lemma PiE_cong:
+  assumes "\<And>i. i\<in>I \<Longrightarrow> A i = B i"
+  shows "Pi\<^isub>E I A = Pi\<^isub>E I B"
+  using assms by (auto intro!: Pi_cong)
+
+lemma sigma_product_algebra_sigma_eq:
+  assumes "finite I"
+  assumes isotone: "\<And>i. i \<in> I \<Longrightarrow> (S i) \<up> (space (E i))"
+  assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)"
+  and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))"
+  shows "sigma (product_algebra (\<lambda>i. sigma (E i)) I) = sigma (product_algebra E I)"
+    (is "?S = ?E")
+proof cases
+  assume "I = {}" then show ?thesis by (simp add: product_algebra_def)
+next
+  assume "I \<noteq> {}"
+  interpret E: sigma_algebra "sigma (E i)" for i
+    using E by (rule sigma_algebra_sigma)
+
+  have into_space[intro]: "\<And>i x A. A \<in> sets (E i) \<Longrightarrow> x i \<in> A \<Longrightarrow> x i \<in> space (E i)"
+    using E by auto
+
+  interpret G: sigma_algebra ?E
+    unfolding product_algebra_def using E
+    by (intro sigma_algebra_sigma) (auto dest: Pi_mem)
+
+  { fix A i assume "i \<in> I" and A: "A \<in> sets (E i)"
+    then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E"
+      using isotone unfolding isoton_def product_algebra_def by (auto dest: Pi_mem)
+    also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))"
+      unfolding product_algebra_def
+      apply simp
+      apply (subst Pi_UN[OF `finite I`])
+      using isotone[THEN isoton_mono_le] apply simp
+      apply (simp add: PiE_Int)
+      apply (intro PiE_cong)
+      using A sets_into by (auto intro!: into_space)
+    also have "\<dots> \<in> sets ?E" unfolding product_algebra_def sets_sigma
+      using sets_into `A \<in> sets (E i)`
+      by (intro sigma_sets.Union)
+         (auto simp: image_subset_iff intro!: sigma_sets.Basic)
+    finally have "(\<lambda>x. x i) -` A \<inter> space ?E \<in> sets ?E" . }
+  then have proj:
+    "\<And>i. i\<in>I \<Longrightarrow> (\<lambda>x. x i) \<in> measurable ?E (sigma (E i))"
+    using E by (subst G.measurable_iff_sigma)
+               (auto simp: product_algebra_def sets_sigma)
+
+  { fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (sigma (E i))"
+    with proj have basic: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) -` (A i) \<inter> space ?E \<in> sets ?E"
+      unfolding measurable_def by simp
+    have "Pi\<^isub>E I A = (\<Inter>i\<in>I. (\<lambda>x. x i) -` (A i) \<inter> space ?E)"
+      using A E.sets_into_space `I \<noteq> {}` unfolding product_algebra_def by auto blast
+    then have "Pi\<^isub>E I A \<in> sets ?E"
+      using G.finite_INT[OF `finite I` `I \<noteq> {}` basic, of "\<lambda>i. i"] by simp }
+  then have "sigma_sets (space ?E) (sets (product_algebra (\<lambda>i. sigma (E i)) I)) \<subseteq> sets ?E"
+    by (intro G.sigma_sets_subset) (auto simp add: sets_sigma product_algebra_def)
+  then have subset: "sets ?S \<subseteq> sets ?E"
+    by (simp add: sets_sigma product_algebra_def)
+
+  have "sets ?S = sets ?E"
+  proof (intro set_eqI iffI)
+    fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S"
+      unfolding sets_sigma
+    proof induct
+      case (Basic A) then show ?case
+        by (auto simp: sets_sigma product_algebra_def intro: sigma_sets.Basic)
+    qed (auto intro: sigma_sets.intros simp: product_algebra_def)
+  next
+    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
+  qed
+  then show ?thesis
+    by (simp add: product_algebra_def sigma_def)
+qed
+
+lemma (in finite_product_sigma_algebra) pair_sigma_algebra_finite_product_space:
+  "sigma (pair_algebra P (M i)) = sigma (pair_algebra G (M i))"
+proof -
+  have "sigma (pair_algebra P (M i)) = sigma (pair_algebra P (sigma (M i)))" by simp
+  also have "\<dots> = sigma (pair_algebra G (M i))"
+  proof (rule pair_sigma_algebra_sigma)
+    show "(\<lambda>_. \<Pi>\<^isub>E i\<in>I. space (M i)) \<up> space G"
+      "(\<lambda>_. space (M i)) \<up> space (M i)"
+      by (simp_all add: isoton_const)
+    show "range (\<lambda>_. \<Pi>\<^isub>E i\<in>I. space (M i)) \<subseteq> sets G" "range (\<lambda>_. space (M i)) \<subseteq> sets (M i)"
+      by (auto intro!: image_eqI[where x="\<lambda>i\<in>I. space (M i)"] dest: Pi_mem
+               simp: product_algebra_def)
+    show "sets G \<subseteq> Pow (space G)" "sets (M i) \<subseteq> Pow (space (M i))"
+      using product_algebra_into_space M.sets_into_space by auto
+  qed
+  finally show ?thesis .
+qed
+
+lemma sets_pair_algebra: "sets (pair_algebra N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
+  unfolding pair_algebra_def by auto
+
+lemma (in finite_product_sigma_algebra) sigma_pair_algebra_sigma_eq:
+  "sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))) =
+   sigma (pair_algebra (product_algebra M I) (product_algebra M J))"
+  using M.sets_into_space
+  by (intro pair_sigma_algebra_sigma[of "\<lambda>_. \<Pi>\<^isub>E i\<in>I. space (M i)", of _ "\<lambda>_. \<Pi>\<^isub>E i\<in>J. space (M i)"])
+     (auto simp: isoton_const product_algebra_def, blast+)
+
+lemma (in product_sigma_algebra) product_product_vimage_algebra:
+  assumes [simp]: "I \<inter> J = {}" and "finite I" "finite J"
+  shows "sigma_algebra.vimage_algebra
+    (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))
+    (space (product_algebra M (I \<union> J))) (\<lambda>x. ((\<lambda>i\<in>I. x i), (\<lambda>i\<in>J. x i))) =
+    sigma (product_algebra M (I \<union> J))"
+    (is "sigma_algebra.vimage_algebra _ (space ?IJ) ?f = sigma ?IJ")
+proof -
+  have "finite (I \<union> J)" using assms by auto
+  interpret I: finite_product_sigma_algebra M I by default fact
+  interpret J: finite_product_sigma_algebra M J by default fact
+  interpret IJ: finite_product_sigma_algebra M "I \<union> J" by default fact
+  interpret pair_sigma_algebra I.P J.P by default
+
+  show "vimage_algebra (space ?IJ) ?f = sigma ?IJ"
+    unfolding I.sigma_pair_algebra_sigma_eq
+  proof (rule vimage_algebra_sigma)
+    from M.sets_into_space
+    show "sets (pair_algebra I.G J.G) \<subseteq> Pow (space (pair_algebra I.G J.G))"
+      by (auto simp: sets_pair_algebra space_pair_algebra product_algebra_def) blast+
+    show "?f \<in> space IJ.G \<rightarrow> space (pair_algebra I.G J.G)"
+      by (auto simp: space_pair_algebra product_algebra_def)
+    let ?F = "\<lambda>A. ?f -` A \<inter> (space IJ.G)"
+    let ?s = "\<lambda>I. Pi\<^isub>E I ` (\<Pi> i\<in>I. sets (M i))"
+    { fix A assume "A \<in> sets IJ.G"
+      then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F" "F \<in> (\<Pi> i\<in>I. sets (M i))" "F \<in> (\<Pi> i\<in>J. sets (M i))"
+        by (auto simp: product_algebra_def)
+      show "A \<in> ?F ` sets (pair_algebra I.G J.G)"
+          using A M.sets_into_space
+          by (auto simp: restrict_Pi_cancel product_algebra_def
+                   intro!: image_eqI[where x="Pi\<^isub>E I F \<times> Pi\<^isub>E J F"]) blast+ }
+    { fix A assume "A \<in> sets (pair_algebra I.G J.G)"
+      then obtain E F where A: "A = Pi\<^isub>E I E \<times> Pi\<^isub>E J F" "E \<in> (\<Pi> i\<in>I. sets (M i))" "F \<in> (\<Pi> i\<in>J. sets (M i))"
+        by (auto simp: product_algebra_def sets_pair_algebra)
+      then show "?F A \<in> sets IJ.G"
+          using A M.sets_into_space
+          by (auto simp: restrict_Pi_cancel product_algebra_def
+                   intro!: image_eqI[where x="merge I E J F"]) blast+ }
+  qed
+qed
+
+lemma (in finite_product_sigma_algebra) sigma_pair_algebra_sigma_M_eq:
+  "sigma (pair_algebra P (M i)) = sigma (pair_algebra G (M i))"
+proof -
+  have "sigma (pair_algebra P (sigma (M i))) = sigma (pair_algebra G (M i))"
+    using M.sets_into_space
+    by (intro pair_sigma_algebra_sigma[of "\<lambda>_. \<Pi>\<^isub>E i\<in>I. space (M i)", of _ "\<lambda>_. space (M i)"])
+       (auto simp: isoton_const product_algebra_def, blast+)
+  then show ?thesis by simp
+qed
+
+lemma (in product_sigma_algebra) product_singleton_vimage_algebra_eq:
+  assumes [simp]: "i \<notin> I" "finite I"
+  shows "sigma_algebra.vimage_algebra
+    (sigma (pair_algebra (sigma (product_algebra M I)) (M i)))
+    (space (product_algebra M (insert i I))) (\<lambda>x. ((\<lambda>i\<in>I. x i), x i)) =
+    sigma (product_algebra M (insert i I))"
+    (is "sigma_algebra.vimage_algebra _ (space ?I') ?f = sigma ?I'")
+proof -
+  have "finite (insert i I)" using assms by auto
+  interpret I: finite_product_sigma_algebra M I by default fact
+  interpret I': finite_product_sigma_algebra M "insert i I" by default fact
+  interpret pair_sigma_algebra I.P "M i" by default
+  show "vimage_algebra (space ?I') ?f = sigma ?I'"
+    unfolding I.sigma_pair_algebra_sigma_M_eq
+  proof (rule vimage_algebra_sigma)
+    from M.sets_into_space
+    show "sets (pair_algebra I.G (M i)) \<subseteq> Pow (space (pair_algebra I.G (M i)))"
+      by (auto simp: sets_pair_algebra space_pair_algebra product_algebra_def) blast
+    show "?f \<in> space I'.G \<rightarrow> space (pair_algebra I.G (M i))"
+      by (auto simp: space_pair_algebra product_algebra_def)
+    let ?F = "\<lambda>A. ?f -` A \<inter> (space I'.G)"
+    { fix A assume "A \<in> sets I'.G"
+      then obtain F where A: "A = Pi\<^isub>E (insert i I) F" "F \<in> (\<Pi> i\<in>I. sets (M i))" "F i \<in> sets (M i)"
+        by (auto simp: product_algebra_def)
+      show "A \<in> ?F ` sets (pair_algebra I.G (M i))"
+          using A M.sets_into_space
+          by (auto simp: restrict_Pi_cancel product_algebra_def
+                   intro!: image_eqI[where x="Pi\<^isub>E I F \<times> F i"]) blast+ }
+    { fix A assume "A \<in> sets (pair_algebra I.G (M i))"
+      then obtain E F where A: "A = Pi\<^isub>E I E \<times> F" "E \<in> (\<Pi> i\<in>I. sets (M i))" "F \<in> sets (M i)"
+        by (auto simp: product_algebra_def sets_pair_algebra)
+      then show "?F A \<in> sets I'.G"
+          using A M.sets_into_space
+          by (auto simp: restrict_Pi_cancel product_algebra_def
+                   intro!: image_eqI[where x="E(i:= F)"]) blast+ }
+  qed
+qed
 
-lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B"
-  by (auto simp add: prod_sets_def)
+lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I"
+  by (auto simp: restrict_def intro!: ext)
+
+lemma bij_betw_restrict_I_i:
+  "i \<notin> I \<Longrightarrow> bij_betw (\<lambda>x. (restrict x I, x i))
+    (space (product_algebra M (insert i I)))
+    (space (pair_algebra (sigma (product_algebra M I)) (M i)))"
+  by (intro bij_betwI[where g="(\<lambda>(x,y). x(i:=y))"])
+     (auto simp: space_pair_algebra extensional_def intro!: ext)
+
+lemma (in product_sigma_algebra) product_singleton_vimage_algebra_inv_eq:
+  assumes [simp]: "i \<notin> I" "finite I"
+  shows "sigma_algebra.vimage_algebra
+    (sigma (product_algebra M (insert i I)))
+    (space (pair_algebra (sigma (product_algebra M I)) (M i))) (\<lambda>(x,y). x(i:=y)) =
+    sigma (pair_algebra (sigma (product_algebra M I)) (M i))"
+proof -
+  have "finite (insert i I)" using `finite I` by auto
+  interpret I: finite_product_sigma_algebra M I by default fact
+  interpret I': finite_product_sigma_algebra M "insert i I" by default fact
+  interpret pair_sigma_algebra I.P "M i" by default
+  show ?thesis
+    unfolding product_singleton_vimage_algebra_eq[OF assms, symmetric]
+    using bij_betw_restrict_I_i[OF `i \<notin> I`, of M]
+    by (rule vimage_vimage_inv[unfolded space_sigma])
+       (auto simp: space_pair_algebra product_algebra_def dest: extensional_restrict)
+qed
+
+locale product_sigma_finite =
+  fixes M :: "'i \<Rightarrow> 'a algebra" and \<mu> :: "'i \<Rightarrow> 'a set \<Rightarrow> pinfreal"
+  assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i) (\<mu> i)"
+
+locale finite_product_sigma_finite = product_sigma_finite M \<mu> for M :: "'i \<Rightarrow> 'a algebra" and \<mu> +
+  fixes I :: "'i set" assumes finite_index': "finite I"
+
+sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" "\<mu> i" for i
+  by (rule sigma_finite_measures)
+
+sublocale product_sigma_finite \<subseteq> product_sigma_algebra
+  by default
+
+sublocale finite_product_sigma_finite \<subseteq> finite_product_sigma_algebra
+  by default (fact finite_index')
+
+lemma (in finite_product_sigma_finite) sigma_finite_pairs:
+  "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
+    (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
+    (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<omega>) \<and>
+    (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<up> space G"
+proof -
+  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> F \<up> space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<omega>)"
+    using M.sigma_finite_up by simp
+  from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
+  then have "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. F i \<up> space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<omega>"
+    by auto
+  let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
+  note space_product_algebra[simp]
+  show ?thesis
+  proof (intro exI[of _ F] conjI allI isotoneI set_eqI iffI ballI)
+    fix i show "range (F i) \<subseteq> sets (M i)" by fact
+  next
+    fix i k show "\<mu> i (F i k) \<noteq> \<omega>" by fact
+  next
+    fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space G"
+      using `\<And>i. range (F i) \<subseteq> sets (M i)` M.sets_into_space by auto blast
+  next
+    fix f assume "f \<in> space G"
+    with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"]
+      `\<And>i. F i \<up> space (M i)`[THEN isotonD(2)]
+      `\<And>i. F i \<up> space (M i)`[THEN isoton_mono_le]
+    show "f \<in> (\<Union>i. ?F i)" by auto
+  next
+    fix i show "?F i \<subseteq> ?F (Suc i)"
+      using `\<And>i. F i \<up> space (M i)`[THEN isotonD(1)] by auto
+  qed
+qed
+
+lemma (in product_sigma_finite) product_measure_exists:
+  assumes "finite I"
+  shows "\<exists>\<nu>. (\<forall>A\<in>(\<Pi> i\<in>I. sets (M i)). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))) \<and>
+     sigma_finite_measure (sigma (product_algebra M I)) \<nu>"
+using `finite I` proof induct
+  case empty then show ?case unfolding product_algebra_def
+    by (auto intro!: exI[of _ "\<lambda>A. if A = {} then 0 else 1"] sigma_algebra_sigma
+                     sigma_algebra.finite_additivity_sufficient
+             simp add: positive_def additive_def sets_sigma sigma_finite_measure_def
+                       sigma_finite_measure_axioms_def)
+next
+  case (insert i I)
+  interpret finite_product_sigma_finite M \<mu> I by default fact
+  have "finite (insert i I)" using `finite I` by auto
+  interpret I': finite_product_sigma_finite M \<mu> "insert i I" by default fact
+  from insert obtain \<nu> where
+    prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))" and
+    "sigma_finite_measure P \<nu>" by auto
+  interpret I: sigma_finite_measure P \<nu> by fact
+  interpret P: pair_sigma_finite P \<nu> "M i" "\<mu> i" ..
+
+  let ?h = "\<lambda>x. (restrict x I, x i)"
+  let ?\<nu> = "\<lambda>A. P.pair_measure (?h ` A)"
+  interpret I': measure_space "sigma (product_algebra M (insert i I))" ?\<nu>
+    unfolding product_singleton_vimage_algebra_eq[OF `i \<notin> I` `finite I`, symmetric]
+    using bij_betw_restrict_I_i[OF `i \<notin> I`, of M]
+    by (intro P.measure_space_isomorphic) auto
+
+  show ?case
+  proof (intro exI[of _ ?\<nu>] conjI ballI)
+    { fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))"
+      moreover then have "A \<in> (\<Pi> i\<in>I. sets (M i))" by auto
+      moreover have "(\<lambda>x. (restrict x I, x i)) ` Pi\<^isub>E (insert i I) A = Pi\<^isub>E I A \<times> A i"
+        using `i \<notin> I`
+        apply auto
+        apply (rule_tac x="a(i:=b)" in image_eqI)
+        apply (auto simp: extensional_def fun_eq_iff)
+        done
+      ultimately show "?\<nu> (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. \<mu> i (A i))"
+        apply simp
+        apply (subst P.pair_measure_times)
+        apply fastsimp
+        apply fastsimp
+        using `i \<notin> I` `finite I` prod[of A] by (auto simp: ac_simps) }
+    note product = this
+
+    show "sigma_finite_measure I'.P ?\<nu>"
+    proof
+      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)"
+        "(\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) \<up> space I'.G"
+        "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<omega>"
+        by blast+
+      let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k"
+      show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and>
+          (\<Union>i. F i) = space I'.P \<and> (\<forall>i. ?\<nu> (F i) \<noteq> \<omega>)"
+      proof (intro exI[of _ ?F] conjI allI)
+        show "range ?F \<subseteq> sets I'.P" using F(1) by auto
+      next
+        from F(2)[THEN isotonD(2)]
+        show "(\<Union>i. ?F i) = space I'.P" by simp
+      next
+        fix j
+        show "?\<nu> (?F j) \<noteq> \<omega>"
+          using F `finite I`
+          by (subst product) (auto simp: setprod_\<omega>)
+      qed
+    qed
+  qed
+qed
+
+definition (in finite_product_sigma_finite)
+  measure :: "('i \<Rightarrow> 'a) set \<Rightarrow> pinfreal" where
+  "measure = (SOME \<nu>.
+     (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))) \<and>
+     sigma_finite_measure P \<nu>)"
+
+sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P measure
+proof -
+  show "sigma_finite_measure P measure"
+    unfolding measure_def
+    by (rule someI2_ex[OF product_measure_exists[OF finite_index]]) auto
+qed
+
+lemma (in finite_product_sigma_finite) measure_times:
+  assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)"
+  shows "measure (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))"
+proof -
+  note ex = product_measure_exists[OF finite_index]
+  show ?thesis
+    unfolding measure_def
+  proof (rule someI2_ex[OF ex], elim conjE)
+    fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))"
+    have "Pi\<^isub>E I A = Pi\<^isub>E I (\<lambda>i\<in>I. A i)" by (auto dest: Pi_mem)
+    then have "\<nu> (Pi\<^isub>E I A) = \<nu> (Pi\<^isub>E I (\<lambda>i\<in>I. A i))" by simp
+    also have "\<dots> = (\<Prod>i\<in>I. \<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto
+    finally show "\<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. \<mu> i (A i))" by simp
+  qed
+qed
 
-lemma sigma_prod_sets_finite:
+abbreviation (in product_sigma_finite)
+  "product_measure I \<equiv> finite_product_sigma_finite.measure M \<mu> I"
+
+abbreviation (in product_sigma_finite)
+  "product_positive_integral I \<equiv>
+    measure_space.positive_integral (sigma (product_algebra M I)) (product_measure I)"
+
+abbreviation (in product_sigma_finite)
+  "product_integral I \<equiv>
+    measure_space.integral (sigma (product_algebra M I)) (product_measure I)"
+
+lemma (in product_sigma_finite) positive_integral_empty:
+  "product_positive_integral {} f = f (\<lambda>k. undefined)"
+proof -
+  interpret finite_product_sigma_finite M \<mu> "{}" by default (fact finite.emptyI)
+  have "\<And>A. measure (Pi\<^isub>E {} A) = 1"
+    using assms by (subst measure_times) auto
+  then show ?thesis
+    unfolding positive_integral_alt simple_function_def simple_integral_def_raw
+  proof (simp add: P_empty, intro antisym)
+    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> f}. f (\<lambda>k. undefined))"
+      by (intro le_SUPI) auto
+    show "(SUP f:{g. g \<le> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)"
+      by (intro SUP_leI) (auto simp: le_fun_def)
+  qed
+qed
+
+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"
+  unfolding merge_def by (auto intro!: ext)
+
+lemma merge_x_x_eq_restrict[simp]:
+  "merge I x J x = restrict x (I \<union> J)"
+  unfolding merge_def by (auto intro!: ext)
+
+lemma bij_betw_restrict_I_J:
+  "I \<inter> J = {} \<Longrightarrow> bij_betw (\<lambda>x. (restrict x I, restrict x J))
+    (space (product_algebra M (I \<union> J)))
+    (space (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J))))"
+  by (intro bij_betwI[where g="\<lambda>(x,y). merge I x J y"])
+     (auto dest: extensional_restrict simp: space_pair_algebra)
+
+lemma (in product_sigma_algebra) product_product_vimage_algebra_eq:
+  assumes [simp]: "I \<inter> J = {}" and "finite I" "finite J"
+  shows "sigma_algebra.vimage_algebra
+    (sigma (product_algebra M (I \<union> J)))
+    (space (sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J)))))
+    (\<lambda>(x, y). merge I x J y) =
+    sigma (pair_algebra (sigma (product_algebra M I)) (sigma (product_algebra M J)))"
+  (is "sigma_algebra.vimage_algebra ?IJ ?S ?m = ?P")
+proof -
+  interpret I: finite_product_sigma_algebra M I by default fact
+  interpret J: finite_product_sigma_algebra M J by default fact
+  have "finite (I \<union> J)" using assms by auto
+  interpret IJ: finite_product_sigma_algebra M "I \<union> J" by default fact
+  interpret P: pair_sigma_algebra I.P J.P by default
+
+  let ?g = "\<lambda>x. (restrict x I, restrict x J)"
+  let ?f = "\<lambda>(x, y). merge I x J y"
+  show "IJ.vimage_algebra (space P.P) ?f = P.P"
+    using bij_betw_restrict_I_J[OF `I \<inter> J = {}`]
+    by (subst P.vimage_vimage_inv[of ?g "space IJ.G" ?f,
+                   unfolded product_product_vimage_algebra[OF assms]])
+       (auto simp: pair_algebra_def dest: extensional_restrict)
+qed
+
+lemma (in product_sigma_finite) measure_fold_left:
+  assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
+  and f: "f \<in> borel_measurable (sigma (product_algebra M (I \<union> J)))"
+  shows "product_positive_integral (I \<union> J) f =
+    product_positive_integral I (\<lambda>x. product_positive_integral J (\<lambda>y. f (merge I x J y)))"
+proof -
+  interpret I: finite_product_sigma_finite M \<mu> I by default fact
+  interpret J: finite_product_sigma_finite M \<mu> J by default fact
+  have "finite (I \<union> J)" using fin by auto
+  interpret IJ: finite_product_sigma_finite M \<mu> "I \<union> J" by default fact
+  interpret P: pair_sigma_finite I.P I.measure J.P J.measure by default
+
+  let ?f = "\<lambda>x. ((\<lambda>i\<in>I. x i), (\<lambda>i\<in>J. x i))"
+
+  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P"
+    by (subst product_product_vimage_algebra_eq[OF IJ fin, symmetric])
+       (auto simp: space_pair_algebra intro!: IJ.measurable_vimage f)
+
+  have split_f_image[simp]: "\<And>F. ?f ` (Pi\<^isub>E (I \<union> J) F) = (Pi\<^isub>E I F) \<times> (Pi\<^isub>E J F)"
+    apply auto apply (rule_tac x="merge I a J b" in image_eqI)
+    by (auto dest: extensional_restrict)
+
+  have "IJ.positive_integral f =  IJ.positive_integral (\<lambda>x. f (restrict x (I \<union> J)))"
+    by (auto intro!: IJ.positive_integral_cong arg_cong[where f=f] dest!: extensional_restrict)
+  also have "\<dots> = I.positive_integral (\<lambda>x. J.positive_integral (\<lambda>y. f (merge I x J y)))"
+    unfolding P.positive_integral_fst_measurable[OF P_borel, simplified]
+    unfolding P.positive_integral_vimage[unfolded space_sigma, OF bij_betw_restrict_I_J[OF IJ]]
+    unfolding product_product_vimage_algebra[OF IJ fin]
+  proof (simp, rule IJ.positive_integral_cong_measure[symmetric])
+    fix A assume *: "A \<in> sets IJ.P"
+    from IJ.sigma_finite_pairs obtain F where
+      F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
+         "(\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) \<up> space IJ.G"
+         "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<omega>"
+      by auto
+    let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
+    show "P.pair_measure (?f ` A) = IJ.measure A"
+    proof (rule measure_unique_Int_stable[OF _ _ _ _ _ _ _ _ *])
+      show "Int_stable IJ.G" by (simp add: PiE_Int Int_stable_def product_algebra_def) auto
+      show "range ?F \<subseteq> sets IJ.G" using F by (simp add: image_subset_iff product_algebra_def)
+      show "?F \<up> space IJ.G " using F(2) by simp
+      show "measure_space IJ.P (\<lambda>A. P.pair_measure (?f ` A))"
+        unfolding product_product_vimage_algebra[OF IJ fin, symmetric]
+        using bij_betw_restrict_I_J[OF IJ, of M]
+        by (auto intro!: P.measure_space_isomorphic)
+      show "measure_space IJ.P IJ.measure" by fact
+    next
+      fix A assume "A \<in> sets IJ.G"
+      then obtain F where A[simp]: "A = Pi\<^isub>E (I \<union> J) F" "F \<in> (\<Pi> i\<in>I \<union> J. sets (M i))"
+        by (auto simp: product_algebra_def)
+      then have F: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets (M i)" "\<And>i. i \<in> J \<Longrightarrow> F i \<in> sets (M i)"
+        by auto
+      have "P.pair_measure (?f ` A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))"
+        using `finite J` `finite I` F
+        by (simp add: P.pair_measure_times I.measure_times J.measure_times)
+      also have "\<dots> = (\<Prod>i\<in>I \<union> J. \<mu> i (F i))"
+        using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
+      also have "\<dots> = IJ.measure A"
+        using `finite J` `finite I` F unfolding A
+        by (intro IJ.measure_times[symmetric]) auto
+      finally show "P.pair_measure (?f ` A) = IJ.measure A" .
+    next
+      fix k
+      have "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i k \<in> sets (M i)" using F by auto
+      then have "P.pair_measure (?f ` ?F k) = (\<Prod>i\<in>I. \<mu> i (F i k)) * (\<Prod>i\<in>J. \<mu> i (F i k))"
+        by (simp add: P.pair_measure_times I.measure_times J.measure_times)
+      then show "P.pair_measure (?f ` ?F k) \<noteq> \<omega>"
+        using `finite I` F by (simp add: setprod_\<omega>)
+    qed simp
+  qed
+  finally show ?thesis .
+qed
+
+lemma (in product_sigma_finite) finite_pair_measure_singleton:
+  assumes f: "f \<in> borel_measurable (M i)"
+  shows "product_positive_integral {i} (\<lambda>x. f (x i)) = M.positive_integral i f"
+proof -
+  interpret I: finite_product_sigma_finite M \<mu> "{i}" by default simp
+  have bij: "bij_betw (\<lambda>x. \<lambda>j\<in>{i}. x) (space (M i)) (space I.P)"
+    by (auto intro!: bij_betwI ext simp: extensional_def)
+  have *: "(\<lambda>x. (\<lambda>x. \<lambda>j\<in>{i}. x) -` Pi\<^isub>E {i} x \<inter> space (M i)) ` (\<Pi> i\<in>{i}. sets (M i)) = sets (M i)"
+  proof (subst image_cong, rule refl)
+    fix x assume "x \<in> (\<Pi> i\<in>{i}. sets (M i))"
+    then show "(\<lambda>x. \<lambda>j\<in>{i}. x) -` Pi\<^isub>E {i} x \<inter> space (M i) = x i"
+      using sets_into_space by auto
+  qed auto
+  have vimage: "I.vimage_algebra (space (M i)) (\<lambda>x. \<lambda>j\<in>{i}. x) = M i"
+    unfolding I.vimage_algebra_def
+    unfolding product_sigma_algebra_def sets_sigma
+    apply (subst sigma_sets_vimage[symmetric])
+    apply (simp_all add: image_image sigma_sets_eq product_algebra_def * del: vimage_Int)
+    using sets_into_space by blast
+  show "I.positive_integral (\<lambda>x. f (x i)) = M.positive_integral i f"
+    unfolding I.positive_integral_vimage[OF bij]
+    unfolding vimage
+    apply (subst positive_integral_cong_measure)
+  proof -
+    fix A assume A: "A \<in> sets (M i)"
+    have "(\<lambda>x. \<lambda>j\<in>{i}. x) ` A = (\<Pi>\<^isub>E i\<in>{i}. A)"
+      by (auto intro!: image_eqI ext[where 'b='a] simp: extensional_def)
+    with A show "product_measure {i} ((\<lambda>x. \<lambda>j\<in>{i}. x) ` A) = \<mu> i A"
+      using I.measure_times[of "\<lambda>i. A"] by simp
+  qed simp
+qed
+
+section "Products on finite spaces"
+
+lemma sigma_sets_pair_algebra_finite:
   assumes "finite A" and "finite B"
-  shows "sigma_sets (A \<times> B) (prod_sets (Pow A) (Pow B)) = Pow (A \<times> B)"
+  shows "sigma_sets (A \<times> B) ((\<lambda>(x,y). x \<times> y) ` (Pow A \<times> Pow B)) = Pow (A \<times> B)"
+  (is "sigma_sets ?prod ?sets = _")
 proof safe
   have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
-
   fix x assume subset: "x \<subseteq> A \<times> B"
   hence "finite x" using fin by (rule finite_subset)
-  from this subset show "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))"
-    (is "x \<in> sigma_sets ?prod ?sets")
+  from this subset show "x \<in> sigma_sets ?prod ?sets"
   proof (induct x)
     case empty show ?case by (rule sigma_sets.Empty)
   next
     case (insert a x)
-    hence "{a} \<in> sigma_sets ?prod ?sets" by (auto simp: prod_sets_def intro!: sigma_sets.Basic)
+    hence "{a} \<in> sigma_sets ?prod ?sets"
+      by (auto simp: pair_algebra_def intro!: sigma_sets.Basic)
     moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
     ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
   qed
 next
   fix x a b
-  assume "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" and "(a, b) \<in> x"
+  assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"
   from sigma_sets_into_sp[OF _ this(1)] this(2)
-  show "a \<in> A" and "b \<in> B"
-    by (auto simp: prod_sets_def)
+  show "a \<in> A" and "b \<in> B" by auto
 qed
 
-lemma (in sigma_algebra) measurable_prod_sigma:
-  assumes sa1: "sigma_algebra a1" and sa2: "sigma_algebra a2"
-  assumes 1: "(fst o f) \<in> measurable M a1" and 2: "(snd o f) \<in> measurable M a2"
-  shows "f \<in> measurable M (sigma ((space a1) \<times> (space a2))
-                          (prod_sets (sets a1) (sets a2)))"
+locale pair_finite_sigma_algebra = M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 for M1 M2
+
+sublocale pair_finite_sigma_algebra \<subseteq> pair_sigma_algebra by default
+
+lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra[simp]:
+  shows "P = (| space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2) |)"
 proof -
-  from 1 have fn1: "fst \<circ> f \<in> space M \<rightarrow> space a1"
-     and q1: "\<forall>y\<in>sets a1. (fst \<circ> f) -` y \<inter> space M \<in> sets M"
-    by (auto simp add: measurable_def)
-  from 2 have fn2: "snd \<circ> f \<in> space M \<rightarrow> space a2"
-     and q2: "\<forall>y\<in>sets a2. (snd \<circ> f) -` y \<inter> space M \<in> sets M"
-    by (auto simp add: measurable_def)
-  show ?thesis
-    proof (rule measurable_sigma)
-      show "prod_sets (sets a1) (sets a2) \<subseteq> Pow (space a1 \<times> space a2)" using sa1 sa2
-        by (auto simp add: prod_sets_def sigma_algebra_iff dest: algebra.space_closed)
-    next
-      show "f \<in> space M \<rightarrow> space a1 \<times> space a2"
-        by (rule prod_final [OF fn1 fn2])
-    next
-      fix z
-      assume z: "z \<in> prod_sets (sets a1) (sets a2)"
-      thus "f -` z \<inter> space M \<in> sets M"
-        proof (auto simp add: prod_sets_def vimage_Times)
-          fix x y
-          assume x: "x \<in> sets a1" and y: "y \<in> sets a2"
-          have "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M =
-                ((fst \<circ> f) -` x \<inter> space M) \<inter> ((snd \<circ> f) -` y \<inter> space M)"
-            by blast
-          also have "...  \<in> sets M" using x y q1 q2
-            by blast
-          finally show "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M \<in> sets M" .
-        qed
-    qed
+  show ?thesis using M1.finite_space M2.finite_space
+    by (simp add: sigma_def space_pair_algebra sets_pair_algebra
+                  sigma_sets_pair_algebra_finite M1.sets_eq_Pow M2.sets_eq_Pow)
+qed
+
+sublocale pair_finite_sigma_algebra \<subseteq> finite_sigma_algebra P
+proof
+  show "finite (space P)" "sets P = Pow (space P)"
+    using M1.finite_space M2.finite_space by auto
 qed
 
-lemma (in sigma_finite_measure) prod_measure_times:
-  assumes "sigma_finite_measure N \<nu>"
-  and "A1 \<in> sets M" "A2 \<in> sets N"
-  shows "prod_measure M \<mu> N \<nu> (A1 \<times> A2) = \<mu> A1 * \<nu> A2"
-  oops
+locale pair_finite_space = M1: finite_measure_space M1 \<mu>1 + M2: finite_measure_space M2 \<mu>2
+  for M1 \<mu>1 M2 \<mu>2
+
+sublocale pair_finite_space \<subseteq> pair_finite_sigma_algebra
+  by default
 
-lemma (in sigma_finite_measure) sigma_finite_prod_measure_space:
-  assumes "sigma_finite_measure N \<nu>"
-  shows "sigma_finite_measure (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)"
-  oops
+sublocale pair_finite_space \<subseteq> pair_sigma_finite
+  by default
 
-lemma (in finite_measure_space) finite_prod_measure_times:
-  assumes "finite_measure_space N \<nu>"
-  and "A1 \<in> sets M" "A2 \<in> sets N"
-  shows "prod_measure M \<mu> N \<nu> (A1 \<times> A2) = \<mu> A1 * \<nu> A2"
+lemma (in pair_finite_space) finite_pair_sigma_algebra[simp]:
+  shows "P = (| space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2) |)"
 proof -
-  interpret N: finite_measure_space N \<nu> by fact
-  have *: "\<And>x. \<nu> (Pair x -` (A1 \<times> A2)) * \<mu> {x} = (if x \<in> A1 then \<nu> A2 * \<mu> {x} else 0)"
-    by (auto simp: vimage_Times comp_def)
-  have "finite A1"
-    using `A1 \<in> sets M` finite_space by (auto simp: sets_eq_Pow intro: finite_subset)
-  then have "\<mu> A1 = (\<Sum>x\<in>A1. \<mu> {x})" using `A1 \<in> sets M`
-    by (auto intro!: measure_finite_singleton simp: sets_eq_Pow)
-  then show ?thesis using `A1 \<in> sets M`
-    unfolding prod_measure_def positive_integral_finite_eq_setsum *
-    by (auto simp add: sets_eq_Pow setsum_right_distrib[symmetric] mult_commute setsum_cases[OF finite_space])
-qed
-
-lemma (in finite_measure_space) finite_prod_measure_space:
-  assumes "finite_measure_space N \<nu>"
-  shows "prod_measure_space M N = \<lparr> space = space M \<times> space N, sets = Pow (space M \<times> space N) \<rparr>"
-proof -
-  interpret N: finite_measure_space N \<nu> by fact
-  show ?thesis using finite_space N.finite_space
-    by (simp add: sigma_def prod_measure_space_def sigma_prod_sets_finite sets_eq_Pow N.sets_eq_Pow)
+  show ?thesis using M1.finite_space M2.finite_space
+    by (simp add: sigma_def space_pair_algebra sets_pair_algebra
+                  sigma_sets_pair_algebra_finite M1.sets_eq_Pow M2.sets_eq_Pow)
 qed
 
-lemma (in finite_measure_space) finite_measure_space_finite_prod_measure:
-  fixes N :: "('c, 'd) algebra_scheme"
-  assumes N: "finite_measure_space N \<nu>"
-  shows "finite_measure_space (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)"
-  unfolding finite_prod_measure_space[OF assms]
-proof (rule finite_measure_spaceI, simp_all)
-  interpret N: finite_measure_space N \<nu> by fact
-  show "finite (space M \<times> space N)" using finite_space N.finite_space by auto
-  show "prod_measure M \<mu> N \<nu> (space M \<times> space N) \<noteq> \<omega>"
-    using finite_prod_measure_times[OF N top N.top] by simp
-  show "prod_measure M \<mu> N \<nu> {} = 0"
-    using finite_prod_measure_times[OF N empty_sets N.empty_sets] by simp
-
-  fix A B :: "('a * 'c) set" assume "A \<inter> B = {}" "A \<subseteq> space M \<times> space N" "B \<subseteq> space M \<times> space N"
-  then show "prod_measure M \<mu> N \<nu> (A \<union> B) = prod_measure M \<mu> N \<nu> A + prod_measure M \<mu> N \<nu> B"
-    apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric]
-                intro!: positive_integral_cong)
-    apply (subst N.measure_additive)
-    apply (auto intro!: arg_cong[where f=\<mu>] simp: N.sets_eq_Pow sets_eq_Pow)
-    done
+lemma (in pair_finite_space) pair_measure_Pair[simp]:
+  assumes "a \<in> space M1" "b \<in> space M2"
+  shows "pair_measure {(a, b)} = \<mu>1 {a} * \<mu>2 {b}"
+proof -
+  have "pair_measure ({a}\<times>{b}) = \<mu>1 {a} * \<mu>2 {b}"
+    using M1.sets_eq_Pow M2.sets_eq_Pow assms
+    by (subst pair_measure_times) auto
+  then show ?thesis by simp
 qed
 
-lemma (in finite_measure_space) finite_measure_space_finite_prod_measure_alterantive:
-  assumes N: "finite_measure_space N \<nu>"
-  shows "finite_measure_space \<lparr> space = space M \<times> space N, sets = Pow (space M \<times> space N) \<rparr> (prod_measure M \<mu> N \<nu>)"
-    (is "finite_measure_space ?M ?m")
-  unfolding finite_prod_measure_space[OF N, symmetric]
-  using finite_measure_space_finite_prod_measure[OF N] .
+lemma (in pair_finite_space) pair_measure_singleton[simp]:
+  assumes "x \<in> space M1 \<times> space M2"
+  shows "pair_measure {x} = \<mu>1 {fst x} * \<mu>2 {snd x}"
+  using pair_measure_Pair assms by (cases x) auto
 
-lemma prod_measure_times_finite:
-  assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N"
-  shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}"
-proof (cases a)
-  case (Pair b c)
-  hence a_eq: "{a} = {b} \<times> {c}" by simp
+sublocale pair_finite_space \<subseteq> finite_measure_space P pair_measure
+  by default auto
 
-  interpret M: finite_measure_space M \<mu> by fact
-  interpret N: finite_measure_space N \<nu> by fact
+lemma (in pair_finite_space) finite_measure_space_finite_prod_measure_alterantive:
+  "finite_measure_space \<lparr>space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2)\<rparr> pair_measure"
+  unfolding finite_pair_sigma_algebra[symmetric]
+  by default
 
-  from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair
-  show ?thesis unfolding a_eq by simp
-qed
-
-end
+end
\ No newline at end of file
--- a/src/HOL/Probability/Radon_Nikodym.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -2,6 +2,14 @@
 imports Lebesgue_Integration
 begin
 
+lemma less_\<omega>_Ex_of_nat: "x < \<omega> \<longleftrightarrow> (\<exists>n. x < of_nat n)"
+proof safe
+  assume "x < \<omega>"
+  then obtain r where "0 \<le> r" "x = Real r" by (cases x) auto
+  moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto
+  ultimately show "\<exists>n. x < of_nat n" by (auto simp: real_eq_of_nat)
+qed auto
+
 lemma (in sigma_finite_measure) Ex_finite_integrable_function:
   shows "\<exists>h\<in>borel_measurable M. positive_integral h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
 proof -
@@ -64,6 +72,21 @@
 definition (in measure_space)
   "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pinfreal))"
 
+lemma (in sigma_finite_measure) absolutely_continuous_AE:
+  assumes "measure_space M \<nu>" "absolutely_continuous \<nu>" "AE x. P x"
+  shows "measure_space.almost_everywhere M \<nu> P"
+proof -
+  interpret \<nu>: measure_space M \<nu> by fact
+  from `AE x. P x` obtain N where N: "N \<in> null_sets" and "{x\<in>space M. \<not> P x} \<subseteq> N"
+    unfolding almost_everywhere_def by auto
+  show "\<nu>.almost_everywhere P"
+  proof (rule \<nu>.AE_I')
+    show "{x\<in>space M. \<not> P x} \<subseteq> N" by fact
+    from `absolutely_continuous \<nu>` show "N \<in> \<nu>.null_sets"
+      using N unfolding absolutely_continuous_def by auto
+  qed
+qed
+
 lemma (in finite_measure_space) absolutely_continuousI:
   assumes "finite_measure_space M \<nu>"
   assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
@@ -542,10 +565,12 @@
   qed simp
 qed
 
-lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
+lemma (in finite_measure) split_space_into_finite_sets_and_rest:
   assumes "measure_space M \<nu>"
-  assumes "absolutely_continuous \<nu>"
-  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+  assumes ac: "absolutely_continuous \<nu>"
+  shows "\<exists>\<Omega>0\<in>sets M. \<exists>\<Omega>::nat\<Rightarrow>'a set. disjoint_family \<Omega> \<and> range \<Omega> \<subseteq> sets M \<and> \<Omega>0 = space M - (\<Union>i. \<Omega> i) \<and>
+    (\<forall>A\<in>sets M. A \<subseteq> \<Omega>0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<omega>)) \<and>
+    (\<forall>i. \<nu> (\<Omega> i) \<noteq> \<omega>)"
 proof -
   interpret v: measure_space M \<nu> by fact
   let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<omega>}"
@@ -604,61 +629,98 @@
   let "?O_0" = "(\<Union>i. ?O i)"
   have "?O_0 \<in> sets M" using Q' by auto
 
-  { fix A assume *: "A \<in> ?Q" "A \<subseteq> space M - ?O_0"
-    then have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
-      using Q' by (auto intro!: measure_additive countable_UN)
-    also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
-    proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
-      show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
-        using * O_sets by auto
-    qed fastsimp
-    also have "\<dots> \<le> ?a"
-    proof (safe intro!: SUPR_bound)
-      fix i have "?O i \<union> A \<in> ?Q"
-      proof (safe del: notI)
-        show "?O i \<union> A \<in> sets M" using O_sets * by auto
-        from O_in_G[of i]
-        moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
-          using v.measure_subadditive[of "?O i" A] * O_sets by auto
-        ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>"
-          using * by auto
-      qed
-      then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
-    qed
-    finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`]
-      by (cases "\<mu> A") (auto simp: pinfreal_noteq_omega_Ex) }
-  note stetic = this
-
-  def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> ?O 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"
-
+  def Q \<equiv> "\<lambda>i. case i of 0 \<Rightarrow> Q' 0 | Suc n \<Rightarrow> ?O (Suc n) - ?O n"
   { fix i have "Q i \<in> sets M" unfolding Q_def using Q'[of 0] by (cases i) (auto intro: O_sets) }
   note Q_sets = this
 
-  { fix i have "\<nu> (Q i) \<noteq> \<omega>"
-    proof (cases i)
-      case 0 then show ?thesis
-        unfolding Q_def using Q'[of 0] by simp
-    next
-      case (Suc n)
-      then show ?thesis unfolding Q_def
-        using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono
-        using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto
-    qed }
-  note Q_omega = this
+  show ?thesis
+  proof (intro bexI exI conjI ballI impI allI)
+    show "disjoint_family Q"
+      by (fastsimp simp: disjoint_family_on_def Q_def
+        split: nat.split_asm)
+    show "range Q \<subseteq> sets M"
+      using Q_sets by auto
 
-  { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
-    proof (induct j)
-      case 0 then show ?case by (simp add: Q_def)
-    next
-      case (Suc j)
-      have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp
-      have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
-      then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)"
-        by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
-      then show ?case using Suc by (auto simp add: eq atMost_Suc)
-    qed }
-  then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
-  then have O_0_eq_Q: "?O_0 = (\<Union>j. Q j)" by fastsimp
+    { fix A assume A: "A \<in> sets M" "A \<subseteq> space M - ?O_0"
+      show "\<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>"
+      proof (rule disjCI, simp)
+        assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<omega>"
+        show "\<mu> A = 0 \<and> \<nu> A = 0"
+        proof cases
+          assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0"
+            unfolding absolutely_continuous_def by auto
+          ultimately show ?thesis by simp
+        next
+          assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<omega>" by auto
+          with A have "\<mu> ?O_0 + \<mu> A = \<mu> (?O_0 \<union> A)"
+            using Q' by (auto intro!: measure_additive countable_UN)
+          also have "\<dots> = (SUP i. \<mu> (?O i \<union> A))"
+          proof (rule continuity_from_below[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
+            show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
+              using `\<nu> A \<noteq> \<omega>` O_sets A by auto
+          qed fastsimp
+          also have "\<dots> \<le> ?a"
+          proof (safe intro!: SUPR_bound)
+            fix i have "?O i \<union> A \<in> ?Q"
+            proof (safe del: notI)
+              show "?O i \<union> A \<in> sets M" using O_sets A by auto
+              from O_in_G[of i]
+              moreover have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
+                using v.measure_subadditive[of "?O i" A] A O_sets by auto
+              ultimately show "\<nu> (?O i \<union> A) \<noteq> \<omega>"
+                using `\<nu> A \<noteq> \<omega>` by auto
+            qed
+            then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule le_SUPI)
+          qed
+          finally have "\<mu> A = 0" unfolding a_eq using finite_measure[OF `?O_0 \<in> sets M`]
+            by (cases "\<mu> A") (auto simp: pinfreal_noteq_omega_Ex)
+          with `\<mu> A \<noteq> 0` show ?thesis by auto
+        qed
+      qed }
+
+    { fix i show "\<nu> (Q i) \<noteq> \<omega>"
+      proof (cases i)
+        case 0 then show ?thesis
+          unfolding Q_def using Q'[of 0] by simp
+      next
+        case (Suc n)
+        then show ?thesis unfolding Q_def
+          using `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q` O_mono
+          using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto
+      qed }
+
+    show "space M - ?O_0 \<in> sets M" using Q'_sets by auto
+
+    { fix j have "(\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q i)"
+      proof (induct j)
+        case 0 then show ?case by (simp add: Q_def)
+      next
+        case (Suc j)
+        have eq: "\<And>j. (\<Union>i\<le>j. ?O i) = (\<Union>i\<le>j. Q' i)" by fastsimp
+        have "{..j} \<union> {..Suc j} = {..Suc j}" by auto
+        then have "(\<Union>i\<le>Suc j. Q' i) = (\<Union>i\<le>j. Q' i) \<union> Q (Suc j)"
+          by (simp add: UN_Un[symmetric] Q_def del: UN_Un)
+        then show ?case using Suc by (auto simp add: eq atMost_Suc)
+      qed }
+    then have "(\<Union>j. (\<Union>i\<le>j. ?O i)) = (\<Union>j. (\<Union>i\<le>j. Q i))" by simp
+    then show "space M - ?O_0 = space M - (\<Union>i. Q i)" by fastsimp
+  qed
+qed
+
+lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
+  assumes "measure_space M \<nu>"
+  assumes "absolutely_continuous \<nu>"
+  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+proof -
+  interpret v: measure_space M \<nu> by fact
+
+  from split_space_into_finite_sets_and_rest[OF assms]
+  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
+    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
+    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> \<nu> A = 0 \<or> 0 < \<mu> A \<and> \<nu> A = \<omega>"
+    and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<omega>" by force
+  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
 
   have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>A\<in>sets M.
     \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f x * indicator (Q i \<inter> A) x))"
@@ -677,7 +739,7 @@
       show "measure_space ?R \<nu>"
         using v.restricted_measure_space Q_sets[of i] by auto
       show "\<nu>  (space ?R) \<noteq> \<omega>"
-        using Q_omega by simp
+        using Q_fin by simp
     qed
     have "R.absolutely_continuous \<nu>"
       using `absolutely_continuous \<nu>` `Q i \<in> sets M`
@@ -697,71 +759,49 @@
       \<nu> (Q i \<inter> A) = positive_integral (\<lambda>x. f i x * indicator (Q i \<inter> A) x)"
     by auto
   let "?f x" =
-    "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator (space M - ?O_0) x"
+    "(\<Sum>\<^isub>\<infinity> i. f i x * indicator (Q i) x) + \<omega> * indicator Q0 x"
   show ?thesis
   proof (safe intro!: bexI[of _ ?f])
     show "?f \<in> borel_measurable M"
       by (safe intro!: borel_measurable_psuminf borel_measurable_pinfreal_times
         borel_measurable_pinfreal_add borel_measurable_indicator
-        borel_measurable_const borel Q_sets O_sets Diff countable_UN)
+        borel_measurable_const borel Q_sets Q0 Diff countable_UN)
     fix A assume "A \<in> sets M"
-    let ?C = "(space M - (\<Union>i. Q i)) \<inter> A"
-    have *: 
+    have *:
       "\<And>x i. indicator A x * (f i x * indicator (Q i) x) =
         f i x * indicator (Q i \<inter> A) x"
-      "\<And>x i. (indicator A x * indicator (space M - (\<Union>i. UNION {..i} Q')) x :: pinfreal) =
-        indicator ?C x" unfolding O_0_eq_Q by (auto simp: indicator_def)
+      "\<And>x i. (indicator A x * indicator Q0 x :: pinfreal) =
+        indicator (Q0 \<inter> A) x" by (auto simp: indicator_def)
     have "positive_integral (\<lambda>x. ?f x * indicator A x) =
-      (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> ?C"
+      (\<Sum>\<^isub>\<infinity> i. \<nu> (Q i \<inter> A)) + \<omega> * \<mu> (Q0 \<inter> A)"
       unfolding f[OF `A \<in> sets M`]
-      apply (simp del: pinfreal_times(2) add: field_simps)
+      apply (simp del: pinfreal_times(2) add: field_simps *)
       apply (subst positive_integral_add)
-      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
-        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
+      apply (fastsimp intro: Q0 `A \<in> sets M`)
+      apply (fastsimp intro: Q_sets `A \<in> sets M` borel_measurable_psuminf borel)
+      apply (subst positive_integral_cmult_indicator)
+      apply (fastsimp intro: Q0 `A \<in> sets M`)
       unfolding psuminf_cmult_right[symmetric]
       apply (subst positive_integral_psuminf)
-      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
-        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
-      apply (subst positive_integral_cmult)
-      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const
-        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
-      unfolding *
-      apply (subst positive_integral_indicator)
-      apply (safe intro!: borel_measurable_pinfreal_times Diff borel_measurable_const Int
-        borel_measurable_psuminf borel_measurable_indicator `A \<in> sets M` Q_sets borel countable_UN Q'_sets)
-      by simp
+      apply (fastsimp intro: `A \<in> sets M` Q_sets borel)
+      apply (simp add: *)
+      done
     moreover have "(\<Sum>\<^isub>\<infinity>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
-    proof (rule v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
-      show "range (\<lambda>i. Q i \<inter> A) \<subseteq> sets M"
-        using Q_sets `A \<in> sets M` by auto
-      show "disjoint_family (\<lambda>i. Q i \<inter> A)"
-        by (fastsimp simp: disjoint_family_on_def Q_def
-          split: nat.split_asm)
+      using Q Q_sets `A \<in> sets M`
+      by (intro v.measure_countably_additive[of "\<lambda>i. Q i \<inter> A", unfolded comp_def, simplified])
+         (auto simp: disjoint_family_on_def)
+    moreover have "\<omega> * \<mu> (Q0 \<inter> A) = \<nu> (Q0 \<inter> A)"
+    proof -
+      have "Q0 \<inter> A \<in> sets M" using Q0(1) `A \<in> sets M` by blast
+      from in_Q0[OF this] show ?thesis by auto
     qed
-    moreover have "\<omega> * \<mu> ?C = \<nu> ?C"
-    proof cases
-      assume null: "\<mu> ?C = 0"
-      hence "?C \<in> null_sets" using Q_sets `A \<in> sets M` by auto
-      with `absolutely_continuous \<nu>` and null
-      show ?thesis by (simp add: absolutely_continuous_def)
-    next
-      assume not_null: "\<mu> ?C \<noteq> 0"
-      have "\<nu> ?C = \<omega>"
-      proof (rule ccontr)
-        assume "\<nu> ?C \<noteq> \<omega>"
-        then have "?C \<in> ?Q"
-          using Q_sets `A \<in> sets M` by auto
-        from stetic[OF this] not_null
-        show False unfolding O_0_eq_Q by auto
-      qed
-      then show ?thesis using not_null by simp
-    qed
-    moreover have "?C \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
-      using Q_sets `A \<in> sets M` by (auto intro!: countable_UN)
-    moreover have "((\<Union>i. Q i) \<inter> A) \<union> ?C = A" "((\<Union>i. Q i) \<inter> A) \<inter> ?C = {}"
-      using `A \<in> sets M` sets_into_space by auto
+    moreover have "Q0 \<inter> A \<in> sets M" "((\<Union>i. Q i) \<inter> A) \<in> sets M"
+      using Q_sets `A \<in> sets M` Q0(1) by (auto intro!: countable_UN)
+    moreover have "((\<Union>i. Q i) \<inter> A) \<union> (Q0 \<inter> A) = A" "((\<Union>i. Q i) \<inter> A) \<inter> (Q0 \<inter> A) = {}"
+      using `A \<in> sets M` sets_into_space Q0 by auto
     ultimately show "\<nu> A = positive_integral (\<lambda>x. ?f x * indicator A x)"
-      using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" ?C] by auto
+      using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"]
+      by simp
   qed
 qed
 
@@ -801,12 +841,283 @@
   qed
 qed
 
+section "Uniqueness of densities"
+
+lemma (in measure_space) density_is_absolutely_continuous:
+  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+  shows "absolutely_continuous \<nu>"
+  using assms unfolding absolutely_continuous_def
+  by (simp add: positive_integral_null_set)
+
+lemma (in measure_space) finite_density_unique:
+  assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  and fin: "positive_integral f < \<omega>"
+  shows "(\<forall>A\<in>sets M. positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. g x * indicator A x))
+    \<longleftrightarrow> (AE x. f x = g x)"
+    (is "(\<forall>A\<in>sets M. ?P f A = ?P g A) \<longleftrightarrow> _")
+proof (intro iffI ballI)
+  fix A assume eq: "AE x. f x = g x"
+  show "?P f A = ?P g A"
+    by (rule positive_integral_cong_AE[OF AE_mp[OF eq]]) simp
+next
+  assume eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+  from this[THEN bspec, OF top] fin
+  have g_fin: "positive_integral g < \<omega>" by (simp cong: positive_integral_cong)
+  { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+      and g_fin: "positive_integral g < \<omega>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
+    let ?N = "{x\<in>space M. g x < f x}"
+    have N: "?N \<in> sets M" using borel by simp
+    have "?P (\<lambda>x. (f x - g x)) ?N = positive_integral (\<lambda>x. f x * indicator ?N x - g x * indicator ?N x)"
+      by (auto intro!: positive_integral_cong simp: indicator_def)
+    also have "\<dots> = ?P f ?N - ?P g ?N"
+    proof (rule positive_integral_diff)
+      show "(\<lambda>x. f x * indicator ?N x) \<in> borel_measurable M" "(\<lambda>x. g x * indicator ?N x) \<in> borel_measurable M"
+        using borel N by auto
+      have "?P g ?N \<le> positive_integral g"
+        by (auto intro!: positive_integral_mono simp: indicator_def)
+      then show "?P g ?N \<noteq> \<omega>" using g_fin by auto
+      fix x assume "x \<in> space M"
+      show "g x * indicator ?N x \<le> f x * indicator ?N x"
+        by (auto simp: indicator_def)
+    qed
+    also have "\<dots> = 0"
+      using eq[THEN bspec, OF N] by simp
+    finally have "\<mu> {x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = 0"
+      using borel N by (subst (asm) positive_integral_0_iff) auto
+    moreover have "{x\<in>space M. (f x - g x) * indicator ?N x \<noteq> 0} = ?N"
+      by (auto simp: pinfreal_zero_le_diff)
+    ultimately have "?N \<in> null_sets" using N by simp }
+  from this[OF borel g_fin eq] this[OF borel(2,1) fin]
+  have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} \<in> null_sets"
+    using eq by (intro null_sets_Un) auto
+  also have "{x\<in>space M. g x < f x} \<union> {x\<in>space M. f x < g x} = {x\<in>space M. f x \<noteq> g x}"
+    by auto
+  finally show "AE x. f x = g x"
+    unfolding almost_everywhere_def by auto
+qed
+
+lemma (in finite_measure) density_unique_finite_measure:
+  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
+  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. f' x * indicator A x)"
+    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
+  shows "AE x. f x = f' x"
+proof -
+  let "?\<nu> A" = "?P f A" and "?\<nu>' A" = "?P f' A"
+  let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x"
+  interpret M: measure_space M ?\<nu>
+    using borel(1) by (rule measure_space_density)
+  have ac: "absolutely_continuous ?\<nu>"
+    using f by (rule density_is_absolutely_continuous)
+  from split_space_into_finite_sets_and_rest[OF `measure_space M ?\<nu>` ac]
+  obtain Q0 and Q :: "nat \<Rightarrow> 'a set"
+    where Q: "disjoint_family Q" "range Q \<subseteq> sets M"
+    and Q0: "Q0 \<in> sets M" "Q0 = space M - (\<Union>i. Q i)"
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> \<mu> A = 0 \<and> ?\<nu> A = 0 \<or> 0 < \<mu> A \<and> ?\<nu> A = \<omega>"
+    and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<omega>" by force
+  from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
+  let ?N = "{x\<in>space M. f x \<noteq> f' x}"
+  have "?N \<in> sets M" using borel by auto
+  have *: "\<And>i x A. \<And>y::pinfreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
+    unfolding indicator_def by auto
+  have 1: "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x"
+    using borel Q_fin Q
+    by (intro finite_density_unique[THEN iffD1] allI)
+       (auto intro!: borel_measurable_pinfreal_times f Int simp: *)
+  have 2: "AE x. ?f Q0 x = ?f' Q0 x"
+  proof (rule AE_I')
+    { fix f :: "'a \<Rightarrow> pinfreal" assume borel: "f \<in> borel_measurable M"
+        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+      let "?A i" = "Q0 \<inter> {x \<in> space M. f x < of_nat i}"
+      have "(\<Union>i. ?A i) \<in> null_sets"
+      proof (rule null_sets_UN)
+        fix i have "?A i \<in> sets M"
+          using borel Q0(1) by auto
+        have "?\<nu> (?A i) \<le> positive_integral (\<lambda>x. of_nat i * indicator (?A i) x)"
+          unfolding eq[OF `?A i \<in> sets M`]
+          by (auto intro!: positive_integral_mono simp: indicator_def)
+        also have "\<dots> = of_nat i * \<mu> (?A i)"
+          using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
+        also have "\<dots> < \<omega>"
+          using `?A i \<in> sets M`[THEN finite_measure] by auto
+        finally have "?\<nu> (?A i) \<noteq> \<omega>" by simp
+        then show "?A i \<in> null_sets" using in_Q0[OF `?A i \<in> sets M`] `?A i \<in> sets M` by auto
+      qed
+      also have "(\<Union>i. ?A i) = Q0 \<inter> {x\<in>space M. f x < \<omega>}"
+        by (auto simp: less_\<omega>_Ex_of_nat)
+      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" by (simp add: pinfreal_less_\<omega>) }
+    from this[OF borel(1) refl] this[OF borel(2) f]
+    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>} \<in> null_sets" by simp_all
+    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>}) \<in> null_sets" by (rule null_sets_Un)
+    show "{x \<in> space M. ?f Q0 x \<noteq> ?f' Q0 x} \<subseteq>
+      (Q0 \<inter> {x\<in>space M. f x \<noteq> \<omega>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<omega>})" by (auto simp: indicator_def)
+  qed
+  have **: "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
+    ?f (space M) x = ?f' (space M) x"
+    by (auto simp: indicator_def Q0)
+  have 3: "AE x. ?f (space M) x = ?f' (space M) x"
+    by (rule AE_mp[OF 1[unfolded all_AE_countable] AE_mp[OF 2]]) (simp add: **)
+  then show "AE x. f x = f' x"
+    by (rule AE_mp) (auto intro!: AE_cong simp: indicator_def)
+qed
+
+lemma (in sigma_finite_measure) density_unique:
+  assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
+  assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. f' x * indicator A x)"
+    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
+  shows "AE x. f x = f' x"
+proof -
+  obtain h where h_borel: "h \<in> borel_measurable M"
+    and fin: "positive_integral h \<noteq> \<omega>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>"
+    using Ex_finite_integrable_function by auto
+  interpret h: measure_space M "\<lambda>A. positive_integral (\<lambda>x. h x * indicator A x)"
+    using h_borel by (rule measure_space_density)
+  interpret h: finite_measure M "\<lambda>A. positive_integral (\<lambda>x. h x * indicator A x)"
+    by default (simp cong: positive_integral_cong add: fin)
+
+  interpret f: measure_space M "\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)"
+    using borel(1) by (rule measure_space_density)
+  interpret f': measure_space M "\<lambda>A. positive_integral (\<lambda>x. f' x * indicator A x)"
+    using borel(2) by (rule measure_space_density)
+
+  { fix A assume "A \<in> sets M"
+    then have " {x \<in> space M. h x \<noteq> 0 \<and> indicator A x \<noteq> (0::pinfreal)} = A"
+      using pos sets_into_space by (force simp: indicator_def)
+    then have "positive_integral (\<lambda>xa. h xa * indicator A xa) = 0 \<longleftrightarrow> A \<in> null_sets"
+      using h_borel `A \<in> sets M` by (simp add: positive_integral_0_iff) }
+  note h_null_sets = this
+
+  { fix A assume "A \<in> sets M"
+    have "positive_integral (\<lambda>x. h x * (f x * indicator A x)) =
+      f.positive_integral (\<lambda>x. h x * indicator A x)"
+      using `A \<in> sets M` h_borel borel
+      by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
+    also have "\<dots> = f'.positive_integral (\<lambda>x. h x * indicator A x)"
+      by (rule f'.positive_integral_cong_measure) (rule f)
+    also have "\<dots> = positive_integral (\<lambda>x. h x * (f' x * indicator A x))"
+      using `A \<in> sets M` h_borel borel
+      by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong)
+    finally have "positive_integral (\<lambda>x. h x * (f x * indicator A x)) = positive_integral (\<lambda>x. h x * (f' x * indicator A x))" . }
+  then have "h.almost_everywhere (\<lambda>x. f x = f' x)"
+    using h_borel borel
+    by (intro h.density_unique_finite_measure[OF borel])
+       (simp add: positive_integral_translated_density)
+  then show "AE x. f x = f' x"
+    unfolding h.almost_everywhere_def almost_everywhere_def
+    by (auto simp add: h_null_sets)
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
+  assumes \<nu>: "measure_space M \<nu>" and f: "f \<in> borel_measurable M"
+    and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+  shows "sigma_finite_measure M \<nu> \<longleftrightarrow> (AE x. f x \<noteq> \<omega>)"
+proof
+  assume "sigma_finite_measure M \<nu>"
+  then interpret \<nu>: sigma_finite_measure M \<nu> .
+  from \<nu>.Ex_finite_integrable_function obtain h where
+    h: "h \<in> borel_measurable M" "\<nu>.positive_integral h \<noteq> \<omega>"
+    and fin: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<omega>" by auto
+  have "AE x. f x * h x \<noteq> \<omega>"
+  proof (rule AE_I')
+    have "\<nu>.positive_integral h = positive_integral (\<lambda>x. f x * h x)"
+      by (simp add: \<nu>.positive_integral_cong_measure[symmetric, OF eq[symmetric]])
+         (intro positive_integral_translated_density f h)
+    then have "positive_integral (\<lambda>x. f x * h x) \<noteq> \<omega>"
+      using h(2) by simp
+    then show "(\<lambda>x. f x * h x) -` {\<omega>} \<inter> space M \<in> null_sets"
+      using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage)
+  qed auto
+  then show "AE x. f x \<noteq> \<omega>"
+  proof (rule AE_mp, intro AE_cong)
+    fix x assume "x \<in> space M" from this[THEN fin]
+    show "f x * h x \<noteq> \<omega> \<longrightarrow> f x \<noteq> \<omega>" by auto
+  qed
+next
+  assume AE: "AE x. f x \<noteq> \<omega>"
+  from sigma_finite guess Q .. note Q = this
+  interpret \<nu>: measure_space M \<nu> by fact
+  def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<omega>} | Suc n \<Rightarrow> {.. of_nat (Suc n)}) \<inter> space M"
+  { fix i j have "A i \<inter> Q j \<in> sets M"
+    unfolding A_def using f Q
+    apply (rule_tac Int)
+    by (cases i) (auto intro: measurable_sets[OF f]) }
+  note A_in_sets = this
+  let "?A n" = "case prod_decode n of (i,j) \<Rightarrow> A i \<inter> Q j"
+  show "sigma_finite_measure M \<nu>"
+  proof (default, intro exI conjI subsetI allI)
+    fix x assume "x \<in> range ?A"
+    then obtain n where n: "x = ?A n" by auto
+    then show "x \<in> sets M" using A_in_sets by (cases "prod_decode n") auto
+  next
+    have "(\<Union>i. ?A i) = (\<Union>i j. A i \<inter> Q j)"
+    proof safe
+      fix x i j assume "x \<in> A i" "x \<in> Q j"
+      then show "x \<in> (\<Union>i. case prod_decode i of (i, j) \<Rightarrow> A i \<inter> Q j)"
+        by (intro UN_I[of "prod_encode (i,j)"]) auto
+    qed auto
+    also have "\<dots> = (\<Union>i. A i) \<inter> space M" using Q by auto
+    also have "(\<Union>i. A i) = space M"
+    proof safe
+      fix x assume x: "x \<in> space M"
+      show "x \<in> (\<Union>i. A i)"
+      proof (cases "f x")
+        case infinite then show ?thesis using x unfolding A_def by (auto intro: exI[of _ 0])
+      next
+        case (preal r)
+        with less_\<omega>_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by auto
+        then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"])
+      qed
+    qed (auto simp: A_def)
+    finally show "(\<Union>i. ?A i) = space M" by simp
+  next
+    fix n obtain i j where
+      [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto
+    have "positive_integral (\<lambda>x. f x * indicator (A i \<inter> Q j) x) \<noteq> \<omega>"
+    proof (cases i)
+      case 0
+      have "AE x. f x * indicator (A i \<inter> Q j) x = 0"
+        using AE by (rule AE_mp) (auto intro!: AE_cong simp: A_def `i = 0`)
+      then have "positive_integral (\<lambda>x. f x * indicator (A i \<inter> Q j) x) = 0"
+        using A_in_sets f
+        apply (subst positive_integral_0_iff)
+        apply fast
+        apply (subst (asm) AE_iff_null_set)
+        apply (intro borel_measurable_pinfreal_neq_const)
+        apply fast
+        by simp
+      then show ?thesis by simp
+    next
+      case (Suc n)
+      then have "positive_integral (\<lambda>x. f x * indicator (A i \<inter> Q j) x) \<le>
+        positive_integral (\<lambda>x. of_nat (Suc n) * indicator (Q j) x)"
+        by (auto intro!: positive_integral_mono simp: indicator_def A_def)
+      also have "\<dots> = of_nat (Suc n) * \<mu> (Q j)"
+        using Q by (auto intro!: positive_integral_cmult_indicator)
+      also have "\<dots> < \<omega>"
+        using Q by auto
+      finally show ?thesis by simp
+    qed
+    then show "\<nu> (?A n) \<noteq> \<omega>"
+      using A_in_sets Q eq by auto
+  qed
+qed
+
 section "Radon Nikodym derivative"
 
 definition (in sigma_finite_measure)
   "RN_deriv \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and>
     (\<forall>A \<in> sets M. \<nu> A = positive_integral (\<lambda>x. f x * indicator A x))"
 
+lemma (in sigma_finite_measure) RN_deriv_cong:
+  assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> \<mu>' A = \<mu> A" "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
+  shows "sigma_finite_measure.RN_deriv M \<mu>' \<nu>' x = RN_deriv \<nu> x"
+proof -
+  interpret \<mu>': sigma_finite_measure M \<mu>'
+    using cong(1) by (rule sigma_finite_measure_cong)
+  show ?thesis
+    unfolding RN_deriv_def \<mu>'.RN_deriv_def
+    by (simp add: cong positive_integral_cong_measure[OF cong(1)])
+qed
+
 lemma (in sigma_finite_measure) RN_deriv:
   assumes "measure_space M \<nu>"
   assumes "absolutely_continuous \<nu>"
@@ -821,6 +1132,107 @@
     by (rule someI2_ex) (simp add: `A \<in> sets M`)
 qed
 
+lemma (in sigma_finite_measure) RN_deriv_positive_integral:
+  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
+    and f: "f \<in> borel_measurable M"
+  shows "measure_space.positive_integral M \<nu> f = positive_integral (\<lambda>x. RN_deriv \<nu> x * f x)"
+proof -
+  interpret \<nu>: measure_space M \<nu> by fact
+  have "\<nu>.positive_integral f =
+    measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. RN_deriv \<nu> x * indicator A x)) f"
+    by (intro \<nu>.positive_integral_cong_measure[symmetric] RN_deriv(2)[OF \<nu>, symmetric])
+  also have "\<dots> = positive_integral (\<lambda>x. RN_deriv \<nu> x * f x)"
+    by (intro positive_integral_translated_density RN_deriv[OF \<nu>] f)
+  finally show ?thesis .
+qed
+
+lemma (in sigma_finite_measure) RN_deriv_unique:
+  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
+  and f: "f \<in> borel_measurable M"
+  and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = positive_integral (\<lambda>x. f x * indicator A x)"
+  shows "AE x. f x = RN_deriv \<nu> x"
+proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]])
+  fix A assume A: "A \<in> sets M"
+  show "positive_integral (\<lambda>x. f x * indicator A x) = positive_integral (\<lambda>x. RN_deriv \<nu> x * indicator A x)"
+    unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] ..
+qed
+
+lemma the_inv_into_in:
+  assumes "inj_on f A" and x: "x \<in> f`A"
+  shows "the_inv_into A f x \<in> A"
+  using assms by (auto simp: the_inv_into_f_f)
+
+lemma (in sigma_finite_measure) RN_deriv_vimage:
+  fixes f :: "'b \<Rightarrow> 'a"
+  assumes f: "bij_betw f S (space M)"
+  assumes \<nu>: "measure_space M \<nu>" "absolutely_continuous \<nu>"
+  shows "AE x.
+    sigma_finite_measure.RN_deriv (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>A. \<nu> (f ` A)) (the_inv_into S f x) = RN_deriv \<nu> x"
+proof (rule RN_deriv_unique[OF \<nu>])
+  interpret sf: sigma_finite_measure "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)"
+    using f by (rule sigma_finite_measure_isomorphic)
+  interpret \<nu>: measure_space M \<nu> using \<nu>(1) .
+  have \<nu>': "measure_space (vimage_algebra S f) (\<lambda>A. \<nu> (f ` A))"
+    using f by (rule \<nu>.measure_space_isomorphic)
+  { fix A assume "A \<in> sets M" then have "f ` (f -` A \<inter> S) = A"
+      using sets_into_space f[unfolded bij_betw_def]
+      by (intro image_vimage_inter_eq[where T="space M"]) auto }
+  note A_f = this
+  then have ac: "sf.absolutely_continuous (\<lambda>A. \<nu> (f ` A))"
+    using \<nu>(2) by (auto simp: sf.absolutely_continuous_def absolutely_continuous_def)
+  show "(\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) (the_inv_into S f x)) \<in> borel_measurable M"
+    using sf.RN_deriv(1)[OF \<nu>' ac]
+    unfolding measurable_vimage_iff_inv[OF f] comp_def .
+  fix A assume "A \<in> sets M"
+  then have *: "\<And>x. x \<in> space M \<Longrightarrow> indicator (f -` A \<inter> S) (the_inv_into S f x) = (indicator A x :: pinfreal)"
+    using f[unfolded bij_betw_def]
+    unfolding indicator_def by (auto simp: f_the_inv_into_f the_inv_into_in)
+  have "\<nu> (f ` (f -` A \<inter> S)) = sf.positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) x * indicator (f -` A \<inter> S) x)"
+    using `A \<in> sets M` by (force intro!: sf.RN_deriv(2)[OF \<nu>' ac])
+  also have "\<dots> = positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) (the_inv_into S f x) * indicator A x)"
+    unfolding positive_integral_vimage_inv[OF f]
+    by (simp add: * cong: positive_integral_cong)
+  finally show "\<nu> A = positive_integral (\<lambda>x. sf.RN_deriv (\<lambda>A. \<nu> (f ` A)) (the_inv_into S f x) * indicator A x)"
+    unfolding A_f[OF `A \<in> sets M`] .
+qed
+
+lemma (in sigma_finite_measure) RN_deriv_finite:
+  assumes sfm: "sigma_finite_measure M \<nu>" and ac: "absolutely_continuous \<nu>"
+  shows "AE x. RN_deriv \<nu> x \<noteq> \<omega>"
+proof -
+  interpret \<nu>: sigma_finite_measure M \<nu> by fact
+  have \<nu>: "measure_space M \<nu>" by default
+  from sfm show ?thesis
+    using sigma_finite_iff_density_finite[OF \<nu> RN_deriv[OF \<nu> ac]] by simp
+qed
+
+lemma (in sigma_finite_measure)
+  assumes \<nu>: "sigma_finite_measure M \<nu>" "absolutely_continuous \<nu>"
+    and f: "f \<in> borel_measurable M"
+  shows RN_deriv_integral: "measure_space.integral M \<nu> f = integral (\<lambda>x. real (RN_deriv \<nu> x) * f x)" (is ?integral)
+    and RN_deriv_integrable: "measure_space.integrable M \<nu> f \<longleftrightarrow> integrable (\<lambda>x. real (RN_deriv \<nu> x) * f x)" (is ?integrable)
+proof -
+  interpret \<nu>: sigma_finite_measure M \<nu> by fact
+  have ms: "measure_space M \<nu>" by default
+  have minus_cong: "\<And>A B A' B'::pinfreal. A = A' \<Longrightarrow> B = B' \<Longrightarrow> real A - real B = real A' - real B'" by simp
+  have f': "(\<lambda>x. - f x) \<in> borel_measurable M" using f by auto
+  { fix f :: "'a \<Rightarrow> real" assume "f \<in> borel_measurable M"
+    { fix x assume *: "RN_deriv \<nu> x \<noteq> \<omega>"
+      have "Real (real (RN_deriv \<nu> x)) * Real (f x) = Real (real (RN_deriv \<nu> x) * f x)"
+        by (simp add: mult_le_0_iff)
+      then have "RN_deriv \<nu> x * Real (f x) = Real (real (RN_deriv \<nu> x) * f x)"
+        using * by (simp add: Real_real) }
+    note * = this
+    have "positive_integral (\<lambda>x. RN_deriv \<nu> x * Real (f x)) = positive_integral (\<lambda>x. Real (real (RN_deriv \<nu> x) * f x))"
+      apply (rule positive_integral_cong_AE)
+      apply (rule AE_mp[OF RN_deriv_finite[OF \<nu>]])
+      by (auto intro!: AE_cong simp: *) }
+  with this[OF f] this[OF f'] f f'
+  show ?integral ?integrable
+    unfolding \<nu>.integral_def integral_def \<nu>.integrable_def integrable_def
+    by (auto intro!: RN_deriv(1)[OF ms \<nu>(2)] minus_cong simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
+qed
+
 lemma (in sigma_finite_measure) RN_deriv_singleton:
   assumes "measure_space M \<nu>"
   and ac: "absolutely_continuous \<nu>"
--- a/src/HOL/Probability/Sigma_Algebra.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -23,7 +23,7 @@
   sets :: "'a set set"
 
 locale algebra =
-  fixes M
+  fixes M :: "'a algebra"
   assumes space_closed: "sets M \<subseteq> Pow (space M)"
      and  empty_sets [iff]: "{} \<in> sets M"
      and  compl_sets [intro]: "!!a. a \<in> sets M \<Longrightarrow> space M - a \<in> sets M"
@@ -110,13 +110,6 @@
   assumes countable_nat_UN [intro]:
          "!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
 
-lemma sigma_algebra_cong:
-  fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
-  assumes *: "sigma_algebra M"
-  and cong: "space M = space M'" "sets M = sets M'"
-  shows "sigma_algebra M'"
-using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
-
 lemma countable_UN_eq:
   fixes A :: "'i::countable \<Rightarrow> 'a set"
   shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow>
@@ -222,15 +215,13 @@
   | Compl: "a \<in> sigma_sets sp A \<Longrightarrow> sp - a \<in> sigma_sets sp A"
   | Union: "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Union>i. a i) \<in> sigma_sets sp A"
 
+definition
+  "sigma M = (| space = space M, sets = sigma_sets (space M) (sets M) |)"
 
-definition
-  sigma  where
-  "sigma sp A = (| space = sp, sets = sigma_sets sp A |)"
-
-lemma sets_sigma: "sets (sigma A B) = sigma_sets A B"
+lemma sets_sigma: "sets (sigma M) = sigma_sets (space M) (sets M)"
   unfolding sigma_def by simp
 
-lemma space_sigma [simp]: "space (sigma X B) = X"
+lemma space_sigma [simp]: "space (sigma M) = space M"
   by (simp add: sigma_def)
 
 lemma sigma_sets_top: "sp \<in> sigma_sets sp A"
@@ -242,7 +233,7 @@
 lemma sigma_sets_Un:
   "a \<in> sigma_sets sp A \<Longrightarrow> b \<in> sigma_sets sp A \<Longrightarrow> a \<union> b \<in> sigma_sets sp A"
 apply (simp add: Un_range_binary range_binary_eq)
-apply (rule Union, simp add: binary_def fun_upd_apply)
+apply (rule Union, simp add: binary_def)
 done
 
 lemma sigma_sets_Inter:
@@ -306,15 +297,34 @@
   done
 
 lemma sigma_algebra_sigma:
-     "a \<subseteq> Pow X \<Longrightarrow> sigma_algebra (sigma X a)"
+    "sets M \<subseteq> Pow (space M) \<Longrightarrow> sigma_algebra (sigma M)"
   apply (rule sigma_algebra_sigma_sets)
   apply (auto simp add: sigma_def)
   done
 
 lemma (in sigma_algebra) sigma_subset:
-     "a \<subseteq> sets M ==> sets (sigma (space M) a) \<subseteq> (sets M)"
+    "sets N \<subseteq> sets M \<Longrightarrow> space N = space M \<Longrightarrow> sets (sigma N) \<subseteq> (sets M)"
   by (simp add: sigma_def sigma_sets_subset)
 
+lemma sigma_sets_least_sigma_algebra:
+  assumes "A \<subseteq> Pow S"
+  shows "sigma_sets S A = \<Inter>{B. A \<subseteq> B \<and> sigma_algebra \<lparr>space = S, sets = B\<rparr>}"
+proof safe
+  fix B X assume "A \<subseteq> B" and sa: "sigma_algebra \<lparr> space = S, sets = B \<rparr>"
+    and X: "X \<in> sigma_sets S A"
+  from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `A \<subseteq> B`] X
+  show "X \<in> B" by auto
+next
+  fix X assume "X \<in> \<Inter>{B. A \<subseteq> B \<and> sigma_algebra \<lparr>space = S, sets = B\<rparr>}"
+  then have [intro!]: "\<And>B. A \<subseteq> B \<Longrightarrow> sigma_algebra \<lparr>space = S, sets = B\<rparr> \<Longrightarrow> X \<in> B"
+     by simp
+  have "A \<subseteq> sigma_sets S A" using assms
+    by (auto intro!: sigma_sets.Basic)
+  moreover have "sigma_algebra \<lparr>space = S, sets = sigma_sets S A\<rparr>"
+    using assms by (intro sigma_algebra_sigma_sets[of A]) auto
+  ultimately show "X \<in> sigma_sets S A" by auto
+qed
+
 lemma (in sigma_algebra) restriction_in_sets:
   fixes A :: "nat \<Rightarrow> 'a set"
   assumes "S \<in> sets M"
@@ -340,22 +350,69 @@
   show "(\<Union>i. A i) \<in> sets (restricted_space S)" by simp
 qed
 
+lemma sigma_sets_Int:
+  assumes "A \<in> sigma_sets sp st"
+  shows "op \<inter> A ` sigma_sets sp st = sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+proof (intro equalityI subsetI)
+  fix x assume "x \<in> op \<inter> A ` sigma_sets sp st"
+  then obtain y where "y \<in> sigma_sets sp st" "x = y \<inter> A" by auto
+  then show "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+  proof (induct arbitrary: x)
+    case (Compl a)
+    then show ?case
+      by (force intro!: sigma_sets.Compl simp: Diff_Int_distrib ac_simps)
+  next
+    case (Union a)
+    then show ?case
+      by (auto intro!: sigma_sets.Union
+               simp add: UN_extend_simps simp del: UN_simps)
+  qed (auto intro!: sigma_sets.intros)
+next
+  fix x assume "x \<in> sigma_sets (A \<inter> sp) (op \<inter> A ` st)"
+  then show "x \<in> op \<inter> A ` sigma_sets sp st"
+  proof induct
+    case (Compl a)
+    then obtain x where "a = A \<inter> x" "x \<in> sigma_sets sp st" by auto
+    then show ?case
+      by (force simp add: image_iff intro!: bexI[of _ "sp - x"] sigma_sets.Compl)
+  next
+    case (Union a)
+    then have "\<forall>i. \<exists>x. x \<in> sigma_sets sp st \<and> a i = A \<inter> x"
+      by (auto simp: image_iff Bex_def)
+    from choice[OF this] guess f ..
+    then show ?case
+      by (auto intro!: bexI[of _ "(\<Union>x. f x)"] sigma_sets.Union
+               simp add: image_iff)
+  qed (auto intro!: sigma_sets.intros)
+qed
+
+lemma sigma_sets_single[simp]: "sigma_sets {X} {{X}} = {{}, {X}}"
+proof (intro set_eqI iffI)
+  fix x assume "x \<in> sigma_sets {X} {{X}}"
+  from sigma_sets_into_sp[OF _ this]
+  show "x \<in> {{}, {X}}" by auto
+next
+  fix x assume "x \<in> {{}, {X}}"
+  then show "x \<in> sigma_sets {X} {{X}}"
+    by (auto intro: sigma_sets.Empty sigma_sets_top)
+qed
+
 section {* Measurable functions *}
 
 definition
   "measurable A B = {f \<in> space A -> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
 
 lemma (in sigma_algebra) measurable_sigma:
-  assumes B: "B \<subseteq> Pow X"
-      and f: "f \<in> space M -> X"
-      and ba: "\<And>y. y \<in> B \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
-  shows "f \<in> measurable M (sigma X B)"
+  assumes B: "sets N \<subseteq> Pow (space N)"
+      and f: "f \<in> space M -> space N"
+      and ba: "\<And>y. y \<in> sets N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M (sigma N)"
 proof -
-  have "sigma_sets X B \<subseteq> {y . (f -` y) \<inter> space M \<in> sets M & y \<subseteq> X}"
+  have "sigma_sets (space N) (sets N) \<subseteq> {y . (f -` y) \<inter> space M \<in> sets M & y \<subseteq> space N}"
     proof clarify
       fix x
-      assume "x \<in> sigma_sets X B"
-      thus "f -` x \<inter> space M \<in> sets M \<and> x \<subseteq> X"
+      assume "x \<in> sigma_sets (space N) (sets N)"
+      thus "f -` x \<inter> space M \<in> sets M \<and> x \<subseteq> space N"
         proof induct
           case (Basic a)
           thus ?case
@@ -366,15 +423,14 @@
             by auto
         next
           case (Compl a)
-          have [simp]: "f -` X \<inter> space M = space M"
+          have [simp]: "f -` space N \<inter> space M = space M"
             by (auto simp add: funcset_mem [OF f])
           thus ?case
             by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl)
         next
           case (Union a)
           thus ?case
-            by (simp add: vimage_UN, simp only: UN_extend_simps(4))
-               (blast intro: countable_UN)
+            by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast
         qed
     qed
   thus ?thesis
@@ -397,7 +453,7 @@
    unfolding measurable_def by auto
 
 lemma (in sigma_algebra) measurable_subset:
-     "(\<And>S. S \<in> sets A \<Longrightarrow> S \<subseteq> space A) \<Longrightarrow> measurable M A \<subseteq> measurable M (sigma (space A) (sets A))"
+     "(\<And>S. S \<in> sets A \<Longrightarrow> S \<subseteq> space A) \<Longrightarrow> measurable M A \<subseteq> measurable M (sigma A)"
   by (auto intro: measurable_sigma measurable_sets measurable_space)
 
 lemma measurable_eqI:
@@ -477,7 +533,7 @@
   by (auto simp add: measurable_def)
 
 lemma measurable_up_sigma:
-  "measurable A M \<subseteq> measurable (sigma (space A) (sets A)) M"
+  "measurable A M \<subseteq> measurable (sigma A) M"
   unfolding measurable_def
   by (auto simp: sigma_def intro: sigma_sets.Basic)
 
@@ -495,32 +551,11 @@
     \<Longrightarrow> f \<in> measurable M \<lparr>space = f ` space M, sets = Pow (f ` space M)\<rparr>"
   by (simp add: measurable_def sigma_algebra_Pow) intro_locales
 
-lemma (in sigma_algebra) sigma_algebra_preimages:
-  fixes f :: "'x \<Rightarrow> 'a"
-  assumes "f \<in> A \<rightarrow> space M"
-  shows "sigma_algebra \<lparr> space = A, sets = (\<lambda>M. f -` M \<inter> A) ` sets M \<rparr>"
-    (is "sigma_algebra \<lparr> space = _, sets = ?F ` sets M \<rparr>")
-proof (simp add: sigma_algebra_iff2, safe)
-  show "{} \<in> ?F ` sets M" by blast
-next
-  fix S assume "S \<in> sets M"
-  moreover have "A - ?F S = ?F (space M - S)"
-    using assms by auto
-  ultimately show "A - ?F S \<in> ?F ` sets M"
-    by blast
-next
-  fix S :: "nat \<Rightarrow> 'x set" assume *: "range S \<subseteq> ?F ` sets M"
-  have "\<forall>i. \<exists>b. b \<in> sets M \<and> S i = ?F b"
-  proof safe
-    fix i
-    have "S i \<in> ?F ` sets M" using * by auto
-    then show "\<exists>b. b \<in> sets M \<and> S i = ?F b" by auto
-  qed
-  from choice[OF this] obtain b where b: "range b \<subseteq> sets M" "\<And>i. S i = ?F (b i)"
-    by auto
-  then have "(\<Union>i. S i) = ?F (\<Union>i. b i)" by auto
-  then show "(\<Union>i. S i) \<in> ?F ` sets M" using b(1) by blast
-qed
+lemma (in sigma_algebra) measurable_iff_sigma:
+  assumes "sets E \<subseteq> Pow (space E)" and "f \<in> space M \<rightarrow> space E"
+  shows "f \<in> measurable M (sigma E) \<longleftrightarrow> (\<forall>A\<in>sets E. f -` A \<inter> space M \<in> sets M)"
+  using measurable_sigma[OF assms]
+  by (fastsimp simp: measurable_def sets_sigma intro: sigma_sets.intros)
 
 section "Disjoint families"
 
@@ -544,6 +579,12 @@
      "disjoint_family A \<Longrightarrow> (!!x. B x \<subseteq> A x) \<Longrightarrow> disjoint_family B"
   by (force simp add: disjoint_family_on_def)
 
+lemma disjoint_family_on_bisimulation:
+  assumes "disjoint_family_on f S"
+  and "\<And>n m. n \<in> S \<Longrightarrow> m \<in> S \<Longrightarrow> n \<noteq> m \<Longrightarrow> f n \<inter> f m = {} \<Longrightarrow> g n \<inter> g m = {}"
+  shows "disjoint_family_on g S"
+  using assms unfolding disjoint_family_on_def by auto
+
 lemma disjoint_family_on_mono:
   "A \<subseteq> B \<Longrightarrow> disjoint_family_on f B \<Longrightarrow> disjoint_family_on f A"
   unfolding disjoint_family_on_def by auto
@@ -661,38 +702,163 @@
   "space (vimage_algebra S f) = S"
   by (simp add: vimage_algebra_def)
 
+lemma (in sigma_algebra) sigma_algebra_preimages:
+  fixes f :: "'x \<Rightarrow> 'a"
+  assumes "f \<in> A \<rightarrow> space M"
+  shows "sigma_algebra \<lparr> space = A, sets = (\<lambda>M. f -` M \<inter> A) ` sets M \<rparr>"
+    (is "sigma_algebra \<lparr> space = _, sets = ?F ` sets M \<rparr>")
+proof (simp add: sigma_algebra_iff2, safe)
+  show "{} \<in> ?F ` sets M" by blast
+next
+  fix S assume "S \<in> sets M"
+  moreover have "A - ?F S = ?F (space M - S)"
+    using assms by auto
+  ultimately show "A - ?F S \<in> ?F ` sets M"
+    by blast
+next
+  fix S :: "nat \<Rightarrow> 'x set" assume *: "range S \<subseteq> ?F ` sets M"
+  have "\<forall>i. \<exists>b. b \<in> sets M \<and> S i = ?F b"
+  proof safe
+    fix i
+    have "S i \<in> ?F ` sets M" using * by auto
+    then show "\<exists>b. b \<in> sets M \<and> S i = ?F b" by auto
+  qed
+  from choice[OF this] obtain b where b: "range b \<subseteq> sets M" "\<And>i. S i = ?F (b i)"
+    by auto
+  then have "(\<Union>i. S i) = ?F (\<Union>i. b i)" by auto
+  then show "(\<Union>i. S i) \<in> ?F ` sets M" using b(1) by blast
+qed
+
 lemma (in sigma_algebra) sigma_algebra_vimage:
   fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
   shows "sigma_algebra (vimage_algebra S f)"
-proof
-  fix A assume "A \<in> sets (vimage_algebra S f)"
-  then guess B unfolding in_vimage_algebra ..
-  then show "space (vimage_algebra S f) - A \<in> sets (vimage_algebra S f)"
-    using sets_into_space assms
-    by (auto intro!: bexI[of _ "space M - B"])
-next
-  fix A assume "A \<in> sets (vimage_algebra S f)"
-  then guess A' unfolding in_vimage_algebra .. note A' = this
-  fix B assume "B \<in> sets (vimage_algebra S f)"
-  then guess B' unfolding in_vimage_algebra .. note B' = this
-  then show "A \<union> B \<in> sets (vimage_algebra S f)"
-    using sets_into_space assms A' B'
-    by (auto intro!: bexI[of _ "A' \<union> B'"])
-next
-  fix A::"nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets (vimage_algebra S f)"
-  then have "\<forall>i. \<exists>B. A i = f -` B \<inter> S \<and> B \<in> sets M"
-    by (simp add: subset_eq) blast
-  from this[THEN choice] obtain B
-    where B: "\<And>i. A i = f -` B i \<inter> S" "range B \<subseteq> sets M" by auto
-  show "(\<Union>i. A i) \<in> sets (vimage_algebra S f)"
-    using B by (auto intro!: bexI[of _ "\<Union>i. B i"])
-qed auto
+proof -
+  from sigma_algebra_preimages[OF assms]
+  show ?thesis unfolding vimage_algebra_def by (auto simp: sigma_algebra_iff2)
+qed
 
 lemma (in sigma_algebra) measurable_vimage_algebra:
   fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
   shows "f \<in> measurable (vimage_algebra S f) M"
     unfolding measurable_def using assms by force
 
+lemma (in sigma_algebra) measurable_vimage:
+  fixes g :: "'a \<Rightarrow> 'c" and f :: "'d \<Rightarrow> 'a"
+  assumes "g \<in> measurable M M2" "f \<in> S \<rightarrow> space M"
+  shows "(\<lambda>x. g (f x)) \<in> measurable (vimage_algebra S f) M2"
+proof -
+  note measurable_vimage_algebra[OF assms(2)]
+  from measurable_comp[OF this assms(1)]
+  show ?thesis by (simp add: comp_def)
+qed
+
+lemma (in sigma_algebra) vimage_vimage_inv:
+  assumes f: "bij_betw f S (space M)"
+  assumes [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f (g x) = x" and g: "g \<in> space M \<rightarrow> S"
+  shows "sigma_algebra.vimage_algebra (vimage_algebra S f) (space M) g = M"
+proof -
+  interpret T: sigma_algebra "vimage_algebra S f"
+    using f by (safe intro!: sigma_algebra_vimage bij_betw_imp_funcset)
+
+  have inj: "inj_on f S" and [simp]: "f`S = space M"
+    using f unfolding bij_betw_def by auto
+
+  { fix A assume A: "A \<in> sets M"
+    have "g -` f -` A \<inter> g -` S \<inter> space M = (f \<circ> g) -` A \<inter> space M"
+      using g by auto
+    also have "\<dots> = A"
+      using `A \<in> sets M`[THEN sets_into_space] by auto
+    finally have "g -` f -` A \<inter> g -` S \<inter> space M = A" . }
+  note X = this
+  show ?thesis
+    unfolding T.vimage_algebra_def unfolding vimage_algebra_def
+    by (simp add: image_compose[symmetric] comp_def X cong: image_cong)
+qed
+
+lemma (in sigma_algebra) measurable_vimage_iff:
+  fixes f :: "'b \<Rightarrow> 'a" assumes f: "bij_betw f S (space M)"
+  shows "g \<in> measurable M M' \<longleftrightarrow> (g \<circ> f) \<in> measurable (vimage_algebra S f) M'"
+proof
+  assume "g \<in> measurable M M'"
+  from measurable_vimage[OF this f[THEN bij_betw_imp_funcset]]
+  show "(g \<circ> f) \<in> measurable (vimage_algebra S f) M'" unfolding comp_def .
+next
+  interpret v: sigma_algebra "vimage_algebra S f"
+    using f[THEN bij_betw_imp_funcset] by (rule sigma_algebra_vimage)
+  note f' = f[THEN bij_betw_the_inv_into]
+  assume "g \<circ> f \<in> measurable (vimage_algebra S f) M'"
+  from v.measurable_vimage[OF this, unfolded space_vimage_algebra, OF f'[THEN bij_betw_imp_funcset]]
+  show "g \<in> measurable M M'"
+    using f f'[THEN bij_betw_imp_funcset] f[unfolded bij_betw_def]
+    by (subst (asm) vimage_vimage_inv)
+       (simp_all add: f_the_inv_into_f cong: measurable_cong)
+qed
+
+lemma (in sigma_algebra) measurable_vimage_iff_inv:
+  fixes f :: "'b \<Rightarrow> 'a" assumes f: "bij_betw f S (space M)"
+  shows "g \<in> measurable (vimage_algebra S f) M' \<longleftrightarrow> (g \<circ> the_inv_into S f) \<in> measurable M M'"
+  unfolding measurable_vimage_iff[OF f]
+  using f[unfolded bij_betw_def]
+  by (auto intro!: measurable_cong simp add: the_inv_into_f_f)
+
+lemma sigma_sets_vimage:
+  assumes "f \<in> S' \<rightarrow> S" and "A \<subseteq> Pow S"
+  shows "sigma_sets S' ((\<lambda>X. f -` X \<inter> S') ` A) = (\<lambda>X. f -` X \<inter> S') ` sigma_sets S A"
+proof (intro set_eqI iffI)
+  let ?F = "\<lambda>X. f -` X \<inter> S'"
+  fix X assume "X \<in> sigma_sets S' (?F ` A)"
+  then show "X \<in> ?F ` sigma_sets S A"
+  proof induct
+    case (Basic X) then obtain X' where "X = ?F X'" "X' \<in> A"
+      by auto
+    then show ?case by (auto intro!: sigma_sets.Basic)
+  next
+    case Empty then show ?case
+      by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty)
+  next
+    case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \<in> sigma_sets S A"
+      by auto
+    then have "S - X' \<in> sigma_sets S A"
+      by (auto intro!: sigma_sets.Compl)
+    then show ?case
+      using X assms by (auto intro!: image_eqI[where x="S - X'"])
+  next
+    case (Union F)
+    then have "\<forall>i. \<exists>F'.  F' \<in> sigma_sets S A \<and> F i = f -` F' \<inter> S'"
+      by (auto simp: image_iff Bex_def)
+    from choice[OF this] obtain F' where
+      "\<And>i. F' i \<in> sigma_sets S A" and "\<And>i. F i = f -` F' i \<inter> S'"
+      by auto
+    then show ?case
+      by (auto intro!: sigma_sets.Union image_eqI[where x="\<Union>i. F' i"])
+  qed
+next
+  let ?F = "\<lambda>X. f -` X \<inter> S'"
+  fix X assume "X \<in> ?F ` sigma_sets S A"
+  then obtain X' where "X' \<in> sigma_sets S A" "X = ?F X'" by auto
+  then show "X \<in> sigma_sets S' (?F ` A)"
+  proof (induct arbitrary: X)
+    case (Basic X') then show ?case by (auto intro: sigma_sets.Basic)
+  next
+    case Empty then show ?case by (auto intro: sigma_sets.Empty)
+  next
+    case (Compl X')
+    have "S' - (S' - X) \<in> sigma_sets S' (?F ` A)"
+      apply (rule sigma_sets.Compl)
+      using assms by (auto intro!: Compl.hyps simp: Compl.prems)
+    also have "S' - (S' - X) = X"
+      using assms Compl by auto
+    finally show ?case .
+  next
+    case (Union F)
+    have "(\<Union>i. f -` F i \<inter> S') \<in> sigma_sets S' (?F ` A)"
+      by (intro sigma_sets.Union Union.hyps) simp
+    also have "(\<Union>i. f -` F i \<inter> S') = X"
+      using assms Union by auto
+    finally show ?case .
+  qed
+qed
+
 section {* Conditional space *}
 
 definition (in algebra)
@@ -981,4 +1147,268 @@
     by blast
 qed
 
+section {* Dynkin systems *}
+
+
+locale dynkin_system =
+  fixes M :: "'a algebra"
+  assumes space_closed: "sets M \<subseteq> Pow (space M)"
+    and   space: "space M \<in> sets M"
+    and   compl[intro!]: "\<And>A. A \<in> sets M \<Longrightarrow> space M - A \<in> sets M"
+    and   UN[intro!]: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> sets M
+                           \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+
+lemma (in dynkin_system) sets_into_space: "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
+  using space_closed by auto
+
+lemma (in dynkin_system) empty[intro, simp]: "{} \<in> sets M"
+  using space compl[of "space M"] by simp
+
+lemma (in dynkin_system) diff:
+  assumes sets: "D \<in> sets M" "E \<in> sets M" and "D \<subseteq> E"
+  shows "E - D \<in> sets M"
+proof -
+  let ?f = "\<lambda>x. if x = 0 then D else if x = Suc 0 then space M - E else {}"
+  have "range ?f = {D, space M - E, {}}"
+    by (auto simp: image_iff)
+  moreover have "D \<union> (space M - E) = (\<Union>i. ?f i)"
+    by (auto simp: image_iff split: split_if_asm)
+  moreover
+  then have "disjoint_family ?f" unfolding disjoint_family_on_def
+    using `D \<in> sets M`[THEN sets_into_space] `D \<subseteq> E` by auto
+  ultimately have "space M - (D \<union> (space M - E)) \<in> sets M"
+    using sets by auto
+  also have "space M - (D \<union> (space M - E)) = E - D"
+    using assms sets_into_space by auto
+  finally show ?thesis .
+qed
+
+lemma dynkin_systemI:
+  assumes "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M" "space M \<in> sets M"
+  assumes "\<And> A. A \<in> sets M \<Longrightarrow> space M - A \<in> sets M"
+  assumes "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> sets M
+          \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+  shows "dynkin_system M"
+  using assms by (auto simp: dynkin_system_def)
+
+lemma dynkin_system_trivial:
+  shows "dynkin_system \<lparr> space = A, sets = Pow A \<rparr>"
+  by (rule dynkin_systemI) auto
+
+lemma sigma_algebra_imp_dynkin_system:
+  assumes "sigma_algebra M" shows "dynkin_system M"
+proof -
+  interpret sigma_algebra M by fact
+  show ?thesis using sets_into_space by (fastsimp intro!: dynkin_systemI)
+qed
+
+subsection "Intersection stable algebras"
+
+definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)"
+
+lemma (in algebra) Int_stable: "Int_stable M"
+  unfolding Int_stable_def by auto
+
+lemma (in dynkin_system) sigma_algebra_eq_Int_stable:
+  "sigma_algebra M \<longleftrightarrow> Int_stable M"
+proof
+  assume "sigma_algebra M" then show "Int_stable M"
+    unfolding sigma_algebra_def using algebra.Int_stable by auto
+next
+  assume "Int_stable M"
+  show "sigma_algebra M"
+    unfolding sigma_algebra_disjoint_iff algebra_def
+  proof (intro conjI ballI allI impI)
+    show "sets M \<subseteq> Pow (space M)" using sets_into_space by auto
+  next
+    fix A B assume "A \<in> sets M" "B \<in> sets M"
+    then have "A \<union> B = space M - ((space M - A) \<inter> (space M - B))"
+              "space M - A \<in> sets M" "space M - B \<in> sets M"
+      using sets_into_space by auto
+    then show "A \<union> B \<in> sets M"
+      using `Int_stable M` unfolding Int_stable_def by auto
+  qed auto
+qed
+
+subsection "Smallest Dynkin systems"
+
+definition dynkin :: "'a algebra \<Rightarrow> 'a algebra" where
+  "dynkin M = \<lparr> space = space M,
+     sets =  \<Inter>{D. dynkin_system \<lparr> space = space M, sets = D\<rparr> \<and> sets M \<subseteq> D}\<rparr>"
+
+lemma dynkin_system_dynkin:
+  fixes M :: "'a algebra"
+  assumes "sets M \<subseteq> Pow (space M)"
+  shows "dynkin_system (dynkin M)"
+proof (rule dynkin_systemI)
+  fix A assume "A \<in> sets (dynkin M)"
+  moreover
+  { fix D assume "A \<in> D" and d: "dynkin_system \<lparr> space = space M, sets = D \<rparr>"
+    from dynkin_system.sets_into_space[OF d] `A \<in> D`
+    have "A \<subseteq> space M" by auto }
+  moreover have "{D. dynkin_system \<lparr> space = space M, sets = D\<rparr> \<and> sets M \<subseteq> D} \<noteq> {}"
+    using assms dynkin_system_trivial by fastsimp
+  ultimately show "A \<subseteq> space (dynkin M)"
+    unfolding dynkin_def using assms
+    by simp (metis dynkin_system.sets_into_space in_mono mem_def)
+next
+  show "space (dynkin M) \<in> sets (dynkin M)"
+    unfolding dynkin_def using dynkin_system.space by fastsimp
+next
+  fix A assume "A \<in> sets (dynkin M)"
+  then show "space (dynkin M) - A \<in> sets (dynkin M)"
+    unfolding dynkin_def using dynkin_system.compl by force
+next
+  fix A :: "nat \<Rightarrow> 'a set"
+  assume A: "disjoint_family A" "range A \<subseteq> sets (dynkin M)"
+  show "(\<Union>i. A i) \<in> sets (dynkin M)" unfolding dynkin_def
+  proof (simp, safe)
+    fix D assume "dynkin_system \<lparr>space = space M, sets = D\<rparr>" "sets M \<subseteq> D"
+    with A have "(\<Union>i. A i) \<in> sets \<lparr>space = space M, sets = D\<rparr>"
+      by (intro dynkin_system.UN) (auto simp: dynkin_def)
+    then show "(\<Union>i. A i) \<in> D" by auto
+  qed
+qed
+
+lemma dynkin_Basic[intro]:
+  "A \<in> sets M \<Longrightarrow> A \<in> sets (dynkin M)"
+  unfolding dynkin_def by auto
+
+lemma dynkin_space[simp]:
+  "space (dynkin M) = space M"
+  unfolding dynkin_def by auto
+
+lemma (in dynkin_system) restricted_dynkin_system:
+  assumes "D \<in> sets M"
+  shows "dynkin_system \<lparr> space = space M,
+                         sets = {Q. Q \<subseteq> space M \<and> Q \<inter> D \<in> sets M} \<rparr>"
+proof (rule dynkin_systemI, simp_all)
+  have "space M \<inter> D = D"
+    using `D \<in> sets M` sets_into_space by auto
+  then show "space M \<inter> D \<in> sets M"
+    using `D \<in> sets M` by auto
+next
+  fix A assume "A \<subseteq> space M \<and> A \<inter> D \<in> sets M"
+  moreover have "(space M - A) \<inter> D = (space M - (A \<inter> D)) - (space M - D)"
+    by auto
+  ultimately show "space M - A \<subseteq> space M \<and> (space M - A) \<inter> D \<in> sets M"
+    using  `D \<in> sets M` by (auto intro: diff)
+next
+  fix A :: "nat \<Rightarrow> 'a set"
+  assume "disjoint_family A" "range A \<subseteq> {Q. Q \<subseteq> space M \<and> Q \<inter> D \<in> sets M}"
+  then have "\<And>i. A i \<subseteq> space M" "disjoint_family (\<lambda>i. A i \<inter> D)"
+    "range (\<lambda>i. A i \<inter> D) \<subseteq> sets M" "(\<Union>x. A x) \<inter> D = (\<Union>x. A x \<inter> D)"
+    by ((fastsimp simp: disjoint_family_on_def)+)
+  then show "(\<Union>x. A x) \<subseteq> space M \<and> (\<Union>x. A x) \<inter> D \<in> sets M"
+    by (auto simp del: UN_simps)
+qed
+
+lemma (in dynkin_system) dynkin_subset:
+  fixes N :: "'a algebra"
+  assumes "sets N \<subseteq> sets M"
+  assumes "space N = space M"
+  shows "sets (dynkin N) \<subseteq> sets M"
+proof -
+  have *: "\<lparr>space = space N, sets = sets M\<rparr> = M"
+    unfolding `space N = space M` by simp
+  have "dynkin_system M" by default
+  then have "dynkin_system \<lparr>space = space N, sets = sets M\<rparr>"
+    using assms unfolding * by simp
+  with `sets N \<subseteq> sets M` show ?thesis by (auto simp add: dynkin_def)
+qed
+
+lemma sigma_eq_dynkin:
+  fixes M :: "'a algebra"
+  assumes sets: "sets M \<subseteq> Pow (space M)"
+  assumes "Int_stable M"
+  shows "sigma M = dynkin M"
+proof -
+  have "sets (dynkin M) \<subseteq> sigma_sets (space M) (sets M)"
+    using sigma_algebra_imp_dynkin_system
+    unfolding dynkin_def sigma_def sigma_sets_least_sigma_algebra[OF sets] by auto
+  moreover
+  interpret dynkin_system "dynkin M"
+    using dynkin_system_dynkin[OF sets] .
+  have "sigma_algebra (dynkin M)"
+    unfolding sigma_algebra_eq_Int_stable Int_stable_def
+  proof (intro ballI)
+    fix A B assume "A \<in> sets (dynkin M)" "B \<in> sets (dynkin M)"
+    let "?D E" = "\<lparr> space = space M,
+                    sets = {Q. Q \<subseteq> space M \<and> Q \<inter> E \<in> sets (dynkin M)} \<rparr>"
+    have "sets M \<subseteq> sets (?D B)"
+    proof
+      fix E assume "E \<in> sets M"
+      then have "sets M \<subseteq> sets (?D E)" "E \<in> sets (dynkin M)"
+        using sets_into_space `Int_stable M` by (auto simp: Int_stable_def)
+      then have "sets (dynkin M) \<subseteq> sets (?D E)"
+        using restricted_dynkin_system `E \<in> sets (dynkin M)`
+        by (intro dynkin_system.dynkin_subset) simp_all
+      then have "B \<in> sets (?D E)"
+        using `B \<in> sets (dynkin M)` by auto
+      then have "E \<inter> B \<in> sets (dynkin M)"
+        by (subst Int_commute) simp
+      then show "E \<in> sets (?D B)"
+        using sets `E \<in> sets M` by auto
+    qed
+    then have "sets (dynkin M) \<subseteq> sets (?D B)"
+      using restricted_dynkin_system `B \<in> sets (dynkin M)`
+      by (intro dynkin_system.dynkin_subset) simp_all
+    then show "A \<inter> B \<in> sets (dynkin M)"
+      using `A \<in> sets (dynkin M)` sets_into_space by auto
+  qed
+  from sigma_algebra.sigma_sets_subset[OF this, of "sets M"]
+  have "sigma_sets (space M) (sets M) \<subseteq> sets (dynkin M)" by auto
+  ultimately have "sigma_sets (space M) (sets M) = sets (dynkin M)" by auto
+  then show ?thesis
+    by (intro algebra.equality) (simp_all add: sigma_def)
+qed
+
+lemma (in dynkin_system) dynkin_idem:
+  "dynkin M = M"
+proof -
+  have "sets (dynkin M) = sets M"
+  proof
+    show "sets M \<subseteq> sets (dynkin M)"
+      using dynkin_Basic by auto
+    show "sets (dynkin M) \<subseteq> sets M"
+      by (intro dynkin_subset) auto
+  qed
+  then show ?thesis
+    by (auto intro!: algebra.equality)
+qed
+
+lemma (in dynkin_system) dynkin_lemma:
+  fixes E :: "'a algebra"
+  assumes "Int_stable E" and E: "sets E \<subseteq> sets M" "space E = space M"
+  and "sets M \<subseteq> sets (sigma E)"
+  shows "sigma E = M"
+proof -
+  have "sets E \<subseteq> Pow (space E)"
+    using E sets_into_space by auto
+  then have "sigma E = dynkin E"
+    using `Int_stable E` by (rule sigma_eq_dynkin)
+  moreover then have "sets (dynkin E) = sets M"
+    using assms dynkin_subset[OF E] by simp
+  ultimately show ?thesis
+    using E by simp
+qed
+
+locale finite_sigma_algebra = sigma_algebra +
+  assumes finite_space: "finite (space M)"
+  and sets_eq_Pow[simp]: "sets M = Pow (space M)"
+
+lemma (in finite_sigma_algebra) sets_image_space_eq_Pow:
+  "sets (image_space X) = Pow (space (image_space X))"
+proof safe
+  fix x S assume "S \<in> sets (image_space X)" "x \<in> S"
+  then show "x \<in> space (image_space X)"
+    using sets_into_space by (auto intro!: imageI simp: image_space_def)
+next
+  fix S assume "S \<subseteq> space (image_space X)"
+  then obtain S' where "S = X`S'" "S'\<in>sets M"
+    by (auto simp: subset_image_iff sets_eq_Pow image_space_def)
+  then show "S \<in> sets (image_space X)"
+    by (auto simp: image_space_def)
+qed
+
 end
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Wed Dec 01 06:50:54 2010 -0800
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -2,12 +2,6 @@
 imports Information
 begin
 
-lemma finite_information_spaceI:
-  "\<lbrakk> finite_measure_space M \<mu> ; \<mu> (space M) = 1 ; 1 < b \<rbrakk> \<Longrightarrow> finite_information_space M \<mu> b"
-  unfolding finite_information_space_def finite_measure_space_def finite_measure_space_axioms_def
-    finite_prob_space_def prob_space_def prob_space_axioms_def finite_information_space_axioms_def
-  by auto
-
 locale finite_space =
   fixes S :: "'a set"
   assumes finite[simp]: "finite S"
@@ -21,19 +15,17 @@
   and sets_M[simp]: "sets M = Pow S"
   by (simp_all add: M_def)
 
-sublocale finite_space \<subseteq> finite_information_space M \<mu> 2
-proof (rule finite_information_spaceI)
-  let ?measure = "\<lambda>s::'a set. real (card s) / real (card S)"
+sublocale finite_space \<subseteq> finite_measure_space M \<mu>
+proof (rule finite_measure_spaceI)
+  fix A B :: "'a set" assume "A \<inter> B = {}" "A \<subseteq> space M" "B \<subseteq> space M"
+  then show "\<mu> (A \<union> B) = \<mu> A + \<mu> B"
+    by (simp add: inverse_eq_divide field_simps Real_real
+                  divide_le_0_iff zero_le_divide_iff
+                  card_Un_disjoint finite_subset[OF _ finite])
+qed auto
 
-  show "finite_measure_space M \<mu>"
-  proof (rule finite_measure_spaceI)
-    fix A B :: "'a set" assume "A \<inter> B = {}" "A \<subseteq> space M" "B \<subseteq> space M"
-    then show "\<mu> (A \<union> B) = \<mu> A + \<mu> B"
-      by (simp add: inverse_eq_divide field_simps Real_real
-                    divide_le_0_iff zero_le_divide_iff
-                    card_Un_disjoint finite_subset[OF _ finite])
-  qed auto
-qed simp_all
+sublocale finite_space \<subseteq> information_space M \<mu> 2
+  by default simp_all
 
 lemma set_of_list_extend:
   "{xs. length xs = Suc n \<and> (\<forall>x\<in>set xs. x \<in> A)} =
@@ -477,21 +469,19 @@
 qed
 
 notation (in dining_cryptographers_space)
-  finite_mutual_information ("\<I>'( _ ; _ ')")
+  mutual_information_Pow ("\<I>'( _ ; _ ')")
 
 notation (in dining_cryptographers_space)
-  finite_entropy ("\<H>'( _ ')")
+  entropy_Pow ("\<H>'( _ ')")
 
 notation (in dining_cryptographers_space)
-  finite_conditional_entropy ("\<H>'( _ | _ ')")
+  conditional_entropy_Pow ("\<H>'( _ | _ ')")
 
 theorem (in dining_cryptographers_space)
   "\<I>( inversion ; payer ) = 0"
 proof -
   have n: "0 < n" using n_gt_3 by auto
-
   have lists: "{xs. length xs = n} \<noteq> {}" using Ex_list_of_length by auto
-
   have card_image_inversion:
     "real (card (inversion ` dc_crypto)) = 2^n / 2"
     unfolding card_image_inversion using `0 < n` by (cases n) auto
@@ -502,7 +492,7 @@
 
   { have "\<H>(inversion|payer) =
         - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. real (?dIP {(x, z)}) * log 2 (real (?dIP {(x, z)}) / real (?dP {z}))))"
-      unfolding conditional_entropy_eq
+      unfolding conditional_entropy_eq[OF simple_function_finite simple_function_finite]
       by (simp add: image_payer_dc_crypto setsum_Sigma)
     also have "... =
         - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. 2 / (real n * 2^n) * (1 - real n)))"
@@ -538,7 +528,7 @@
     finally have "\<H>(inversion|payer) = real n - 1" . }
   moreover
   { have "\<H>(inversion) = - (\<Sum>x \<in> inversion`dc_crypto. real (?dI {x}) * log 2 (real (?dI {x})))"
-      unfolding entropy_eq by simp
+      unfolding entropy_eq[OF simple_function_finite] by simp
     also have "... = - (\<Sum>x \<in> inversion`dc_crypto. 2 * (1 - real n) / 2^n)"
       unfolding neg_equal_iff_equal
     proof (rule setsum_cong[OF refl])
@@ -556,7 +546,7 @@
     finally have "\<H>(inversion) = real n - 1" .
   }
   ultimately show ?thesis
-    unfolding mutual_information_eq_entropy_conditional_entropy
+    unfolding mutual_information_eq_entropy_conditional_entropy[OF simple_function_finite simple_function_finite]
     by simp
 qed
 
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Wed Dec 01 19:20:30 2010 +0100
@@ -0,0 +1,540 @@
+(* Author: Johannes Hoelzl, TU Muenchen *)
+
+header {* Formalization of a Countermeasure by Koepf & Duermuth 2009 *}
+
+theory Koepf_Duermuth_Countermeasure
+  imports Information Permutation
+begin
+
+lemma
+  fixes p u :: "'a \<Rightarrow> real"
+  assumes "1 < b"
+  assumes sum: "(\<Sum>i\<in>S. p i) = (\<Sum>i\<in>S. u i)"
+  and pos: "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> p i" "\<And>i. i \<in> S \<Longrightarrow> 0 \<le> u i"
+  and cont: "\<And>i. i \<in> S \<Longrightarrow> 0 < p i \<Longrightarrow> 0 < u i"
+  shows "(\<Sum>i\<in>S. p i * log b (u i)) \<le> (\<Sum>i\<in>S. p i * log b (p i))"
+proof -
+  have "(\<Sum>i\<in>S. p i * ln (u i) - p i * ln (p i)) \<le> (\<Sum>i\<in>S. u i - p i)"
+  proof (intro setsum_mono)
+    fix i assume [intro, simp]: "i \<in> S"
+    show "p i * ln (u i) - p i * ln (p i) \<le> u i - p i"
+    proof cases
+      assume "p i = 0" with pos[of i] show ?thesis by simp
+    next
+      assume "p i \<noteq> 0"
+      then have "0 < p i" "0 < u i" using pos[of i] cont[of i] by auto
+      then have *: "0 < u i / p i" by (blast intro: divide_pos_pos cont)
+      from `0 < p i` `0 < u i`
+      have "p i * ln (u i) - p i * ln (p i) = p i * ln (u i / p i)"
+        by (simp add: ln_div field_simps)
+      also have "\<dots> \<le> p i * (u i / p i - 1)"
+        using exp_ge_add_one_self[of "ln (u i / p i)"] pos[of i] exp_ln[OF *]
+        by (auto intro!: mult_left_mono)
+      also have "\<dots> = u i - p i"
+        using `p i \<noteq> 0` by (simp add: field_simps)
+      finally show ?thesis by simp
+    qed
+  qed
+  also have "\<dots> = 0" using sum by (simp add: setsum_subtractf)
+  finally show ?thesis using `1 < b` unfolding log_def setsum_subtractf
+    by (auto intro!: divide_right_mono
+             simp: times_divide_eq_right setsum_divide_distrib[symmetric])
+qed
+
+definition (in prob_space)
+  "ordered_variable_partition X = (SOME f. bij_betw f {0..<card (X`space M)} (X`space M) \<and>
+    (\<forall>i<card (X`space M). \<forall>j<card (X`space M). i \<le> j \<longrightarrow> distribution X {f i} \<le> distribution X {f j}))"
+
+lemma ex_ordered_bij_betw_nat_finite:
+  fixes order :: "nat \<Rightarrow> 'a\<Colon>linorder"
+  assumes "finite S"
+  shows "\<exists>f. bij_betw f {0..<card S} S \<and> (\<forall>i<card S. \<forall>j<card S. i \<le> j \<longrightarrow> order (f i) \<le> order (f j))"
+proof -
+  from ex_bij_betw_nat_finite [OF `finite S`] guess f .. note f = this
+  let ?xs = "sort_key order (map f [0 ..< card S])"
+
+  have "?xs <~~> map f [0 ..< card S]"
+    unfolding multiset_of_eq_perm[symmetric] by (rule multiset_of_sort)
+  from permutation_Ex_bij [OF this]
+  obtain g where g: "bij_betw g {0..<card S} {0..<card S}" and
+    map: "\<And>i. i<card S \<Longrightarrow> ?xs ! i = map f [0 ..< card S] ! g i"
+    by (auto simp: atLeast0LessThan)
+
+  { fix i assume "i < card S"
+    then have "g i < card S" using g by (auto simp: bij_betw_def)
+    with map [OF `i < card S`] have "f (g i) = ?xs ! i" by simp }
+  note this[simp]
+
+  show ?thesis
+  proof (intro exI allI conjI impI)
+    show "bij_betw (f\<circ>g) {0..<card S} S"
+      using g f by (rule bij_betw_trans)
+    fix i j assume [simp]: "i < card S" "j < card S" "i \<le> j"
+    from sorted_nth_mono[of "map order ?xs" i j]
+    show "order ((f\<circ>g) i) \<le> order ((f\<circ>g) j)" by simp
+  qed
+qed
+
+lemma (in prob_space)
+  assumes "finite (X`space M)"
+  shows "bij_betw (ordered_variable_partition X) {0..<card (X`space M)} (X`space M)"
+  and "\<And>i j. \<lbrakk> i < card (X`space M) ; j < card (X`space M) ; i \<le> j \<rbrakk> \<Longrightarrow>
+    distribution X {ordered_variable_partition X i} \<le> distribution X {ordered_variable_partition X j}"
+proof -
+  
+qed
+
+definition (in prob_space)
+  "order_distribution X i = real (distribution X {ordered_variable_partition X i})"
+
+definition (in prob_space)
+  "guessing_entropy b X = (\<Sum>i<card(X`space M). real i * log b (order_distribution X i))"
+
+abbreviation (in finite_information_space)
+  finite_guessing_entropy ("\<G>'(_')") where
+  "\<G>(X) \<equiv> guessing_entropy b X"
+
+
+
+lemma zero_notin_Suc_image[simp]: "0 \<notin> Suc ` A"
+  by auto
+
+definition extensional :: "'b \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'b) set" where
+  "extensional d A = {f. \<forall>x. x \<notin> A \<longrightarrow> f x = d}"
+
+lemma extensional_empty[simp]: "extensional d {} = {\<lambda>x. d}"
+  unfolding extensional_def by (simp add: expand_set_eq expand_fun_eq)
+
+lemma funset_eq_UN_fun_upd_I:
+  assumes *: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f(a := d) \<in> F A"
+  and **: "\<And>f. f \<in> F (insert a A) \<Longrightarrow> f a \<in> G (f(a:=d))"
+  and ***: "\<And>f x. \<lbrakk> f \<in> F A ; x \<in> G f \<rbrakk> \<Longrightarrow> f(a:=x) \<in> F (insert a A)"
+  shows "F (insert a A) = (\<Union>f\<in>F A. fun_upd f a ` (G f))"
+proof safe
+  fix f assume f: "f \<in> F (insert a A)"
+  show "f \<in> (\<Union>f\<in>F A. fun_upd f a ` G f)"
+  proof (rule UN_I[of "f(a := d)"])
+    show "f(a := d) \<in> F A" using *[OF f] .
+    show "f \<in> fun_upd (f(a:=d)) a ` G (f(a:=d))"
+    proof (rule image_eqI[of _ _ "f a"])
+      show "f a \<in> G (f(a := d))" using **[OF f] .
+    qed simp
+  qed
+next
+  fix f x assume "f \<in> F A" "x \<in> G f"
+  from ***[OF this] show "f(a := x) \<in> F (insert a A)" .
+qed
+
+lemma extensional_insert[simp]:
+  assumes "a \<notin> A"
+  shows "extensional d (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional d A \<inter> (A \<rightarrow> B). (\<lambda>b. f(a := b)) ` B)"
+  apply (rule funset_eq_UN_fun_upd_I)
+  using assms
+  by (auto intro!: inj_onI dest: inj_onD split: split_if_asm simp: extensional_def)
+
+lemma finite_extensional_funcset[simp, intro]:
+  assumes "finite A" "finite B"
+  shows "finite (extensional d A \<inter> (A \<rightarrow> B))"
+  using assms by induct auto
+
+lemma fun_upd_eq_iff: "f(a := b) = g(a := b') \<longleftrightarrow> b = b' \<and> f(a := d) = g(a := d)"
+  by (auto simp: expand_fun_eq)
+
+lemma card_funcset:
+  assumes "finite A" "finite B"
+  shows "card (extensional d A \<inter> (A \<rightarrow> B)) = card B ^ card A"
+using `finite A` proof induct
+  case (insert a A) thus ?case unfolding extensional_insert[OF `a \<notin> A`]
+  proof (subst card_UN_disjoint, safe, simp_all)
+    show "finite (extensional d A \<inter> (A \<rightarrow> B))" "\<And>f. finite (fun_upd f a ` B)"
+      using `finite B` `finite A` by simp_all
+  next
+    fix f g b b' assume "f \<noteq> g" and eq: "f(a := b) = g(a := b')" and
+      ext: "f \<in> extensional d A" "g \<in> extensional d A"
+    have "f a = d" "g a = d"
+      using ext `a \<notin> A` by (auto simp add: extensional_def)
+    with `f \<noteq> g` eq show False unfolding fun_upd_eq_iff[of _ _ b _ _ d]
+      by (auto simp: fun_upd_idem fun_upd_eq_iff)
+  next
+    { fix f assume "f \<in> extensional d A \<inter> (A \<rightarrow> B)"
+      have "card (fun_upd f a ` B) = card B"
+      proof (auto intro!: card_image inj_onI)
+        fix b b' assume "f(a := b) = f(a := b')"
+        from fun_upd_eq_iff[THEN iffD1, OF this] show "b = b'" by simp
+      qed }
+    then show "(\<Sum>i\<in>extensional d A \<inter> (A \<rightarrow> B). card (fun_upd i a ` B)) = card B * card B ^ card A"
+      using insert by simp
+  qed
+qed simp
+
+lemma set_of_list_extend:
+  "{xs. length xs = Suc n \<and> (\<forall>x\<in>set xs. x \<in> A)} =
+    (\<lambda>(xs, n). n#xs) ` ({xs. length xs = n \<and> (\<forall>x\<in>set xs. x \<in> A)} \<times> A)"
+  by (auto simp: length_Suc_conv)
+
+lemma
+  assumes "finite A"
+  shows finite_lists:
+    "finite {xs. length xs = n \<and> (\<forall>x\<in>set xs. x \<in> A)}" (is "finite (?lists n)")
+  and card_list_length:
+    "card {xs. length xs = n \<and> (\<forall>x\<in>set xs. x \<in> A)} = (card A)^n"
+proof -
+  from `finite A`
+  have "(card {xs. length xs = n \<and> (\<forall>x\<in>set xs. x \<in> A)} = (card A)^n) \<and>
+         finite {xs. length xs = n \<and> (\<forall>x\<in>set xs. x \<in> A)}"
+  proof (induct n)
+    case (Suc n)
+    moreover note set_of_list_extend[of n A]
+    moreover have "inj_on (\<lambda>(xs, n). n#xs) (?lists n \<times> A)"
+      by (auto intro!: inj_onI)
+    ultimately show ?case using assms by (auto simp: card_image)
+  qed (simp cong: conj_cong)
+  then show "finite (?lists n)" "card (?lists n) = card A ^ n"
+    by auto
+qed
+
+locale finite_information =
+  fixes \<Omega> :: "'a set"
+    and p :: "'a \<Rightarrow> real"
+    and b :: real
+  assumes finite_space[simp, intro]: "finite \<Omega>"
+  and p_sums_1[simp]: "(\<Sum>\<omega>\<in>\<Omega>. p \<omega>) = 1"
+  and positive_p[simp, intro]: "\<And>x. 0 \<le> p x"
+  and b_gt_1[simp, intro]: "1 < b"
+
+lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
+   by (auto intro!: setsum_nonneg)
+
+sublocale finite_information \<subseteq> finite_information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega> \<rparr>" "\<lambda>S. Real (setsum p S)" b
+proof -
+  show "finite_information_space \<lparr> space = \<Omega>, sets = Pow \<Omega> \<rparr> (\<lambda>S. Real (setsum p S)) b"
+    unfolding finite_information_space_def finite_information_space_axioms_def
+    unfolding finite_prob_space_def prob_space_def prob_space_axioms_def
+    unfolding finite_measure_space_def finite_measure_space_axioms_def
+    by (force intro!: sigma_algebra.finite_additivity_sufficient
+              simp: additive_def sigma_algebra_Pow positive_def Real_eq_Real
+                    setsum.union_disjoint finite_subset)
+qed
+
+lemma (in prob_space) prob_space_subalgebra:
+  assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
+  shows "prob_space (M\<lparr> sets := N \<rparr>) \<mu>" sorry
+
+lemma (in measure_space) measure_space_subalgebra:
+  assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
+  shows "measure_space (M\<lparr> sets := N \<rparr>) \<mu>" sorry
+
+locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b
+    for b :: real
+    and keys :: "'key set" and K :: "'key \<Rightarrow> real"
+    and messages :: "'message set" and M :: "'message \<Rightarrow> real" +
+  fixes observe :: "'key \<Rightarrow> 'message \<Rightarrow> 'observation"
+    and n :: nat
+begin
+
+definition msgs :: "('key \<times> 'message list) set" where
+  "msgs = keys \<times> {ms. length ms = n \<and> (\<forall>M\<in>set ms. M \<in> messages)}"
+
+definition P :: "('key \<times> 'message list) \<Rightarrow> real" where
+  "P = (\<lambda>(k, ms). K k * (\<Prod>i<length ms. M (ms ! i)))"
+
+end
+
+sublocale koepf_duermuth \<subseteq> finite_information msgs P b
+proof default
+  show "finite msgs" unfolding msgs_def
+    using finite_lists[OF M.finite_space, of n]
+    by auto
+
+  have [simp]: "\<And>A. inj_on (\<lambda>(xs, n). n # xs) A" by (force intro!: inj_onI)
+
+  note setsum_right_distrib[symmetric, simp]
+  note setsum_left_distrib[symmetric, simp]
+  note setsum_cartesian_product'[simp]
+
+  have "(\<Sum>ms | length ms = n \<and> (\<forall>M\<in>set ms. M \<in> messages). \<Prod>x<length ms. M (ms ! x)) = 1"
+  proof (induct n)
+    case 0 then show ?case by (simp cong: conj_cong)
+  next
+    case (Suc n)
+    then show ?case
+      by (simp add: comp_def set_of_list_extend
+                    lessThan_eq_Suc_image setsum_reindex setprod_reindex)
+  qed
+  then show "setsum P msgs = 1"
+    unfolding msgs_def P_def by simp
+
+  fix x
+  have "\<And> A f. 0 \<le> (\<Prod>x\<in>A. M (f x))" by (auto simp: setprod_nonneg)
+  then show "0 \<le> P x"
+    unfolding P_def by (auto split: prod.split simp: zero_le_mult_iff)
+qed auto
+
+lemma SIGMA_image_vimage:
+  "snd ` (SIGMA x:f`X. f -` {x} \<inter> X) = X"
+  by (auto simp: image_iff)
+
+lemma zero_le_eq_True: "0 \<le> (x::pinfreal) \<longleftrightarrow> True" by simp
+
+lemma Real_setprod:
+  assumes"\<And>i. i\<in>X \<Longrightarrow> 0 \<le> f i"
+  shows "(\<Prod>i\<in>X. Real (f i)) = Real (\<Prod>i\<in>X. f i)"
+proof cases
+  assume "finite X"
+  from this assms show ?thesis by induct (auto simp: mult_le_0_iff)
+qed simp
+
+lemma inj_Cons[simp]: "\<And>X. inj_on (\<lambda>(xs, x). x#xs) X" by (auto intro!: inj_onI)
+
+lemma setprod_setsum_distrib_lists:
+  fixes P and S :: "'a set" and f :: "'a \<Rightarrow> _::semiring_0" assumes "finite S"
+  shows "(\<Sum>ms\<in>{ms. length ms = n \<and> set ms \<subseteq> S \<and> (\<forall>i<n. P i (ms!i))}. \<Prod>x<n. f (ms ! x)) =
+         (\<Prod>i<n. \<Sum>m\<in>{m\<in>S. P i m}. f m)"
+proof (induct n arbitrary: P)
+  case 0 then show ?case by (simp cong: conj_cong)
+next
+  case (Suc n)
+  have *: "{ms. length ms = Suc n \<and> set ms \<subseteq> S \<and> (\<forall>i<Suc n. P i (ms ! i))} =
+    (\<lambda>(xs, x). x#xs) ` ({ms. length ms = n \<and> set ms \<subseteq> S \<and> (\<forall>i<n. P (Suc i) (ms ! i))} \<times> {m\<in>S. P 0 m})"
+    apply (auto simp: image_iff length_Suc_conv)
+    apply force+
+    apply (case_tac i)
+    by force+
+  show ?case unfolding *
+    using Suc[of "\<lambda>i. P (Suc i)"]
+    by (simp add: setsum_reindex split_conv setsum_cartesian_product'
+      lessThan_eq_Suc_image setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
+qed
+
+context koepf_duermuth
+begin
+
+definition observations :: "'observation set" where
+  "observations = (\<Union>f\<in>observe ` keys. f ` messages)"
+
+lemma finite_observations[simp, intro]: "finite observations"
+  unfolding observations_def by auto
+
+definition OB :: "'key \<times> 'message list \<Rightarrow> 'observation list" where
+  "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)"
+
+lemma card_T_bound: "card ((t\<circ>OB)`msgs) \<le> (n+1)^card observations"
+proof -
+  have "(t\<circ>OB)`msgs \<subseteq> extensional 0 observations \<inter> (observations \<rightarrow> {.. n})"
+    unfolding observations_def extensional_def OB_def msgs_def
+    by (auto simp add: t_def comp_def image_iff)
+  with finite_extensional_funcset
+  have "card ((t\<circ>OB)`msgs) \<le> card (extensional 0 observations \<inter> (observations \<rightarrow> {.. n}))"
+    by (rule card_mono) auto
+  also have "\<dots> = (n + 1) ^ card observations"
+    by (subst card_funcset) auto
+  finally show ?thesis .
+qed
+
+abbreviation
+ "p A \<equiv> setsum P A"
+
+abbreviation probability ("\<P>'(_') _") where
+ "\<P>(X) x \<equiv> real (distribution X x)"
+
+abbreviation joint_probability ("\<P>'(_, _') _") where
+ "\<P>(X, Y) x \<equiv> real (joint_distribution X Y x)"
+
+abbreviation conditional_probability ("\<P>'(_|_') _") where
+ "\<P>(X|Y) x \<equiv> \<P>(X, Y) x / \<P>(Y) (snd`x)"
+
+notation
+  finite_entropy ("\<H>'( _ ')")
+
+notation
+  finite_conditional_entropy ("\<H>'( _ | _ ')")
+
+notation
+  finite_mutual_information ("\<I>'( _ ; _ ')")
+
+lemma t_eq_imp_bij_func:
+  assumes "t xs = t ys"
+  shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))"
+proof -
+  have "count (multiset_of xs) = count (multiset_of ys)"
+    using assms by (simp add: expand_fun_eq count_multiset_of t_def)
+  then have "xs <~~> ys" unfolding multiset_of_eq_perm count_inject .
+  then show ?thesis by (rule permutation_Ex_func)
+qed
+
+lemma \<P>_k: assumes "k \<in> keys" shows "\<P>(fst) {k} = K k"
+proof -
+  from assms have *:
+      "fst -` {k} \<inter> msgs = {k}\<times>{ms. length ms = n \<and> (\<forall>M\<in>set ms. M \<in> messages)}"
+    unfolding msgs_def by auto
+  show "\<P>(fst) {k} = K k" unfolding distribution_def
+    apply (simp add: *) unfolding P_def
+    apply (simp add: setsum_cartesian_product')
+    using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"]
+    by (simp add: setsum_right_distrib[symmetric] subset_eq setprod_1)
+qed
+
+lemma fst_image_msgs[simp]: "fst`msgs = keys"
+proof -
+  from M.not_empty obtain m where "m \<in> messages" by auto
+  then have *: "{m. length m = n \<and> (\<forall>x\<in>set m. x\<in>messages)} \<noteq> {}"
+    by (auto intro!: exI[of _ "replicate n m"])
+  then show ?thesis
+    unfolding msgs_def fst_image_times if_not_P[OF *] by simp
+qed
+
+lemma ce_OB_eq_ce_t: "\<H>(fst | OB) = \<H>(fst | t\<circ>OB)"
+proof -
+  txt {* Lemma 2 *}
+
+  { fix k obs obs'
+    assume "k \<in> keys" "K k \<noteq> 0" and obs': "obs' \<in> OB ` msgs" and obs: "obs \<in> OB ` msgs"
+    assume "t obs = t obs'"
+    from t_eq_imp_bij_func[OF this]
+    obtain t_f where "bij_betw t_f {..<n} {..<n}" and
+      obs_t_f: "\<And>i. i<n \<Longrightarrow> obs!i = obs' ! t_f i"
+      using obs obs' unfolding OB_def msgs_def by auto
+    then have t_f: "inj_on t_f {..<n}" "{..<n} = t_f`{..<n}" unfolding bij_betw_def by auto
+
+    { fix obs assume "obs \<in> OB`msgs"
+      then have **: "\<And>ms. length ms = n \<Longrightarrow> OB (k, ms) = obs \<longleftrightarrow> (\<forall>i<n. observe k (ms!i) = obs ! i)"
+        unfolding OB_def msgs_def by (simp add: image_iff list_eq_iff_nth_eq)
+
+      have "\<P>(OB, fst) {(obs, k)} / K k =
+          p ({k}\<times>{ms. (k,ms) \<in> msgs \<and> OB (k,ms) = obs}) / K k"
+        unfolding distribution_def by (auto intro!: arg_cong[where f=p])
+      also have "\<dots> =
+          (\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)"
+        unfolding P_def using `K k \<noteq> 0` `k \<in> keys`
+        apply (simp add: setsum_cartesian_product' setsum_divide_distrib msgs_def ** cong: conj_cong)
+        apply (subst setprod_setsum_distrib_lists[OF M.finite_space, unfolded subset_eq]) ..
+      finally have "\<P>(OB, fst) {(obs, k)} / K k =
+            (\<Prod>i<n. \<Sum>m\<in>{m\<in>messages. observe k m = obs ! i}. M m)" . }
+    note * = this
+
+    have "\<P>(OB, fst) {(obs, k)} / K k = \<P>(OB, fst) {(obs', k)} / K k"
+      unfolding *[OF obs] *[OF obs']
+      using t_f(1) obs_t_f by (subst (2) t_f(2)) (simp add: setprod_reindex)
+    then have "\<P>(OB, fst) {(obs, k)} = \<P>(OB, fst) {(obs', k)}"
+      using `K k \<noteq> 0` by auto }
+  note t_eq_imp = this
+
+  let "?S obs" = "t -`{t obs} \<inter> OB`msgs"
+  { fix k obs assume "k \<in> keys" "K k \<noteq> 0" and obs: "obs \<in> OB`msgs"
+    have *: "((\<lambda>x. (t (OB x), fst x)) -` {(t obs, k)} \<inter> msgs) =
+      (\<Union>obs'\<in>?S obs. ((\<lambda>x. (OB x, fst x)) -` {(obs', k)} \<inter> msgs))" by auto
+    have df: "disjoint_family_on (\<lambda>obs'. (\<lambda>x. (OB x, fst x)) -` {(obs', k)} \<inter> msgs) (?S obs)"
+      unfolding disjoint_family_on_def by auto
+    have "\<P>(t\<circ>OB, fst) {(t obs, k)} = (\<Sum>obs'\<in>?S obs. \<P>(OB, fst) {(obs', k)})"
+      unfolding distribution_def comp_def
+      using real_finite_measure_finite_Union[OF _ df]
+      by (force simp add: * intro!: setsum_nonneg)
+    also have "(\<Sum>obs'\<in>?S obs. \<P>(OB, fst) {(obs', k)}) = real (card (?S obs)) * \<P>(OB, fst) {(obs, k)}"
+      by (simp add: t_eq_imp[OF `k \<in> keys` `K k \<noteq> 0` obs] real_eq_of_nat)
+    finally have "\<P>(t\<circ>OB, fst) {(t obs, k)} = real (card (?S obs)) * \<P>(OB, fst) {(obs, k)}" .}
+  note P_t_eq_P_OB = this
+
+  { fix k obs assume "k \<in> keys" and obs: "obs \<in> OB`msgs"
+    have "\<P>(t\<circ>OB | fst) {(t obs, k)} =
+      real (card (t -` {t obs} \<inter> OB ` msgs)) * \<P>(OB | fst) {(obs, k)}"
+      using \<P>_k[OF `k \<in> keys`] P_t_eq_P_OB[OF `k \<in> keys` _ obs] by auto }
+  note CP_t_K = this
+
+  { fix k obs assume "k \<in> keys" and obs: "obs \<in> OB`msgs"
+    then have "t -`{t obs} \<inter> OB`msgs \<noteq> {}" (is "?S \<noteq> {}") by auto
+    then have "real (card ?S) \<noteq> 0" by auto
+
+    have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / \<P>(t\<circ>OB) {t obs}"
+      using real_distribution_order'[of fst k "t\<circ>OB" "t obs"]
+      by (subst joint_distribution_commute) auto
+    also have "\<P>(t\<circ>OB) {t obs} = (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'})"
+      using setsum_real_distribution(2)[of "t\<circ>OB" fst "t obs", symmetric]
+      using real_distribution_order'[of fst _ "t\<circ>OB" "t obs"] by (auto intro!: setsum_cong)
+    also have "\<P>(t\<circ>OB | fst) {(t obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(t\<circ>OB|fst) {(t obs, k')} * \<P>(fst) {k'}) =
+      \<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / (\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'})"
+      using CP_t_K[OF `k\<in>keys` obs] CP_t_K[OF _ obs] `real (card ?S) \<noteq> 0`
+      by (simp only: setsum_right_distrib[symmetric] ac_simps
+          mult_divide_mult_cancel_left[OF `real (card ?S) \<noteq> 0`]
+        cong: setsum_cong)
+    also have "(\<Sum>k'\<in>keys. \<P>(OB|fst) {(obs, k')} * \<P>(fst) {k'}) = \<P>(OB) {obs}"
+      using setsum_real_distribution(2)[of OB fst obs, symmetric]
+      using real_distribution_order'[of fst _ OB obs] by (auto intro!: setsum_cong)
+    also have "\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / \<P>(OB) {obs} = \<P>(fst | OB) {(k, obs)}"
+      using real_distribution_order'[of fst k OB obs]
+      by (subst joint_distribution_commute) auto
+    finally have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(fst | OB) {(k, obs)}" . }
+  note CP_T_eq_CP_O = this
+
+  let "?H obs" = "(\<Sum>k\<in>keys. \<P>(fst|OB) {(k, obs)} * log b (\<P>(fst|OB) {(k, obs)})) :: real"
+  let "?Ht obs" = "(\<Sum>k\<in>keys. \<P>(fst|t\<circ>OB) {(k, obs)} * log b (\<P>(fst|t\<circ>OB) {(k, obs)})) :: real"
+
+  note [[simproc del: finite_information_space.mult_log]]
+
+  { fix obs assume obs: "obs \<in> OB`msgs"
+    have "?H obs = (\<Sum>k\<in>keys. \<P>(fst|t\<circ>OB) {(k, t obs)} * log b (\<P>(fst|t\<circ>OB) {(k, t obs)}))"
+      using CP_T_eq_CP_O[OF _ obs]
+      by simp
+    then have "?H obs = ?Ht (t obs)" . }
+  note * = this
+
+  have **: "\<And>x f A. (\<Sum>y\<in>t-`{x}\<inter>A. f y (t y)) = (\<Sum>y\<in>t-`{x}\<inter>A. f y x)" by auto
+
+  { fix x
+    have *: "(\<lambda>x. t (OB x)) -` {t (OB x)} \<inter> msgs =
+      (\<Union>obs\<in>?S (OB x). OB -` {obs} \<inter> msgs)" by auto
+    have df: "disjoint_family_on (\<lambda>obs. OB -` {obs} \<inter> msgs) (?S (OB x))"
+      unfolding disjoint_family_on_def by auto
+    have "\<P>(t\<circ>OB) {t (OB x)} = (\<Sum>obs\<in>?S (OB x). \<P>(OB) {obs})"
+      unfolding distribution_def comp_def
+      using real_finite_measure_finite_Union[OF _ df]
+      by (force simp add: * intro!: setsum_nonneg) }
+  note P_t_sum_P_O = this
+
+  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)"
+    apply (subst SIGMA_image_vimage[symmetric, of "OB`msgs" t])
+    apply (subst setsum_reindex)
+    apply (fastsimp intro!: inj_onI)
+    apply simp
+    apply (subst setsum_Sigma[symmetric, unfolded split_def])
+    using finite_space apply fastsimp
+    using finite_space apply fastsimp
+    apply (safe intro!: setsum_cong)
+    using P_t_sum_P_O
+    by (simp add: setsum_divide_distrib[symmetric] field_simps **
+                  setsum_right_distrib[symmetric] setsum_left_distrib[symmetric])
+  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 .
+qed
+
+theorem "\<I>(fst ; OB) \<le> real (card observations) * log b (real n + 1)"
+proof -
+  have "\<I>(fst ; OB) = \<H>(fst) - \<H>(fst | OB)"
+    using mutual_information_eq_entropy_conditional_entropy .
+  also have "\<dots> = \<H>(fst) - \<H>(fst | t\<circ>OB)"
+    unfolding ce_OB_eq_ce_t ..
+  also have "\<dots> = \<H>(t\<circ>OB) - \<H>(t\<circ>OB | fst)"
+    unfolding entropy_chain_rule[symmetric] sign_simps
+    by (subst entropy_commute) simp
+  also have "\<dots> \<le> \<H>(t\<circ>OB)"
+    using conditional_entropy_positive[of "t\<circ>OB" fst] by simp
+  also have "\<dots> \<le> log b (real (card ((t\<circ>OB)`msgs)))"
+    using entropy_le_card[of "t\<circ>OB"] by simp
+  also have "\<dots> \<le> log b (real (n + 1)^card observations)"
+    using card_T_bound not_empty
+    by (auto intro!: log_le simp: card_gt_0_iff power_real_of_nat simp del: real_of_nat_power)
+  also have "\<dots> = real (card observations) * log b (real n + 1)"
+    by (simp add: log_nat_power real_of_nat_Suc)
+  finally show ?thesis  .
+qed
+
+end
+
+end
\ No newline at end of file